merged.
authorhuffman
Mon, 05 Jan 2009 07:54:16 -0800
changeset 29354 6ef5ddf22d3a
parent 29353 3d2e35c23c66 (current diff)
parent 29350 c7735554d291 (diff)
child 29355 642cac18e155
child 29371 bab4e907d881
merged.
src/FOL/ex/LocaleTest.thy
src/HOL/Complex/Fundamental_Theorem_Algebra.thy
src/HOL/Complex/README.html
src/HOL/Complex/document/root.tex
src/HOL/Hyperreal/SEQ.thy
src/HOL/Library/Dense_Linear_Order.thy
src/HOL/Real/HahnBanach/Bounds.thy
src/HOL/Real/HahnBanach/FunctionNorm.thy
src/HOL/Real/HahnBanach/FunctionOrder.thy
src/HOL/Real/HahnBanach/HahnBanach.thy
src/HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
src/HOL/Real/HahnBanach/HahnBanachLemmas.thy
src/HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
src/HOL/Real/HahnBanach/Linearform.thy
src/HOL/Real/HahnBanach/NormedSpace.thy
src/HOL/Real/HahnBanach/README.html
src/HOL/Real/HahnBanach/ROOT.ML
src/HOL/Real/HahnBanach/Subspace.thy
src/HOL/Real/HahnBanach/VectorSpace.thy
src/HOL/Real/HahnBanach/ZornLemma.thy
src/HOL/Real/HahnBanach/document/root.bib
src/HOL/Real/HahnBanach/document/root.tex
src/HOL/Real/RealVector.thy
src/Pure/Thy/thy_edit.ML
--- a/Admin/isatest/isatest-stats	Mon Dec 29 11:04:27 2008 -0800
+++ b/Admin/isatest/isatest-stats	Mon Jan 05 07:54:16 2009 -0800
@@ -31,28 +31,33 @@
   HOL-Word \
   HOL-ex \
   ZF \
-  ZF-Constructible\
+  ZF-Constructible \
   ZF-UNITY"
 
 AFP_SESSIONS="\
-  CoreC++\
-  LinearQuantifierElim\
-  HOL-DiskPaxos\
-  HOL-Fermat3_4\
-  HOL-Flyspeck-Tame\
-  HOL-Group-Ring-Module\
-  HOL-JinjaThreads\
-  HOL-Jinja\
-  HOL-JiveDataStoreModel\
-  HOL-POPLmark-deBruijn\
-  HOL-Program-Conflict-Analysis\
-  HOL-RSAPSS\
-  HOL-Recursion-Theory-I\
-  HOL-SumSquares\
-  HOL-Topology\
-  HOL-Valuation\
-  Simpl-BDD\
-  Simpl"
+  CoreC++ \
+  HOL-BytecodeLogicJmlTypes \
+  HOL-DiskPaxos \
+  HOL-Fermat3_4 \
+  HOL-Flyspeck-Tame \
+  HOL-Group-Ring-Module \
+  HOL-Jinja \
+  HOL-JinjaThreads \
+  HOL-JiveDataStoreModel \
+  HOL-POPLmark-deBruijn \
+  HOL-Program-Conflict-Analysis \
+  HOL-RSAPSS \
+  HOL-Recursion-Theory-I \
+  HOL-SATSolverVerification \
+  HOL-SIFPL \
+  HOL-SenSocialChoice \
+  HOL-Slicing \
+  HOL-SumSquares \
+  HOL-Topology \
+  HOL-Valuation \
+  LinearQuantifierElim \
+  Simpl \
+  Simpl-BDD"
 
 for PLATFORM in $PLATFORMS
 do
--- a/Admin/isatest/settings/at-mac-poly-5.1-para	Mon Dec 29 11:04:27 2008 -0800
+++ b/Admin/isatest/settings/at-mac-poly-5.1-para	Mon Jan 05 07:54:16 2009 -0800
@@ -4,7 +4,7 @@
   ML_SYSTEM="polyml-5.2.1"
   ML_PLATFORM="x86-darwin"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
-  ML_OPTIONS="--immutable 800 --mutable 1200"
+  ML_OPTIONS="--mutable 500 --immutable 1500"
 
 
 ISABELLE_HOME_USER=~/isabelle-at-mac-poly-e
--- a/NEWS	Mon Dec 29 11:04:27 2008 -0800
+++ b/NEWS	Mon Jan 05 07:54:16 2009 -0800
@@ -130,7 +130,7 @@
 consider declaring a new locale with additional type constraints on the
 parameters (context element "constrains").
 
-* Dropped "locale (open)".  INCOMPATBILITY.
+* Dropped "locale (open)".  INCOMPATIBILITY.
 
 * Interpretation commands no longer attempt to simplify goal.
 INCOMPATIBILITY: in rare situations the generated goal differs.  Use
@@ -139,6 +139,36 @@
 * Interpretation commands no longer accept interpretation attributes.
 INCOMPATBILITY.
 
+* Complete re-implementation of locales.  INCOMPATIBILITY.
+The most important changes are listed below.  See documentation
+(forthcoming) and tutorial (also forthcoming) for details.
+
+- In locale expressions, instantiation replaces renaming.  Parameters
+must be declared in a for clause.  To aid compatibility with previous
+parameter inheritance, in locale declarations, parameters that are not
+'touched' (instantiation position "_" or omitted) are implicitly added
+with their syntax at the beginning of the for clause.
+
+- Syntax from abbreviations and definitions in locales is available in
+locale expressions and context elements.  The latter is particularly
+useful in locale declarations.
+
+- 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.
+
+- "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.
+
 * Command 'instance': attached definitions no longer accepted.
 INCOMPATIBILITY, use proper 'instantiation' target.
 
@@ -157,11 +187,12 @@
 
 *** HOL ***
 
-* Made repository layout more coherent with logical
-distribution structure:
+* Made source layout more coherent with logical distribution
+structure:
 
     src/HOL/Library/RType.thy ~> src/HOL/Typerep.thy
     src/HOL/Library/Code_Message.thy ~> src/HOL/
+    src/HOL/Library/Dense_Linear_Order.thy ~> src/HOL/
     src/HOL/Library/GCD.thy ~> src/HOL/
     src/HOL/Library/Order_Relation.thy ~> src/HOL/
     src/HOL/Library/Parity.thy ~> src/HOL/
@@ -177,6 +208,7 @@
     src/HOL/Complex/Complex_Main.thy ~> src/HOL/
     src/HOL/Complex/Complex.thy ~> src/HOL/
     src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/
+    src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/
     src/HOL/Hyperreal/Deriv.thy ~> src/HOL/
     src/HOL/Hyperreal/Fact.thy ~> src/HOL/
     src/HOL/Hyperreal/Integration.thy ~> src/HOL/
@@ -186,9 +218,12 @@
     src/HOL/Hyperreal/MacLaurin.thy ~> src/HOL/
     src/HOL/Hyperreal/NthRoot.thy ~> src/HOL/
     src/HOL/Hyperreal/Series.thy ~> src/HOL/
+    src/HOL/Hyperreal/SEQ.thy ~> src/HOL/
     src/HOL/Hyperreal/Taylor.thy ~> src/HOL/
     src/HOL/Hyperreal/Transcendental.thy ~> src/HOL/
     src/HOL/Real/Float ~> src/HOL/Library/
+    src/HOL/Real/HahnBanach ~> src/HOL/HahnBanach
+    src/HOL/Real/RealVector.thy ~> src/HOL/
 
     src/HOL/arith_data.ML ~> src/HOL/Tools
     src/HOL/hologic.ML ~> src/HOL/Tools
--- a/doc-src/IsarAdvanced/Classes/Thy/Classes.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Classes/Thy/Classes.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -375,8 +375,8 @@
 
 text {* \noindent together with a corresponding interpretation: *}
 
-interpretation %quote idem_class:
-  idem ["f \<Colon> (\<alpha>\<Colon>idem) \<Rightarrow> \<alpha>"]
+interpretation %quote idem_class':    (* FIXME proper prefix? *)
+  idem "f \<Colon> (\<alpha>\<Colon>idem) \<Rightarrow> \<alpha>"
 proof qed (rule idem)
 
 text {*
@@ -459,7 +459,7 @@
   of monoids for lists:
 *}
 
-interpretation %quote list_monoid: monoid [append "[]"]
+class_interpretation %quote list_monoid: monoid [append "[]"]
   proof qed auto
 
 text {*
@@ -474,10 +474,10 @@
   "replicate 0 _ = []"
   | "replicate (Suc n) xs = xs @ replicate n xs"
 
-interpretation %quote list_monoid: monoid [append "[]"] where
+class_interpretation %quote list_monoid: monoid [append "[]"] where
   "monoid.pow_nat append [] = replicate"
 proof -
-  interpret monoid [append "[]"] ..
+  class_interpret monoid [append "[]"] ..
   show "monoid.pow_nat append [] = replicate"
   proof
     fix n
--- a/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -681,8 +681,8 @@
 %
 \isatagquote
 \isacommand{interpretation}\isamarkupfalse%
-\ idem{\isacharunderscore}class{\isacharcolon}\isanewline
-\ \ idem\ {\isacharbrackleft}{\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ idem{\isacharunderscore}class{\isacharprime}{\isacharcolon}\ \ \ \ \isanewline
+\ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ \isacommand{qed}\isamarkupfalse%
 \ {\isacharparenleft}rule\ idem{\isacharparenright}%
@@ -843,7 +843,7 @@
 \endisadelimquote
 %
 \isatagquote
-\isacommand{interpretation}\isamarkupfalse%
+\isacommand{class{\isacharunderscore}interpretation}\isamarkupfalse%
 \ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
 \ \isacommand{qed}\isamarkupfalse%
@@ -874,12 +874,12 @@
 \ \ {\isachardoublequoteopen}replicate\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 \ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline
 \isanewline
-\isacommand{interpretation}\isamarkupfalse%
+\isacommand{class{\isacharunderscore}interpretation}\isamarkupfalse%
 \ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\ \isakeyword{where}\isanewline
 \ \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
-\ \ \isacommand{interpret}\isamarkupfalse%
+\ \ \isacommand{class{\isacharunderscore}interpret}\isamarkupfalse%
 \ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{show}\isamarkupfalse%
@@ -1231,7 +1231,7 @@
 \hspace*{0pt}\\
 \hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\
 \hspace*{0pt}\\
-\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int);\\
+\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int)\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\
 \hspace*{0pt}\\
@@ -1258,7 +1258,7 @@
 \hspace*{0pt} ~~~else inverse A{\char95}~(pow{\char95}nat (monoid{\char95}group A{\char95}) (nat (IntInf.{\char126}~k)) x));\\
 \hspace*{0pt}\\
 \hspace*{0pt}val example :~IntInf.int =\\
-\hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int);\\
+\hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int)\\
 \hspace*{0pt}\\
 \hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
--- a/doc-src/IsarAdvanced/Codegen/Thy/ML.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Codegen/Thy/ML.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -24,9 +24,9 @@
   \begin{mldecls}
   @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
   @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
-  @{index_ML Code.add_eqnl: "string * (thm * bool) list Lazy.T -> theory -> theory"} \\
-  @{index_ML Code.map_pre: "(MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory"} \\
-  @{index_ML Code.map_post: "(MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory"} \\
+  @{index_ML Code.add_eqnl: "string * (thm * bool) list lazy -> theory -> theory"} \\
+  @{index_ML Code.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
+  @{index_ML Code.map_post: "(simpset -> simpset) -> theory -> theory"} \\
   @{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option)
     -> theory -> theory"} \\
   @{index_ML Code.del_functrans: "string -> theory -> theory"} \\
@@ -80,7 +80,7 @@
   \begin{mldecls}
   @{index_ML Code_Unit.read_const: "theory -> string -> string"} \\
   @{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\
-  @{index_ML Code_Unit.rewrite_eqn: "MetaSimplifier.simpset -> thm -> thm"} \\
+  @{index_ML Code_Unit.rewrite_eqn: "simpset -> thm -> thm"} \\
   \end{mldecls}
 
   \begin{description}
--- a/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -5,7 +5,7 @@
 
 ML {* no_document use_thys
   ["Efficient_Nat", "Code_Char_chr", "Product_ord", "Imperative_HOL",
-   "~~/src/HOL/Complex/ex/ReflectedFerrack"] *}
+   "~~/src/HOL/ex/ReflectedFerrack"] *}
 
 ML_val {* Code_Target.code_width := 74 *}
 
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -281,9 +281,9 @@
 \hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\
 \hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = True;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun in{\char95}interval (k, l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\
+\hspace*{0pt}fun in{\char95}interval (k,~l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -357,9 +357,9 @@
 \hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\
 \hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun in{\char95}interval (k, l) n = (less{\char95}eq{\char95}nat k n) andalso (less{\char95}eq{\char95}nat n l);\\
+\hspace*{0pt}fun in{\char95}interval (k,~l) n = (less{\char95}eq{\char95}nat k n) andalso (less{\char95}eq{\char95}nat n l);\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -414,9 +414,9 @@
 \hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\
 \hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun in{\char95}interval (k, l) n = less{\char95}eq{\char95}nat k n andalso less{\char95}eq{\char95}nat n l;\\
+\hspace*{0pt}fun in{\char95}interval (k,~l) n = less{\char95}eq{\char95}nat k n andalso less{\char95}eq{\char95}nat n l;\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -164,20 +164,20 @@
 \hspace*{0pt}\\
 \hspace*{0pt}datatype 'a queue = Queue of 'a list * 'a list;\\
 \hspace*{0pt}\\
-\hspace*{0pt}val empty :~'a queue = Queue ([], []);\\
+\hspace*{0pt}val empty :~'a queue = Queue ([],~[])\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun dequeue (Queue ([], [])) = (NONE, Queue ([], []))\\
-\hspace*{0pt} ~| dequeue (Queue (xs, y ::~ys)) = (SOME y, Queue (xs, ys))\\
-\hspace*{0pt} ~| dequeue (Queue (v ::~va, [])) =\\
+\hspace*{0pt}fun dequeue (Queue ([],~[])) = (NONE,~Queue ([],~[]))\\
+\hspace*{0pt} ~| dequeue (Queue (xs,~y ::~ys)) = (SOME y,~Queue (xs,~ys))\\
+\hspace*{0pt} ~| dequeue (Queue (v ::~va,~[])) =\\
 \hspace*{0pt} ~~~let\\
 \hspace*{0pt} ~~~~~val y ::~ys = rev (v ::~va);\\
 \hspace*{0pt} ~~~in\\
-\hspace*{0pt} ~~~~~(SOME y, Queue ([], ys))\\
+\hspace*{0pt} ~~~~~(SOME y,~Queue ([],~ys))\\
 \hspace*{0pt} ~~~end;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun enqueue x (Queue (xs, ys)) = Queue (x ::~xs, ys);\\
+\hspace*{0pt}fun enqueue x (Queue (xs,~ys)) = Queue (x ::~xs,~ys);\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -233,31 +233,31 @@
 \hspace*{0pt}module Example where {\char123}\\
 \hspace*{0pt}\\
 \hspace*{0pt}\\
-\hspace*{0pt}foldla ::~forall a b. (a -> b -> a) -> a -> [b] -> a;\\
+\hspace*{0pt}foldla ::~forall a b.~(a -> b -> a) -> a -> [b] -> a;\\
 \hspace*{0pt}foldla f a [] = a;\\
 \hspace*{0pt}foldla f a (x :~xs) = foldla f (f a x) xs;\\
 \hspace*{0pt}\\
-\hspace*{0pt}rev ::~forall a. [a] -> [a];\\
+\hspace*{0pt}rev ::~forall a.~[a] -> [a];\\
 \hspace*{0pt}rev xs = foldla ({\char92}~xsa x -> x :~xsa) [] xs;\\
 \hspace*{0pt}\\
-\hspace*{0pt}list{\char95}case ::~forall t a. t -> (a -> [a] -> t) -> [a] -> t;\\
+\hspace*{0pt}list{\char95}case ::~forall t a.~t -> (a -> [a] -> t) -> [a] -> t;\\
 \hspace*{0pt}list{\char95}case f1 f2 (a :~list) = f2 a list;\\
 \hspace*{0pt}list{\char95}case f1 f2 [] = f1;\\
 \hspace*{0pt}\\
 \hspace*{0pt}data Queue a = Queue [a] [a];\\
 \hspace*{0pt}\\
-\hspace*{0pt}empty ::~forall a. Queue a;\\
+\hspace*{0pt}empty ::~forall a.~Queue a;\\
 \hspace*{0pt}empty = Queue [] [];\\
 \hspace*{0pt}\\
-\hspace*{0pt}dequeue ::~forall a. Queue a -> (Maybe a, Queue a);\\
-\hspace*{0pt}dequeue (Queue [] []) = (Nothing, Queue [] []);\\
-\hspace*{0pt}dequeue (Queue xs (y :~ys)) = (Just y, Queue xs ys);\\
+\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\
+\hspace*{0pt}dequeue (Queue [] []) = (Nothing,~Queue [] []);\\
+\hspace*{0pt}dequeue (Queue xs (y :~ys)) = (Just y,~Queue xs ys);\\
 \hspace*{0pt}dequeue (Queue (v :~va) []) =\\
 \hspace*{0pt} ~let {\char123}\\
 \hspace*{0pt} ~~~(y :~ys) = rev (v :~va);\\
-\hspace*{0pt} ~{\char125}~in (Just y, Queue [] ys);\\
+\hspace*{0pt} ~{\char125}~in (Just y,~Queue [] ys);\\
 \hspace*{0pt}\\
-\hspace*{0pt}enqueue ::~forall a. a -> Queue a -> Queue a;\\
+\hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\
 \hspace*{0pt}enqueue x (Queue xs ys) = Queue (x :~xs) ys;\\
 \hspace*{0pt}\\
 \hspace*{0pt}{\char125}%
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -54,9 +54,9 @@
 \begin{mldecls}
   \indexml{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\
   \indexml{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\
-  \indexml{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list Lazy.T -> theory -> theory| \\
-  \indexml{Code.map\_pre}\verb|Code.map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\
-  \indexml{Code.map\_post}\verb|Code.map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\
+  \indexml{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list lazy -> theory -> theory| \\
+  \indexml{Code.map\_pre}\verb|Code.map_pre: (simpset -> simpset) -> theory -> theory| \\
+  \indexml{Code.map\_post}\verb|Code.map_post: (simpset -> simpset) -> theory -> theory| \\
   \indexml{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
 \verb|    -> theory -> theory| \\
   \indexml{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\
@@ -126,7 +126,7 @@
 \begin{mldecls}
   \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\
   \indexml{Code\_Unit.head\_eqn}\verb|Code_Unit.head_eqn: theory -> thm -> string * ((string * sort) list * typ)| \\
-  \indexml{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: MetaSimplifier.simpset -> thm -> thm| \\
+  \indexml{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: simpset -> thm -> thm| \\
   \end{mldecls}
 
   \begin{description}
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -87,10 +87,10 @@
 \begin{isamarkuptext}%
 \isatypewriter%
 \noindent%
-\hspace*{0pt}dequeue ::~forall a. Queue a -> (Maybe a, Queue a);\\
-\hspace*{0pt}dequeue (Queue xs (y :~ys)) = (Just y, Queue xs ys);\\
+\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\
+\hspace*{0pt}dequeue (Queue xs (y :~ys)) = (Just y,~Queue xs ys);\\
 \hspace*{0pt}dequeue (Queue xs []) =\\
-\hspace*{0pt} ~(if nulla xs then (Nothing, Queue [] [])\\
+\hspace*{0pt} ~(if nulla xs then (Nothing,~Queue [] [])\\
 \hspace*{0pt} ~~~else dequeue (Queue [] (rev xs)));%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -286,7 +286,7 @@
 \hspace*{0pt} ~neutral ::~a;\\
 \hspace*{0pt}{\char125};\\
 \hspace*{0pt}\\
-\hspace*{0pt}pow ::~forall a. (Monoid a) => Nat -> a -> a;\\
+\hspace*{0pt}pow ::~forall a.~(Monoid a) => Nat -> a -> a;\\
 \hspace*{0pt}pow Zero{\char95}nat a = neutral;\\
 \hspace*{0pt}pow (Suc n) a = mult a (pow n a);\\
 \hspace*{0pt}\\
@@ -346,7 +346,7 @@
 \hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\
 \hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\
 \hspace*{0pt}\\
-\hspace*{0pt}type 'a monoid = {\char123}Program{\char95}{\char95}semigroup{\char95}monoid :~'a semigroup, neutral :~'a{\char125};\\
+\hspace*{0pt}type 'a monoid = {\char123}Program{\char95}{\char95}semigroup{\char95}monoid :~'a semigroup,~neutral :~'a{\char125};\\
 \hspace*{0pt}fun semigroup{\char95}monoid (A{\char95}:'a monoid) = {\char35}Program{\char95}{\char95}semigroup{\char95}monoid A{\char95};\\
 \hspace*{0pt}fun neutral (A{\char95}:'a monoid) = {\char35}neutral A{\char95};\\
 \hspace*{0pt}\\
@@ -356,7 +356,7 @@
 \hspace*{0pt}fun plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n)\\
 \hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\
 \hspace*{0pt}\\
-\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat;\\
+\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat\\
 \hspace*{0pt} ~| mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\
@@ -364,12 +364,12 @@
 \hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\
 \hspace*{0pt}\\
 \hspace*{0pt}val monoid{\char95}nat =\\
-\hspace*{0pt} ~{\char123}Program{\char95}{\char95}semigroup{\char95}monoid = semigroup{\char95}nat, neutral = neutral{\char95}nat{\char125}~:\\
+\hspace*{0pt} ~{\char123}Program{\char95}{\char95}semigroup{\char95}monoid = semigroup{\char95}nat,~neutral = neutral{\char95}nat{\char125}~:\\
 \hspace*{0pt} ~nat monoid;\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun bexp n = pow monoid{\char95}nat n (Suc (Suc Zero{\char95}nat));\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -675,7 +675,7 @@
 \hspace*{0pt} ~| plus{\char95}nat m Zero{\char95}nat = m\\
 \hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -794,7 +794,7 @@
 \hspace*{0pt} ~~~~~~~~~~~~else collect{\char95}duplicates A{\char95}~xs (z ::~ys) zs)\\
 \hspace*{0pt} ~~~~~else collect{\char95}duplicates A{\char95}~(z ::~xs) (z ::~ys) zs);\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -918,7 +918,7 @@
 \hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\
 \hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\
 \hspace*{0pt}\\
-\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool, less :~'a -> 'a -> bool{\char125};\\
+\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool,~less :~'a -> 'a -> bool{\char125};\\
 \hspace*{0pt}fun less{\char95}eq (A{\char95}:'a ord) = {\char35}less{\char95}eq A{\char95};\\
 \hspace*{0pt}fun less (A{\char95}:'a ord) = {\char35}less A{\char95};\\
 \hspace*{0pt}\\
@@ -930,16 +930,16 @@
 \hspace*{0pt}type 'a order = {\char123}Orderings{\char95}{\char95}preorder{\char95}order :~'a preorder{\char125};\\
 \hspace*{0pt}fun preorder{\char95}order (A{\char95}:'a order) = {\char35}Orderings{\char95}{\char95}preorder{\char95}order A{\char95};\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun less{\char95}eqa (A1{\char95}, A2{\char95}) B{\char95}~(x1, y1) (x2, y2) =\\
+\hspace*{0pt}fun less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\
 \hspace*{0pt} ~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\
 \hspace*{0pt} ~~~eqop A1{\char95}~x1 x2 andalso\\
 \hspace*{0pt} ~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2\\
-\hspace*{0pt} ~| less{\char95}eqa (A1{\char95}, A2{\char95}) B{\char95}~(x1, y1) (x2, y2) =\\
+\hspace*{0pt} ~| less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\
 \hspace*{0pt} ~~~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\
 \hspace*{0pt} ~~~~~eqop A1{\char95}~x1 x2 andalso\\
 \hspace*{0pt} ~~~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2;\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1052,10 +1052,10 @@
 \hspace*{0pt} ~| list{\char95}all2 p xs [] = null xs\\
 \hspace*{0pt} ~| list{\char95}all2 p [] ys = null ys;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun eq{\char95}monotype (Mono (tyco1, typargs1)) (Mono (tyco2, typargs2)) =\\
+\hspace*{0pt}fun eq{\char95}monotype (Mono (tyco1,~typargs1)) (Mono (tyco2,~typargs2)) =\\
 \hspace*{0pt} ~eq{\char95}nat tyco1 tyco2 andalso list{\char95}all2 eq{\char95}monotype typargs1 typargs2;\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1108,12 +1108,12 @@
 \begin{isamarkuptext}%
 \isatypewriter%
 \noindent%
-\hspace*{0pt}strict{\char95}dequeue ::~forall a. Queue a -> (a, Queue a);\\
-\hspace*{0pt}strict{\char95}dequeue (Queue xs (y :~ys)) = (y, Queue xs ys);\\
+\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\
+\hspace*{0pt}strict{\char95}dequeue (Queue xs (y :~ys)) = (y,~Queue xs ys);\\
 \hspace*{0pt}strict{\char95}dequeue (Queue xs []) =\\
 \hspace*{0pt} ~let {\char123}\\
 \hspace*{0pt} ~~~(y :~ys) = rev xs;\\
-\hspace*{0pt} ~{\char125}~in (y, Queue [] ys);%
+\hspace*{0pt} ~{\char125}~in (y,~Queue [] ys);%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1204,13 +1204,13 @@
 \begin{isamarkuptext}%
 \isatypewriter%
 \noindent%
-\hspace*{0pt}empty{\char95}queue ::~forall a. a;\\
+\hspace*{0pt}empty{\char95}queue ::~forall a.~a;\\
 \hspace*{0pt}empty{\char95}queue = error {\char34}empty{\char95}queue{\char34};\\
 \hspace*{0pt}\\
-\hspace*{0pt}strict{\char95}dequeue' ::~forall a. Queue a -> (a, Queue a);\\
+\hspace*{0pt}strict{\char95}dequeue' ::~forall a.~Queue a -> (a,~Queue a);\\
 \hspace*{0pt}strict{\char95}dequeue' (Queue xs []) =\\
 \hspace*{0pt} ~(if nulla xs then empty{\char95}queue else strict{\char95}dequeue' (Queue [] (rev xs)));\\
-\hspace*{0pt}strict{\char95}dequeue' (Queue xs (y :~ys)) = (y, Queue xs ys);%
+\hspace*{0pt}strict{\char95}dequeue' (Queue xs (y :~ys)) = (y,~Queue xs ys);%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarAdvanced/Codegen/Thy/examples/Example.hs	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Codegen/Thy/examples/Example.hs	Mon Jan 05 07:54:16 2009 -0800
@@ -1,3 +1,5 @@
+{-# OPTIONS_GHC -fglasgow-exts #-}
+
 module Example where {
 
 
--- a/doc-src/IsarAdvanced/Codegen/Thy/examples/example.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Codegen/Thy/examples/example.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -11,7 +11,7 @@
 
 datatype 'a queue = Queue of 'a list * 'a list;
 
-val empty : 'a queue = Queue ([], []);
+val empty : 'a queue = Queue ([], [])
 
 fun dequeue (Queue ([], [])) = (NONE, Queue ([], []))
   | dequeue (Queue (xs, y :: ys)) = (SOME y, Queue (xs, ys))
--- a/doc-src/IsarAdvanced/Functions/Thy/document/Functions.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarAdvanced/Functions/Thy/document/Functions.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -440,7 +440,7 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The the key to this error message is the matrix at the bottom. The rows
+The key to this error message is the matrix at the bottom. The rows
   of that matrix correspond to the different recursive calls (In our
   case, there is just one). The columns are the function's arguments 
   (expressed through different measure functions, which map the
@@ -674,7 +674,7 @@
 {\isacharbar}\ {\isachardoublequoteopen}And\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
 This definition is useful, because the equations can directly be used
-  as simplification rules rules. But the patterns overlap: For example,
+  as simplification rules. But the patterns overlap: For example,
   the expression \isa{And\ T\ T} is matched by both the first and
   the second equation. By default, Isabelle makes the patterns disjoint by
   splitting them up, producing instances:%
@@ -829,13 +829,21 @@
   either \isa{{\isadigit{0}}}, \isa{{\isadigit{1}}} or \isa{n\ {\isacharplus}\ {\isadigit{2}}}:
 
   \begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}n{\isachardot}\ x\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P%
+\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}n{\isachardot}\ x\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\isanewline
+\ {\isadigit{2}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline
+\ {\isadigit{3}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline
+\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline
+\ {\isadigit{5}}{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline
+\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline
+\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}n\ na{\isachardot}\isanewline
+\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }n\ {\isacharplus}\ {\isadigit{2}}\ {\isacharequal}\ na\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\isanewline
+\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ na\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ na{\isacharparenright}%
 \end{isabelle}
 
   This is an arithmetic triviality, but unfortunately the
   \isa{arith} method cannot handle this specific form of an
   elimination rule. However, we can use the method \isa{atomize{\isacharunderscore}elim} to do an ad-hoc conversion to a disjunction of
-  existentials, which can then be soved by the arithmetic decision procedure.
+  existentials, which can then be solved by the arithmetic decision procedure.
   Pattern compatibility and termination are automatic as usual.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
--- a/doc-src/IsarOverview/Isar/Logic.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarOverview/Isar/Logic.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -34,54 +34,51 @@
 be enclosed in double quotes. However, we will continue to do so for
 uniformity.
 
-Trivial proofs, in particular those by assumption, should be trivial
-to perform. Proof ``.'' does just that (and a bit more). Thus
-naming of assumptions is often superfluous: *}
+Instead of applying fact @{text a} via the @{text rule} method, we can
+also push it directly onto our goal.  The proof is then immediate,
+which is formally written as ``.'' in Isar: *}
 lemma "A \<longrightarrow> A"
 proof
-  assume "A"
-  show "A" .
+  assume a: "A"
+  from a show "A" .
 qed
 
-text{* To hide proofs by assumption further, \isakeyword{by}@{text"(method)"}
-first applies @{text method} and then tries to solve all remaining subgoals
-by assumption: *}
+text{* We can also push several facts towards a goal, and put another
+rule in between to establish some result that is one step further
+removed.  We illustrate this by introducing a trivial conjunction: *}
 lemma "A \<longrightarrow> A \<and> A"
 proof
-  assume "A"
-  show "A \<and> A" by(rule conjI)
+  assume a: "A"
+  from a and a show "A \<and> A" by(rule conjI)
 qed
 text{*\noindent Rule @{thm[source]conjI} is of course @{thm conjI}.
-A drawback of implicit proofs by assumption is that it
-is no longer obvious where an assumption is used.
 
 Proofs of the form \isakeyword{by}@{text"(rule"}~\emph{name}@{text")"}
-can be abbreviated to ``..''  if \emph{name} refers to one of the
+can be abbreviated to ``..'' if \emph{name} refers to one of the
 predefined introduction rules (or elimination rules, see below): *}
 
 lemma "A \<longrightarrow> A \<and> A"
 proof
-  assume "A"
-  show "A \<and> A" ..
+  assume a: "A"
+  from a and a show "A \<and> A" ..
 qed
 text{*\noindent
 This is what happens: first the matching introduction rule @{thm[source]conjI}
-is applied (first ``.''), then the two subgoals are solved by assumption
-(second ``.''). *}
+is applied (first ``.''), the remaining problem is solved immediately (second ``.''). *}
 
 subsubsection{*Elimination rules*}
 
 text{*A typical elimination rule is @{thm[source]conjE}, $\land$-elimination:
 @{thm[display,indent=5]conjE}  In the following proof it is applied
 by hand, after its first (\emph{major}) premise has been eliminated via
-@{text"[OF AB]"}: *}
+@{text"[OF ab]"}: *}
 lemma "A \<and> B \<longrightarrow> B \<and> A"
 proof
-  assume AB: "A \<and> B"
+  assume ab: "A \<and> B"
   show "B \<and> A"
-  proof (rule conjE[OF AB])  -- {*@{text"conjE[OF AB]"}: @{thm conjE[OF AB]} *}
-    assume "A" "B"
-    show ?thesis ..
+  proof (rule conjE[OF ab])  -- {*@{text"conjE[OF ab]"}: @{thm conjE[OF ab]} *}
+    assume a: "A" and b: "B"
+    from b and a show ?thesis ..
   qed
 qed
 text{*\noindent Note that the term @{text"?thesis"} always stands for the
@@ -106,11 +103,11 @@
 
 lemma "A \<and> B \<longrightarrow> B \<and> A"
 proof
-  assume AB: "A \<and> B"
-  from AB show "B \<and> A"
+  assume ab: "A \<and> B"
+  from ab show "B \<and> A"
   proof
-    assume "A" "B"
-    show ?thesis ..
+    assume a: "A" and b: "B"
+    from b and a show ?thesis ..
   qed
 qed
 
@@ -120,15 +117,16 @@
 such that the proof of each proposition builds on the previous proposition.
 \end{quote}
 The previous proposition can be referred to via the fact @{text this}.
-This greatly reduces the need for explicit naming of propositions:
+This greatly reduces the need for explicit naming of propositions.  We also
+rearrange the additional inner assumptions into proper order for immediate use:
 *}
 lemma "A \<and> B \<longrightarrow> B \<and> A"
 proof
   assume "A \<and> B"
   from this show "B \<and> A"
   proof
-    assume "A" "B"
-    show ?thesis ..
+    assume "B" "A"
+    from this show ?thesis ..
   qed
 qed
 
@@ -199,11 +197,11 @@
     assume nn: "\<not> (\<not> A \<or> \<not> B)"
     have "\<not> A"
     proof
-      assume "A"
+      assume a: "A"
       have "\<not> B"
       proof
-        assume "B"
-        have "A \<and> B" ..
+        assume b: "B"
+        from a and b have "A \<and> B" ..
         with n show False ..
       qed
       hence "\<not> A \<or> \<not> B" ..
@@ -282,28 +280,28 @@
 \isakeyword{assumes} and \isakeyword{shows} elements which allow direct
 naming of assumptions: *}
 
-lemma assumes AB: "large_A \<and> large_B"
+lemma assumes ab: "large_A \<and> large_B"
   shows "large_B \<and> large_A" (is "?B \<and> ?A")
 proof
-  from AB show "?B" ..
+  from ab show "?B" ..
 next
-  from AB show "?A" ..
+  from ab show "?A" ..
 qed
 text{*\noindent Note the difference between @{text ?AB}, a term, and
-@{text AB}, a fact.
+@{text ab}, a fact.
 
 Finally we want to start the proof with $\land$-elimination so we
 don't have to perform it twice, as above. Here is a slick way to
 achieve this: *}
 
-lemma assumes AB: "large_A \<and> large_B"
+lemma assumes ab: "large_A \<and> large_B"
   shows "large_B \<and> large_A" (is "?B \<and> ?A")
-using AB
+using ab
 proof
-  assume "?A" "?B" show ?thesis ..
+  assume "?B" "?A" thus ?thesis ..
 qed
 text{*\noindent Command \isakeyword{using} can appear before a proof
-and adds further facts to those piped into the proof. Here @{text AB}
+and adds further facts to those piped into the proof. Here @{text ab}
 is the only such fact and it triggers $\land$-elimination. Another
 frequent idiom is as follows:
 \begin{center}
@@ -319,23 +317,23 @@
 not be what we had in mind.
 A simple ``@{text"-"}'' prevents this \emph{faux pas}: *}
 
-lemma assumes AB: "A \<or> B" shows "B \<or> A"
+lemma assumes ab: "A \<or> B" shows "B \<or> A"
 proof -
-  from AB show ?thesis
+  from ab show ?thesis
   proof
-    assume A show ?thesis ..
+    assume A thus ?thesis ..
   next
-    assume B show ?thesis ..
+    assume B thus ?thesis ..
   qed
 qed
 text{*\noindent Alternatively one can feed @{prop"A \<or> B"} directly
 into the proof, thus triggering the elimination rule: *}
-lemma assumes AB: "A \<or> B" shows "B \<or> A"
-using AB
+lemma assumes ab: "A \<or> B" shows "B \<or> A"
+using ab
 proof
-  assume A show ?thesis ..
+  assume A thus ?thesis ..
 next
-  assume B show ?thesis ..
+  assume B thus ?thesis ..
 qed
 text{* \noindent Remember that eliminations have priority over
 introductions.
@@ -416,7 +414,7 @@
   proof              -- "@{thm[source]exE}: @{thm exE}"
     fix x
     assume "P(f x)"
-    show ?thesis ..  -- "@{thm[source]exI}: @{thm exI}"
+    thus ?thesis ..  -- "@{thm[source]exI}: @{thm exI}"
   qed
 qed
 text{*\noindent Explicit $\exists$-elimination as seen above can become
@@ -499,12 +497,12 @@
       assume "y \<in> ?S"
       hence "y \<notin> f y"   by simp
       hence "y \<notin> ?S"    by(simp add: `?S = f y`)
-      thus False         by contradiction
+      with `y \<in> ?S` show False by contradiction
     next
       assume "y \<notin> ?S"
       hence "y \<in> f y"   by simp
       hence "y \<in> ?S"    by(simp add: `?S = f y`)
-      thus False         by contradiction
+      with `y \<notin> ?S` show False by contradiction
     qed
   qed
 qed
--- a/doc-src/IsarOverview/Isar/document/Logic.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/IsarOverview/Isar/document/Logic.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -97,9 +97,9 @@
 be enclosed in double quotes. However, we will continue to do so for
 uniformity.
 
-Trivial proofs, in particular those by assumption, should be trivial
-to perform. Proof ``.'' does just that (and a bit more). Thus
-naming of assumptions is often superfluous:%
+Instead of applying fact \isa{a} via the \isa{rule} method, we can
+also push it directly onto our goal.  The proof is then immediate,
+which is formally written as ``.'' in Isar:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
@@ -113,8 +113,9 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{from}\isamarkupfalse%
+\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -127,9 +128,9 @@
 \endisadelimproof
 %
 \begin{isamarkuptext}%
-To hide proofs by assumption further, \isakeyword{by}\isa{{\isacharparenleft}method{\isacharparenright}}
-first applies \isa{method} and then tries to solve all remaining subgoals
-by assumption:%
+We can also push several facts towards a goal, and put another
+rule in between to establish some result that is one step further
+removed.  We illustrate this by introducing a trivial conjunction:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
@@ -143,8 +144,9 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{from}\isamarkupfalse%
+\ a\ \isakeyword{and}\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}A\ {\isasymand}\ A{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 {\isacharparenleft}rule\ conjI{\isacharparenright}\isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -158,11 +160,9 @@
 %
 \begin{isamarkuptext}%
 \noindent Rule \isa{conjI} is of course \isa{{\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isasymand}\ {\isacharquery}Q}.
-A drawback of implicit proofs by assumption is that it
-is no longer obvious where an assumption is used.
 
 Proofs of the form \isakeyword{by}\isa{{\isacharparenleft}rule}~\emph{name}\isa{{\isacharparenright}}
-can be abbreviated to ``..''  if \emph{name} refers to one of the
+can be abbreviated to ``..'' if \emph{name} refers to one of the
 predefined introduction rules (or elimination rules, see below):%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -177,8 +177,9 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{from}\isamarkupfalse%
+\ a\ \isakeyword{and}\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}A\ {\isasymand}\ A{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -193,8 +194,7 @@
 \begin{isamarkuptext}%
 \noindent
 This is what happens: first the matching introduction rule \isa{conjI}
-is applied (first ``.''), then the two subgoals are solved by assumption
-(second ``.'').%
+is applied (first ``.''), the remaining problem is solved immediately (second ``.'').%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -208,7 +208,7 @@
 \ \ \ \ \ {\isasymlbrakk}{\isacharquery}P\ {\isasymand}\ {\isacharquery}Q{\isacharsemicolon}\ {\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R%
 \end{isabelle}  In the following proof it is applied
 by hand, after its first (\emph{major}) premise has been eliminated via
-\isa{{\isacharbrackleft}OF\ AB{\isacharbrackright}}:%
+\isa{{\isacharbrackleft}OF\ ab{\isacharbrackright}}:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
@@ -222,17 +222,18 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ AB{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
+\ ab{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}B\ {\isasymand}\ A{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
-\ {\isacharparenleft}rule\ conjE{\isacharbrackleft}OF\ AB{\isacharbrackright}{\isacharparenright}\ \ %
-\isamarkupcmt{\isa{conjE{\isacharbrackleft}OF\ AB{\isacharbrackright}}: \isa{{\isacharparenleft}{\isasymlbrakk}A{\isacharsemicolon}\ B{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}R}%
+\ {\isacharparenleft}rule\ conjE{\isacharbrackleft}OF\ ab{\isacharbrackright}{\isacharparenright}\ \ %
+\isamarkupcmt{\isa{conjE{\isacharbrackleft}OF\ ab{\isacharbrackright}}: \isa{{\isacharparenleft}{\isasymlbrakk}A{\isacharsemicolon}\ B{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}R}%
 }
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ \isakeyword{and}\ b{\isacharcolon}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{from}\isamarkupfalse%
+\ b\ \isakeyword{and}\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
@@ -279,15 +280,16 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ AB{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
+\ ab{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{from}\isamarkupfalse%
-\ AB\ \isacommand{show}\isamarkupfalse%
+\ ab\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}B\ {\isasymand}\ A{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ \isakeyword{and}\ b{\isacharcolon}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{from}\isamarkupfalse%
+\ b\ \isakeyword{and}\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
@@ -308,7 +310,8 @@
 such that the proof of each proposition builds on the previous proposition.
 \end{quote}
 The previous proposition can be referred to via the fact \isa{this}.
-This greatly reduces the need for explicit naming of propositions:%
+This greatly reduces the need for explicit naming of propositions.  We also
+rearrange the additional inner assumptions into proper order for immediate use:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
@@ -329,8 +332,9 @@
 \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{from}\isamarkupfalse%
+\ this\ \isacommand{show}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
@@ -455,14 +459,15 @@
 \ \ \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
 \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isasymnot}\ B{\isachardoublequoteclose}\isanewline
 \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
-\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
+\ b{\isacharcolon}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
+\ a\ \isakeyword{and}\ b\ \isacommand{have}\isamarkupfalse%
 \ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
@@ -622,7 +627,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
-\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequoteopen}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequoteclose}\isanewline
+\ \isakeyword{assumes}\ ab{\isacharcolon}\ {\isachardoublequoteopen}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{shows}\ {\isachardoublequoteopen}large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequoteclose}{\isacharparenright}\isanewline
 %
 \isadelimproof
@@ -633,13 +638,13 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{from}\isamarkupfalse%
-\ AB\ \isacommand{show}\isamarkupfalse%
+\ ab\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isacharquery}B{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{next}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{from}\isamarkupfalse%
-\ AB\ \isacommand{show}\isamarkupfalse%
+\ ab\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isacharquery}A{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -653,7 +658,7 @@
 %
 \begin{isamarkuptext}%
 \noindent Note the difference between \isa{{\isacharquery}AB}, a term, and
-\isa{AB}, a fact.
+\isa{ab}, a fact.
 
 Finally we want to start the proof with $\land$-elimination so we
 don't have to perform it twice, as above. Here is a slick way to
@@ -661,7 +666,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
-\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequoteopen}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequoteclose}\isanewline
+\ \isakeyword{assumes}\ ab{\isacharcolon}\ {\isachardoublequoteopen}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{shows}\ {\isachardoublequoteopen}large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequoteclose}{\isacharparenright}\isanewline
 %
 \isadelimproof
@@ -670,11 +675,11 @@
 %
 \isatagproof
 \isacommand{using}\isamarkupfalse%
-\ AB\isanewline
+\ ab\isanewline
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}{\isacharquery}A{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharquery}B{\isachardoublequoteclose}\ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isacharquery}B{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharquery}A{\isachardoublequoteclose}\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -688,7 +693,7 @@
 %
 \begin{isamarkuptext}%
 \noindent Command \isakeyword{using} can appear before a proof
-and adds further facts to those piped into the proof. Here \isa{AB}
+and adds further facts to those piped into the proof. Here \isa{ab}
 is the only such fact and it triggers $\land$-elimination. Another
 frequent idiom is as follows:
 \begin{center}
@@ -706,7 +711,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
-\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{shows}\ {\isachardoublequoteopen}B\ {\isasymor}\ A{\isachardoublequoteclose}\isanewline
+\ \isakeyword{assumes}\ ab{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{shows}\ {\isachardoublequoteopen}B\ {\isasymor}\ A{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
 %
@@ -716,18 +721,18 @@
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
 \ \ \isacommand{from}\isamarkupfalse%
-\ AB\ \isacommand{show}\isamarkupfalse%
+\ ab\ \isacommand{show}\isamarkupfalse%
 \ {\isacharquery}thesis\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ A\ \isacommand{show}\isamarkupfalse%
+\ A\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{next}\isamarkupfalse%
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ B\ \isacommand{show}\isamarkupfalse%
+\ B\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
@@ -747,7 +752,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
-\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{shows}\ {\isachardoublequoteopen}B\ {\isasymor}\ A{\isachardoublequoteclose}\isanewline
+\ \isakeyword{assumes}\ ab{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{shows}\ {\isachardoublequoteopen}B\ {\isasymor}\ A{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
 %
@@ -755,17 +760,17 @@
 %
 \isatagproof
 \isacommand{using}\isamarkupfalse%
-\ AB\isanewline
+\ ab\isanewline
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ A\ \isacommand{show}\isamarkupfalse%
+\ A\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{next}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ B\ \isacommand{show}\isamarkupfalse%
+\ B\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -945,7 +950,7 @@
 \ x\isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
 \ {\isachardoublequoteopen}P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ \ \ \ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \ \ %
 \isamarkupcmt{\isa{exI}: \isa{{\isacharquery}P\ {\isacharquery}x\ {\isasymLongrightarrow}\ {\isasymexists}x{\isachardot}\ {\isacharquery}P\ x}%
@@ -1155,8 +1160,9 @@
 \ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
 \ {\isachardoublequoteopen}y\ {\isasymnotin}\ {\isacharquery}S{\isachardoublequoteclose}\ \ \ \ \isacommand{by}\isamarkupfalse%
 {\isacharparenleft}simp\ add{\isacharcolon}\ {\isacharbackquoteopen}{\isacharquery}S\ {\isacharequal}\ f\ y{\isacharbackquoteclose}{\isacharparenright}\isanewline
-\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
-\ False\ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
+\ {\isacharbackquoteopen}y\ {\isasymin}\ {\isacharquery}S{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse%
+\ False\ \isacommand{by}\isamarkupfalse%
 \ contradiction\isanewline
 \ \ \ \ \isacommand{next}\isamarkupfalse%
 \isanewline
@@ -1168,8 +1174,9 @@
 \ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
 \ {\isachardoublequoteopen}y\ {\isasymin}\ {\isacharquery}S{\isachardoublequoteclose}\ \ \ \ \isacommand{by}\isamarkupfalse%
 {\isacharparenleft}simp\ add{\isacharcolon}\ {\isacharbackquoteopen}{\isacharquery}S\ {\isacharequal}\ f\ y{\isacharbackquoteclose}{\isacharparenright}\isanewline
-\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
-\ False\ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
+\ {\isacharbackquoteopen}y\ {\isasymnotin}\ {\isacharquery}S{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse%
+\ False\ \isacommand{by}\isamarkupfalse%
 \ contradiction\isanewline
 \ \ \ \ \isacommand{qed}\isamarkupfalse%
 \isanewline
--- a/doc-src/LaTeXsugar/Sugar/document/Sugar.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/LaTeXsugar/Sugar/document/Sugar.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -148,7 +148,7 @@
 \begin{isamarkuptext}%
 If you print anything, especially theorems, containing
 schematic variables they are prefixed with a question mark:
-\verb!@!\verb!{thm conjI}! results in \isa{{\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isasymand}\ {\isacharquery}Q}. Most of the time
+\verb!@!\verb!{thm conjI}! results in \isa{{\isasymlbrakk}P{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isasymand}\ Q}. Most of the time
 you would rather not see the question marks. There is an attribute
 \verb!no_vars! that you can attach to the theorem that turns its
 schematic into ordinary free variables: \verb!@!\verb!{thm conjI[no_vars]}!
--- a/doc-src/Locales/Locales/Examples.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/Locales/Locales/Examples.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -610,7 +610,7 @@
   Changes to the locale hierarchy may be declared
   with the \isakeyword{interpretation} command. *}
 
-  interpretation %visible total_order \<subseteq> lattice
+  sublocale %visible total_order \<subseteq> lattice
 
 txt {* This enters the context of locale @{text total_order}, in
   which the goal @{subgoals [display]} must be shown.  First, the
@@ -652,7 +652,7 @@
 
 text {* Similarly, total orders are distributive lattices. *}
 
-  interpretation total_order \<subseteq> distrib_lattice
+  sublocale total_order \<subseteq> distrib_lattice
   proof unfold_locales
     fix %"proof" x y z
     show "x \<sqinter> (y \<squnion> z) = x \<sqinter> y \<squnion> x \<sqinter> z" (is "?l = ?r")
--- a/doc-src/Locales/Locales/Examples1.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/Locales/Locales/Examples1.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -45,7 +45,7 @@
   @{text partial_order} in the global context of the theory.  The
   parameter @{term le} is replaced by @{term "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"}. *} 
 
-  interpretation %visible nat: partial_order ["op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"]
+  interpretation %visible nat: partial_order "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"
 txt {* The locale name is succeeded by a \emph{parameter
   instantiation}.  In general, this is a list of terms, which refer to
   the parameters in the order of declaration in the locale.  The
--- a/doc-src/Locales/Locales/Examples2.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/Locales/Locales/Examples2.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -9,7 +9,7 @@
   \isakeyword{where} and require proofs.  The revised command,
   replacing @{text "\<sqsubset>"} by @{text "<"}, is: *}
 
-interpretation %visible nat: partial_order ["op \<le> :: [nat, nat] \<Rightarrow> bool"]
+interpretation %visible nat: partial_order "op \<le> :: [nat, nat] \<Rightarrow> bool"
   where "partial_order.less op \<le> (x::nat) y = (x < y)"
 proof -
   txt {* The goals are @{subgoals [display]}
--- a/doc-src/Locales/Locales/Examples3.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/Locales/Locales/Examples3.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -16,12 +16,12 @@
   \isakeyword{interpret}).  This interpretation is inside the proof of the global
   interpretation.  The third revision of the example illustrates this.  *}
 
-interpretation %visible nat: partial_order ["op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"]
+interpretation %visible nat: partial_order "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"
   where "partial_order.less (op \<le>) (x::nat) y = (x < y)"
 proof -
   show "partial_order (op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool)"
     by unfold_locales auto
-  then interpret nat: partial_order ["op \<le> :: [nat, nat] \<Rightarrow> bool"] .
+  then interpret nat: partial_order "op \<le> :: [nat, nat] \<Rightarrow> bool" .
   show "partial_order.less (op \<le>) (x::nat) y = (x < y)"
     unfolding nat.less_def by auto
 qed
@@ -48,7 +48,7 @@
   interpretation is reproduced in order to give an example of a more
   elaborate interpretation proof.  *}
 
-interpretation %visible nat: lattice ["op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"]
+interpretation %visible nat: lattice "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"
   where "lattice.meet op \<le> (x::nat) y = min x y"
     and "lattice.join op \<le> (x::nat) y = max x y"
 proof -
@@ -63,7 +63,7 @@
     by arith+
   txt {* In order to show the equations, we put ourselves in a
     situation where the lattice theorems can be used in a convenient way. *}
-  then interpret nat: lattice ["op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"] .
+  then interpret nat: lattice "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool" .
   show "lattice.meet op \<le> (x::nat) y = min x y"
     by (bestsimp simp: nat.meet_def nat.is_inf_def)
   show "lattice.join op \<le> (x::nat) y = max x y"
@@ -73,7 +73,7 @@
 text {* That the relation @{text \<le>} is a total order completes this
   sequence of interpretations. *}
 
-interpretation %visible nat: total_order ["op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"]
+interpretation %visible nat: total_order "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"
   by unfold_locales arith
 
 text {* Theorems that are available in the theory at this point are shown in
@@ -130,12 +130,12 @@
   but not a total order.  Interpretation again proceeds
   incrementally. *}
 
-interpretation nat_dvd: partial_order ["op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"]
+interpretation nat_dvd: partial_order "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
   where "partial_order.less op dvd (x::nat) y = (x dvd y \<and> x \<noteq> y)"
 proof -
   show "partial_order (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool)"
     by unfold_locales (auto simp: dvd_def)
-  then interpret nat_dvd: partial_order ["op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"] .
+  then interpret nat_dvd: partial_order "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool" .
   show "partial_order.less op dvd (x::nat) y = (x dvd y \<and> x \<noteq> y)"
     apply (unfold nat_dvd.less_def)
     apply auto
@@ -145,7 +145,7 @@
 text {* Note that there is no symbol for strict divisibility.  Instead,
   interpretation substitutes @{term "x dvd y \<and> x \<noteq> y"}.   *}
 
-interpretation nat_dvd: lattice ["op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"]
+interpretation nat_dvd: lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
   where nat_dvd_meet_eq:
       "lattice.meet op dvd = gcd"
     and nat_dvd_join_eq:
@@ -159,7 +159,7 @@
     apply (rule_tac x = "lcm x y" in exI)
     apply (auto intro: lcm_dvd1 lcm_dvd2 lcm_least)
     done
-  then interpret nat_dvd: lattice ["op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"] .
+  then interpret nat_dvd: lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool" .
   show "lattice.meet op dvd = gcd"
     apply (auto simp add: expand_fun_eq)
     apply (unfold nat_dvd.meet_def)
@@ -203,7 +203,7 @@
 ML %invisible {* reset quick_and_dirty *}
   
 interpretation %visible nat_dvd:
-  distrib_lattice ["op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"]
+  distrib_lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
   apply unfold_locales
   txt {* @{subgoals [display]} *}
   apply (unfold nat_dvd_meet_eq nat_dvd_join_eq)
@@ -262,7 +262,7 @@
   preserving maps can be declared in the following way.  *}
 
   locale order_preserving =
-    partial_order + partial_order le' (infixl "\<preceq>" 50) +
+    partial_order + po': partial_order le' for le' (infixl "\<preceq>" 50) +
     fixes \<phi> :: "'a \<Rightarrow> 'b"
     assumes hom_le: "x \<sqsubseteq> y \<Longrightarrow> \<phi> x \<preceq> \<phi> y"
 
@@ -288,7 +288,8 @@
   obtained by appending the conclusions of the left locale and of the
   right locale.  *}
 
-text {* The locale @{text order_preserving} contains theorems for both
+text {* % FIXME needs update
+  The locale @{text order_preserving} contains theorems for both
   orders @{text \<sqsubseteq>} and @{text \<preceq>}.  How can one refer to a theorem for
   a particular order, @{text \<sqsubseteq>} or @{text \<preceq>}?  Names in locales are
   qualified by the locale parameters.  More precisely, a name is
@@ -297,20 +298,21 @@
 
 context %invisible order_preserving begin
 
-text {*
-  @{thm [source] le.less_le_trans}: @{thm le.less_le_trans}
+text {* % FIXME needs update?
+  @{thm [source] less_le_trans}: @{thm less_le_trans}
 
-  @{thm [source] le_le'_\<phi>.hom_le}: @{thm le_le'_\<phi>.hom_le}
+  @{thm [source] hom_le}: @{thm hom_le}
   *}
 
 text {* When renaming a locale, the morphism is also applied
   to the qualifiers.  Hence theorems for the partial order @{text \<preceq>}
   are qualified by @{text le'}.  For example, @{thm [source]
-  le'.less_le_trans}: @{thm [display, indent=2] le'.less_le_trans} *}
+  po'.less_le_trans}: @{thm [display, indent=2] po'.less_le_trans} *}
 
 end %invisible
 
-text {* This example reveals that there is no infix syntax for the strict
+text {* % FIXME needs update?
+  This example reveals that there is no infix syntax for the strict
   version of @{text \<preceq>}!  This can, of course, not be introduced
   automatically, but it can be declared manually through an abbreviation.
   *}
@@ -319,7 +321,7 @@
     less' (infixl "\<prec>" 50) where "less' \<equiv> partial_order.less le'"
 
 text {* Now the theorem is displayed nicely as
-  @{thm [locale=order_preserving] le'.less_le_trans}.  *}
+  @{thm [locale=order_preserving] po'.less_le_trans}.  *}
 
 text {* Not only names of theorems are qualified.  In fact, all names
   are qualified, in particular names introduced by definitions and
@@ -331,7 +333,7 @@
 text {* Two more locales illustrate working with locale expressions.
   A map @{text \<phi>} is a lattice homomorphism if it preserves meet and join. *}
 
-  locale lattice_hom = lattice + lattice le' (infixl "\<preceq>" 50) +
+  locale lattice_hom = lattice + lat'!: lattice le' for le' (infixl "\<preceq>" 50) +
     fixes \<phi>
     assumes hom_meet:
 	"\<phi> (lattice.meet le x y) = lattice.meet le' (\<phi> x) (\<phi> y)"
@@ -339,14 +341,14 @@
 	"\<phi> (lattice.join le x y) = lattice.join le' (\<phi> x) (\<phi> y)"
 
   abbreviation (in lattice_hom)
-    meet' (infixl "\<sqinter>''" 50) where "meet' \<equiv> le'.meet"
+    meet' (infixl "\<sqinter>''" 50) where "meet' \<equiv> lat'.meet"
   abbreviation (in lattice_hom)
-    join' (infixl "\<squnion>''" 50) where "join' \<equiv> le'.join"
+    join' (infixl "\<squnion>''" 50) where "join' \<equiv> lat'.join"
 
 text {* A homomorphism is an endomorphism if both orders coincide. *}
 
   locale lattice_end =
-    lattice_hom le (infixl "\<sqsubseteq>" 50) le (infixl "\<sqsubseteq>" 50)
+    lattice_hom le le for le (infixl "\<sqsubseteq>" 50)
 
 text {* The inheritance diagram of the situation we have now is shown
   in Figure~\ref{fig:hom}, where the dashed line depicts an
@@ -395,20 +397,20 @@
   preserving.  As the final example of this section, a locale
   interpretation is used to assert this. *}
 
-  interpretation lattice_hom \<subseteq> order_preserving proof unfold_locales
+  sublocale lattice_hom \<subseteq> order_preserving proof unfold_locales
     fix x y
     assume "x \<sqsubseteq> y"
-    then have "y = (x \<squnion> y)" by (simp add: le.join_connection)
+    then have "y = (x \<squnion> y)" by (simp add: join_connection)
     then have "\<phi> y = (\<phi> x \<squnion>' \<phi> y)" by (simp add: hom_join [symmetric])
-    then show "\<phi> x \<preceq> \<phi> y" by (simp add: le'.join_connection)
+    then show "\<phi> x \<preceq> \<phi> y" by (simp add: lat'.join_connection)
   qed
 
 text {* Theorems and other declarations --- syntax, in particular ---
   from the locale @{text order_preserving} are now active in @{text
   lattice_hom}, for example
 
-  @{thm [locale=lattice_hom, source] le'.less_le_trans}:
-  @{thm [locale=lattice_hom] le'.less_le_trans}
+  @{thm [locale=lattice_hom, source] lat'.less_le_trans}:
+  @{thm [locale=lattice_hom] lat'.less_le_trans}
   *}
 
 
--- a/doc-src/Locales/Locales/document/Examples.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/Locales/Locales/document/Examples.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -1222,7 +1222,7 @@
 \endisadelimvisible
 %
 \isatagvisible
-\isacommand{interpretation}\isamarkupfalse%
+\isacommand{sublocale}\isamarkupfalse%
 \ total{\isacharunderscore}order\ {\isasymsubseteq}\ lattice%
 \begin{isamarkuptxt}%
 This enters the context of locale \isa{total{\isacharunderscore}order}, in
@@ -1325,7 +1325,7 @@
 Similarly, total orders are distributive lattices.%
 \end{isamarkuptext}%
 \isamarkuptrue%
-\ \ \isacommand{interpretation}\isamarkupfalse%
+\ \ \isacommand{sublocale}\isamarkupfalse%
 \ total{\isacharunderscore}order\ {\isasymsubseteq}\ distrib{\isacharunderscore}lattice\isanewline
 %
 \isadelimproof
--- a/doc-src/Locales/Locales/document/Examples1.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/Locales/Locales/document/Examples1.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -74,7 +74,7 @@
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}%
+\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}%
 \begin{isamarkuptxt}%
 The locale name is succeeded by a \emph{parameter
   instantiation}.  In general, this is a list of terms, which refer to
--- a/doc-src/Locales/Locales/document/Examples2.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/Locales/Locales/document/Examples2.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -34,7 +34,7 @@
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}nat{\isacharcomma}\ nat{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}nat{\isacharcomma}\ nat{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}%
--- a/doc-src/Locales/Locales/document/Examples3.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/Locales/Locales/document/Examples3.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -43,7 +43,7 @@
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ {\isacharparenleft}op\ {\isasymle}{\isacharparenright}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
@@ -53,7 +53,7 @@
 \ unfold{\isacharunderscore}locales\ auto\isanewline
 \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{interpret}\isamarkupfalse%
-\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}nat{\isacharcomma}\ nat{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}}\isamarkupfalse%
+\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}nat{\isacharcomma}\ nat{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ {\isacharparenleft}op\ {\isasymle}{\isacharparenright}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
@@ -91,8 +91,8 @@
 \begin{isamarkuptext}%
 Further interpretations are necessary to reuse theorems from
   the other locales.  In \isa{lattice} the operations \isa{{\isasymsqinter}} and
-  \isa{{\isasymsqunion}} are substituted by \isa{ord{\isacharunderscore}class{\isachardot}min} and
-  \isa{ord{\isacharunderscore}class{\isachardot}max}.  The entire proof for the
+  \isa{{\isasymsqunion}} are substituted by \isa{min} and
+  \isa{max}.  The entire proof for the
   interpretation is reproduced in order to give an example of a more
   elaborate interpretation proof.%
 \end{isamarkuptext}%
@@ -104,7 +104,7 @@
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ nat{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{where}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
 \ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
@@ -143,7 +143,7 @@
 \isamarkuptrue%
 \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{interpret}\isamarkupfalse%
-\ nat{\isacharcolon}\ lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}}\isamarkupfalse%
+\ nat{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
@@ -174,7 +174,7 @@
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ total{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ nat{\isacharcolon}\ total{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{by}\isamarkupfalse%
 \ unfold{\isacharunderscore}locales\ arith%
 \endisatagvisible
@@ -196,11 +196,11 @@
   \isa{nat{\isachardot}less{\isacharunderscore}def} from locale \isa{partial{\isacharunderscore}order}: \\
   \quad \isa{{\isacharparenleft}{\isacharquery}x\ {\isacharless}\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}x\ {\isasymle}\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}} \\
   \isa{nat{\isachardot}meet{\isacharunderscore}left} from locale \isa{lattice}: \\
-  \quad \isa{ord{\isacharunderscore}class{\isachardot}min\ {\isacharquery}x\ {\isacharquery}y\ {\isasymle}\ {\isacharquery}x} \\
+  \quad \isa{min\ {\isacharquery}x\ {\isacharquery}y\ {\isasymle}\ {\isacharquery}x} \\
   \isa{nat{\isachardot}join{\isacharunderscore}distr} from locale \isa{distrib{\isacharunderscore}lattice}: \\
-  \quad \isa{ord{\isacharunderscore}class{\isachardot}max\ {\isacharquery}x\ {\isacharparenleft}ord{\isacharunderscore}class{\isachardot}min\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ ord{\isacharunderscore}class{\isachardot}min\ {\isacharparenleft}ord{\isacharunderscore}class{\isachardot}max\ {\isacharquery}x\ {\isacharquery}y{\isacharparenright}\ {\isacharparenleft}ord{\isacharunderscore}class{\isachardot}max\ {\isacharquery}x\ {\isacharquery}z{\isacharparenright}} \\
+  \quad \isa{lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharquery}x\ {\isacharparenleft}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharquery}x\ {\isacharquery}y{\isacharparenright}\ {\isacharparenleft}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharquery}x\ {\isacharquery}z{\isacharparenright}} \\
   \isa{nat{\isachardot}less{\isacharunderscore}total} from locale \isa{total{\isacharunderscore}order}: \\
-  \quad \isa{{\isacharquery}x\ {\isacharless}\ {\isacharquery}y\ {\isasymor}\ {\isacharquery}x\ {\isacharequal}\ {\isacharquery}y\ {\isasymor}\ {\isacharquery}y\ {\isacharless}\ {\isacharquery}x}
+  \quad \isa{partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharquery}x\ {\isacharquery}y\ {\isasymor}\ {\isacharquery}x\ {\isacharequal}\ {\isacharquery}y\ {\isasymor}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharquery}y\ {\isacharquery}x}
 \end{tabular}
 \end{center}
 \hrule
@@ -244,7 +244,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
@@ -260,7 +260,7 @@
 \ unfold{\isacharunderscore}locales\ {\isacharparenleft}auto\ simp{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}\isanewline
 \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{interpret}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}}\isamarkupfalse%
+\ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
@@ -285,7 +285,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{where}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\isanewline
 \ \ \ \ \ \ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ dvd\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
 \ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\isanewline
@@ -316,7 +316,7 @@
 \isanewline
 \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{interpret}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}}\isamarkupfalse%
+\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ dvd\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
@@ -390,7 +390,7 @@
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
 \ nat{\isacharunderscore}dvd{\isacharcolon}\isanewline
-\ \ distrib{\isacharunderscore}lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ \ distrib{\isacharunderscore}lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{apply}\isamarkupfalse%
 \ unfold{\isacharunderscore}locales%
 \begin{isamarkuptxt}%
@@ -434,7 +434,7 @@
   \isa{nat{\isacharunderscore}dvd{\isachardot}meet{\isacharunderscore}left} from locale \isa{lattice}: \\
   \quad \isa{gcd\ {\isacharquery}x\ {\isacharquery}y\ dvd\ {\isacharquery}x} \\
   \isa{nat{\isacharunderscore}dvd{\isachardot}join{\isacharunderscore}distr} from locale \isa{distrib{\isacharunderscore}lattice}: \\
-  \quad \isa{lcm\ {\isacharquery}x\ {\isacharparenleft}gcd\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}lcm\ {\isacharquery}x\ {\isacharquery}y{\isacharparenright}\ {\isacharparenleft}lcm\ {\isacharquery}x\ {\isacharquery}z{\isacharparenright}} \\
+  \quad \isa{lattice{\isachardot}join\ op\ dvd\ {\isacharquery}x\ {\isacharparenleft}lattice{\isachardot}meet\ op\ dvd\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ lattice{\isachardot}meet\ op\ dvd\ {\isacharparenleft}lattice{\isachardot}join\ op\ dvd\ {\isacharquery}x\ {\isacharquery}y{\isacharparenright}\ {\isacharparenleft}lattice{\isachardot}join\ op\ dvd\ {\isacharquery}x\ {\isacharquery}z{\isacharparenright}} \\
 \end{tabular}
 \end{center}
 \hrule
@@ -476,7 +476,7 @@
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
 \ order{\isacharunderscore}preserving\ {\isacharequal}\isanewline
-\ \ \ \ partial{\isacharunderscore}order\ {\isacharplus}\ partial{\isacharunderscore}order\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
+\ \ \ \ partial{\isacharunderscore}order\ {\isacharplus}\ po{\isacharprime}{\isacharcolon}\ partial{\isacharunderscore}order\ le{\isacharprime}\ \isakeyword{for}\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
 \ \ \ \ \isakeyword{fixes}\ {\isasymphi}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequoteclose}\isanewline
 \ \ \ \ \isakeyword{assumes}\ hom{\isacharunderscore}le{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ {\isasymphi}\ x\ {\isasympreceq}\ {\isasymphi}\ y{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
@@ -505,7 +505,8 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The locale \isa{order{\isacharunderscore}preserving} contains theorems for both
+% FIXME needs update
+  The locale \isa{order{\isacharunderscore}preserving} contains theorems for both
   orders \isa{{\isasymsqsubseteq}} and \isa{{\isasympreceq}}.  How can one refer to a theorem for
   a particular order, \isa{{\isasymsqsubseteq}} or \isa{{\isasympreceq}}?  Names in locales are
   qualified by the locale parameters.  More precisely, a name is
@@ -529,16 +530,17 @@
 \endisadeliminvisible
 %
 \begin{isamarkuptext}%
-\isa{le{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymsqsubseteq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}z}
+% FIXME needs update?
+  \isa{less{\isacharunderscore}le{\isacharunderscore}trans}: \isa{{\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}z}
 
-  \isa{le{\isacharunderscore}le{\isacharprime}{\isacharunderscore}{\isasymphi}{\isachardot}hom{\isacharunderscore}le}: \isa{{\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymLongrightarrow}\ {\isasymphi}\ {\isacharquery}x\ {\isasympreceq}\ {\isasymphi}\ {\isacharquery}y}%
+  \isa{hom{\isacharunderscore}le}: \isa{{\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymLongrightarrow}\ {\isasymphi}\ {\isacharquery}x\ {\isasympreceq}\ {\isasymphi}\ {\isacharquery}y}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
 When renaming a locale, the morphism is also applied
   to the qualifiers.  Hence theorems for the partial order \isa{{\isasympreceq}}
-  are qualified by \isa{le{\isacharprime}}.  For example, \isa{le{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}%
+  are qualified by \isa{le{\isacharprime}}.  For example, \isa{po{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}%
 \ \ {\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\isanewline
 \isaindent{\ \ }{\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}z%
 \end{isabelle}%
@@ -560,7 +562,8 @@
 \endisadeliminvisible
 %
 \begin{isamarkuptext}%
-This example reveals that there is no infix syntax for the strict
+% FIXME needs update?
+  This example reveals that there is no infix syntax for the strict
   version of \isa{{\isasympreceq}}!  This can, of course, not be introduced
   automatically, but it can be declared manually through an abbreviation.%
 \end{isamarkuptext}%
@@ -589,7 +592,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
-\ lattice{\isacharunderscore}hom\ {\isacharequal}\ lattice\ {\isacharplus}\ lattice\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
+\ lattice{\isacharunderscore}hom\ {\isacharequal}\ lattice\ {\isacharplus}\ lat{\isacharprime}{\isacharbang}{\isacharcolon}\ lattice\ le{\isacharprime}\ \isakeyword{for}\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
 \ \ \ \ \isakeyword{fixes}\ {\isasymphi}\isanewline
 \ \ \ \ \isakeyword{assumes}\ hom{\isacharunderscore}meet{\isacharcolon}\isanewline
 \ \ \ \ \ \ \ \ {\isachardoublequoteopen}{\isasymphi}\ {\isacharparenleft}lattice{\isachardot}meet\ le\ x\ y{\isacharparenright}\ {\isacharequal}\ lattice{\isachardot}meet\ le{\isacharprime}\ {\isacharparenleft}{\isasymphi}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
@@ -598,17 +601,17 @@
 \isanewline
 \ \ \isacommand{abbreviation}\isamarkupfalse%
 \ {\isacharparenleft}\isakeyword{in}\ lattice{\isacharunderscore}hom{\isacharparenright}\isanewline
-\ \ \ \ meet{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}meet{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}meet{\isachardoublequoteclose}\isanewline
+\ \ \ \ meet{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}meet{\isacharprime}\ {\isasymequiv}\ lat{\isacharprime}{\isachardot}meet{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{abbreviation}\isamarkupfalse%
 \ {\isacharparenleft}\isakeyword{in}\ lattice{\isacharunderscore}hom{\isacharparenright}\isanewline
-\ \ \ \ join{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}join{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}join{\isachardoublequoteclose}%
+\ \ \ \ join{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}join{\isacharprime}\ {\isasymequiv}\ lat{\isacharprime}{\isachardot}join{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
 A homomorphism is an endomorphism if both orders coincide.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
 \ lattice{\isacharunderscore}end\ {\isacharequal}\isanewline
-\ \ \ \ lattice{\isacharunderscore}hom\ le\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubseteq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ le\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubseteq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}%
+\ \ \ \ lattice{\isacharunderscore}hom\ le\ le\ \isakeyword{for}\ le\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubseteq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}%
 \begin{isamarkuptext}%
 The inheritance diagram of the situation we have now is shown
   in Figure~\ref{fig:hom}, where the dashed line depicts an
@@ -659,7 +662,7 @@
   interpretation is used to assert this.%
 \end{isamarkuptext}%
 \isamarkuptrue%
-\ \ \isacommand{interpretation}\isamarkupfalse%
+\ \ \isacommand{sublocale}\isamarkupfalse%
 \ lattice{\isacharunderscore}hom\ {\isasymsubseteq}\ order{\isacharunderscore}preserving%
 \isadelimproof
 \ %
@@ -675,7 +678,7 @@
 \ \ \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{have}\isamarkupfalse%
 \ {\isachardoublequoteopen}y\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}simp\ add{\isacharcolon}\ le{\isachardot}join{\isacharunderscore}connection{\isacharparenright}\isanewline
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ join{\isacharunderscore}connection{\isacharparenright}\isanewline
 \ \ \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{have}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isasymphi}\ y\ {\isacharequal}\ {\isacharparenleft}{\isasymphi}\ x\ {\isasymsqunion}{\isacharprime}\ {\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
@@ -683,7 +686,7 @@
 \ \ \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isasymphi}\ x\ {\isasympreceq}\ {\isasymphi}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}simp\ add{\isacharcolon}\ le{\isacharprime}{\isachardot}join{\isacharunderscore}connection{\isacharparenright}\isanewline
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ lat{\isacharprime}{\isachardot}join{\isacharunderscore}connection{\isacharparenright}\isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
 %
 \endisatagproof
@@ -697,7 +700,7 @@
 Theorems and other declarations --- syntax, in particular ---
   from the locale \isa{order{\isacharunderscore}preserving} are now active in \isa{lattice{\isacharunderscore}hom}, for example
 
-  \isa{le{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}:
+  \isa{lat{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}:
   \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymprec}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymprec}\ {\isacharquery}z}%
 \end{isamarkuptext}%
 \isamarkuptrue%
--- a/doc-src/TutorialI/Documents/document/Documents.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/TutorialI/Documents/document/Documents.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -391,7 +391,7 @@
   session is derived from a single parent, usually an object-logic
   image like \texttt{HOL}.  This results in an overall tree structure,
   which is reflected by the output location in the file system
-  (usually rooted at \verb,~/isabelle/browser_info,).
+  (usually rooted at \verb,~/.isabelle/browser_info,).
 
   \medskip The easiest way to manage Isabelle sessions is via
   \texttt{isabelle mkdir} (generates an initial session source setup)
--- a/doc-src/TutorialI/Types/document/Numbers.tex	Mon Dec 29 11:04:27 2008 -0800
+++ b/doc-src/TutorialI/Types/document/Numbers.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -3,7 +3,6 @@
 \def\isabellecontext{Numbers}%
 %
 \isadelimtheory
-\isanewline
 %
 \endisadelimtheory
 %
--- a/etc/isar-keywords.el	Mon Dec 29 11:04:27 2008 -0800
+++ b/etc/isar-keywords.el	Mon Jan 05 07:54:16 2009 -0800
@@ -46,6 +46,9 @@
     "chapter"
     "class"
     "class_deps"
+    "class_interpret"
+    "class_interpretation"
+    "class_locale"
     "classes"
     "classrel"
     "code_abort"
@@ -420,6 +423,7 @@
     "axiomatization"
     "axioms"
     "class"
+    "class_locale"
     "classes"
     "classrel"
     "code_abort"
@@ -503,6 +507,7 @@
 
 (defconst isar-keywords-theory-goal
   '("ax_specification"
+    "class_interpretation"
     "corollary"
     "cpodef"
     "function"
@@ -541,7 +546,8 @@
     "subsubsect"))
 
 (defconst isar-keywords-proof-goal
-  '("have"
+  '("class_interpret"
+    "have"
     "hence"
     "interpret"
     "invoke"))
--- a/lib/html/isabelle.css	Mon Dec 29 11:04:27 2008 -0800
+++ b/lib/html/isabelle.css	Mon Jan 05 07:54:16 2009 -0800
@@ -20,20 +20,20 @@
 
 /* inner and outer syntax markup */
 
-.tfree, tfree          { color: purple; }
-.tvar, tvar            { color: purple; }
-.free, free            { color: blue; }
-.skolem, skolem        { color: brown; }
-.bound, bound          { color: green; }
-.var, var              { color: blue; }
-.num, num              { }
-.xnum, xnum            { }
-.xstr, xstr            { color: brown; }
-.literal, literal      { font-weight: bold; }
-                      
+.tfree, tfree                 { color: purple; }
+.tvar, tvar                   { color: purple; }
+.free, free                   { color: blue; }
+.skolem, skolem               { color: brown; }
+.bound, bound                 { color: green; }
+.var, var                     { color: blue; }
+.numeral, numeral             { }
+.literal, literal             { font-weight: bold; }
+.inner_string, inner_string   { color: brown; }
+.inner_comment, inner_comment { color: #8B0000; }
+
 .loc, loc              { color: brown; }
 .tclass, tclass        { color: red; }
-          
+
 .keyword, keyword      { font-weight: bold; }
 .command, command      { font-weight: bold; }
 .ident, ident          { }
--- a/src/CCL/Wfd.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/CCL/Wfd.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      CCL/Wfd.thy
-    ID:         $Id$
     Author:     Martin Coen, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 *)
@@ -435,9 +434,9 @@
   | get_bno l n (Bound m) = (m-length(l),n)
 
 (* Not a great way of identifying induction hypothesis! *)
-fun could_IH x = could_unify(x,hd (prems_of @{thm rcallT})) orelse
-                 could_unify(x,hd (prems_of @{thm rcall2T})) orelse
-                 could_unify(x,hd (prems_of @{thm rcall3T}))
+fun could_IH x = Term.could_unify(x,hd (prems_of @{thm rcallT})) orelse
+                 Term.could_unify(x,hd (prems_of @{thm rcall2T})) orelse
+                 Term.could_unify(x,hd (prems_of @{thm rcall3T}))
 
 fun IHinst tac rls = SUBGOAL (fn (Bi,i) =>
   let val bvs = bvars Bi []
@@ -451,7 +450,7 @@
 
 fun is_rigid_prog t =
      case (Logic.strip_assums_concl t) of
-        (Const("Trueprop",_) $ (Const("mem",_) $ a $ _)) => (term_vars a = [])
+        (Const("Trueprop",_) $ (Const("mem",_) $ a $ _)) => null (Term.add_vars a [])
        | _ => false
 in
 
--- a/src/FOL/IsaMakefile	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/FOL/IsaMakefile	Mon Jan 05 07:54:16 2009 -0800
@@ -46,7 +46,7 @@
 FOL-ex: FOL $(LOG)/FOL-ex.gz
 
 $(LOG)/FOL-ex.gz: $(OUT)/FOL ex/First_Order_Logic.thy ex/If.thy		\
-  ex/IffOracle.thy ex/LocaleTest.thy ex/Nat.thy ex/Natural_Numbers.thy	\
+  ex/IffOracle.thy ex/Nat.thy ex/Natural_Numbers.thy	\
   ex/NewLocaleSetup.thy ex/NewLocaleTest.thy    \
   ex/Miniscope.thy ex/Prolog.thy ex/ROOT.ML ex/Classical.thy		\
   ex/document/root.tex ex/Foundation.thy ex/Intuitionistic.thy		\
--- a/src/FOL/ex/LocaleTest.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,765 +0,0 @@
-(*  Title:      FOL/ex/LocaleTest.thy
-    ID:         $Id$
-    Author:     Clemens Ballarin
-    Copyright (c) 2005 by Clemens Ballarin
-
-Collection of regression tests for locales.
-*)
-
-header {* Test of Locale Interpretation *}
-
-theory LocaleTest
-imports FOL
-begin
-
-ML {* set Toplevel.debug *}
-ML {* set show_hyps *}
-ML {* set show_sorts *}
-
-ML {*
-  fun check_thm name = let
-    val thy = the_context ();
-    val thm = PureThy.get_thm thy name;
-    val {prop, hyps, ...} = rep_thm thm;
-    val prems = Logic.strip_imp_prems prop;
-    val _ = if null hyps then ()
-        else error ("Theorem " ^ quote name ^ " has meta hyps.\n" ^
-          "Consistency check of locales package failed.");
-    val _ = if null prems then ()
-        else error ("Theorem " ^ quote name ^ " has premises.\n" ^
-          "Consistency check of locales package failed.");
-  in () end;
-*}
-
-section {* Context Elements and Locale Expressions *}
-
-text {* Naming convention for global objects: prefixes L and l *}
-
-subsection {* Renaming with Syntax *}
-
-locale LS = var mult +
-  assumes "mult(x, y) = mult(y, x)"
-
-print_locale LS
-
-locale LS' = LS mult (infixl "**" 60)
-
-print_locale LS'
-
-locale LT = var mult (infixl "**" 60) +
-  assumes "x ** y = y ** x"
-
-locale LU = LT mult (infixl "**" 60) + LT add (infixl "++" 55) + var h +
-  assumes hom: "h(x ** y) = h(x) ++ h(y)"
-
-
-(* FIXME: graceful handling of type errors?
-locale LY = LT mult (infixl "**" 60) + LT add (binder "++" 55) + var h +
-  assumes "mult(x) == add"
-*)
-
-
-locale LV = LU _ add
-
-locale LW = LU _ mult (infixl "**" 60)
-
-
-subsection {* Constrains *}
-
-locale LZ = fixes a (structure)
-locale LZ' = LZ +
-  constrains a :: "'a => 'b"
-  assumes "a (x :: 'a) = a (y)"
-print_locale LZ'
-
-
-section {* Interpretation *}
-
-text {* Naming convention for global objects: prefixes I and i *}
-
-text {* interpretation input syntax *}
-
-locale IL
-locale IM = fixes a and b and c
-
-interpretation test: IL + IM a b c [x y z] .
-
-print_interps IL    (* output: test *)
-print_interps IM    (* output: test *)
-
-interpretation test: IL print_interps IM .
-
-interpretation IL .
-
-text {* Processing of locale expression *}
-
-locale IA = fixes a assumes asm_A: "a = a"
-
-locale IB = fixes b assumes asm_B [simp]: "b = b"
-
-locale IC = IA + IB + assumes asm_C: "b = b"
-
-locale IC' = IA + IB + assumes asm_C: "c = c"
-  (* independent type var in c *)
-
-locale ID = IA + IB + fixes d defines def_D: "d == (a = b)"
-
-theorem (in ID) True ..
-
-typedecl i
-arities i :: "term"
-
-
-interpretation i1: IC ["X::i" "Y::i"] by unfold_locales auto
-
-print_interps IA  (* output: i1 *)
-
-(* possible accesses *)
-thm i1.a.asm_A thm LocaleTest.IA_locale.i1.a.asm_A thm IA_locale.i1.a.asm_A
-thm i1.asm_A thm LocaleTest.i1.asm_A
-
-ML {* check_thm "i1.a.asm_A" *}
-
-(* without prefix *)
-
-interpretation IC ["W::i" "Z::i"] by intro_locales  (* subsumed by i1: IC *)
-interpretation IC ["W::'a" "Z::i"] by unfold_locales auto
-  (* subsumes i1: IA and i1: IC *)
-
-print_interps IA  (* output: <no prefix>, i1 *)
-
-(* possible accesses *)
-thm asm_C thm a_b.asm_C thm LocaleTest.IC_locale.a_b.asm_C thm IC_locale.a_b.asm_C
-
-ML {* check_thm "asm_C" *}
-
-interpretation i2: ID [X "Y::i" "Y = X"]
-  by (simp add: eq_commute) unfold_locales
-
-print_interps IA  (* output: <no prefix>, i1 *)
-print_interps ID  (* output: i2 *)
-
-
-interpretation i3: ID [X "Y::i"] by simp unfold_locales
-
-(* duplicate: thm not added *)
-
-(* thm i3.a.asm_A *)
-
-
-print_interps IA  (* output: <no prefix>, i1 *)
-print_interps IB  (* output: i1 *)
-print_interps IC  (* output: <no prefix, i1 *)
-print_interps ID  (* output: i2, i3 *)
-
-
-interpretation i10: ID + ID a' b' d' [X "Y::i" _ u "v::i" _] by intro_locales
-
-corollary (in ID) th_x: True ..
-
-(* possible accesses: for each registration *)
-
-thm i2.th_x thm i3.th_x
-
-ML {* check_thm "i2.th_x"; check_thm "i3.th_x" *}
-
-lemma (in ID) th_y: "d == (a = b)" by (rule d_def)
-
-thm i2.th_y thm i3.th_y
-
-ML {* check_thm "i2.th_y"; check_thm "i3.th_y" *}
-
-lemmas (in ID) th_z = th_y
-
-thm i2.th_z
-
-ML {* check_thm "i2.th_z" *}
-
-
-subsection {* Interpretation in Proof Contexts *}
-
-locale IF = fixes f assumes asm_F: "f & f --> f"
-
-consts default :: "'a"
-
-theorem True
-proof -
-  fix alpha::i and beta
-  have alpha_A: "IA(alpha)" by unfold_locales
-  interpret i5: IA [alpha] by intro_locales  (* subsumed *)
-  print_interps IA  (* output: <no prefix>, i1 *)
-  interpret i6: IC [alpha beta] by unfold_locales auto
-  print_interps IA   (* output: <no prefix>, i1 *)
-  print_interps IC   (* output: <no prefix>, i1, i6 *)
-  interpret i11: IF ["default = default"] by (fast intro: IF.intro)
-  thm i11.asm_F [where 'a = i]     (* default has schematic type *)
-qed rule
-
-theorem (in IA) True
-proof -
-  print_interps! IA
-  fix beta and gamma
-  interpret i9: ID [a beta _]
-    apply - apply assumption
-    apply unfold_locales
-    apply (rule refl) done
-qed rule
-
-
-(* Definition involving free variable *)
-
-ML {* reset show_sorts *}
-
-locale IE = fixes e defines e_def: "e(x) == x & x"
-  notes e_def2 = e_def
-
-lemma (in IE) True thm e_def by fast
-
-interpretation i7: IE ["%x. x"] by simp
-
-thm i7.e_def2 (* has no premise *)
-
-ML {* check_thm "i7.e_def2" *}
-
-locale IE' = fixes e defines e_def: "e == (%x. x & x)"
-  notes e_def2 = e_def
-
-interpretation i7': IE' ["(%x. x)"] by simp
-
-thm i7'.e_def2
-
-ML {* check_thm "i7'.e_def2" *}
-
-(* Definition involving free variable in assm *)
-
-locale IG = fixes g assumes asm_G: "g --> x"
-  notes asm_G2 = asm_G
-
-interpretation i8: IG ["False"] by unfold_locales fast
-
-thm i8.asm_G2
-
-ML {* check_thm "i8.asm_G2" *}
-
-text {* Locale without assumptions *}
-
-locale IL1 = notes rev_conjI [intro] = conjI [THEN iffD1 [OF conj_commute]]
-
-lemma "[| P; Q |] ==> P & Q"
-proof -
-  interpret my: IL1 .          txt {* No chained fact required. *}
-  assume Q and P               txt {* order reversed *}
-  then show "P & Q" ..         txt {* Applies @{thm my.rev_conjI}. *}
-qed
-
-locale IL11 = notes rev_conjI = conjI [THEN iffD1 [OF conj_commute]]
-
-
-subsection {* Simple locale with assumptions *}
-
-consts ibin :: "[i, i] => i" (infixl "#" 60)
-
-axioms i_assoc: "(x # y) # z = x # (y # z)"
-  i_comm: "x # y = y # x"
-
-locale IL2 =
-  fixes OP (infixl "+" 60)
-  assumes assoc: "(x + y) + z = x + (y + z)"
-    and comm: "x + y = y + x"
-
-lemma (in IL2) lcomm: "x + (y + z) = y + (x + z)"
-proof -
-  have "x + (y + z) = (x + y) + z" by (simp add: assoc)
-  also have "... = (y + x) + z" by (simp add: comm)
-  also have "... = y + (x + z)" by (simp add: assoc)
-  finally show ?thesis .
-qed
-
-lemmas (in IL2) AC = comm assoc lcomm
-
-lemma "(x::i) # y # z # w = y # x # w # z"
-proof -
-  interpret my: IL2 ["op #"] by (rule IL2.intro [of "op #", OF i_assoc i_comm])
-  show ?thesis by (simp only: my.OP.AC)  (* or my.AC *)
-qed
-
-subsection {* Nested locale with assumptions *}
-
-locale IL3 =
-  fixes OP (infixl "+" 60)
-  assumes assoc: "(x + y) + z = x + (y + z)"
-
-locale IL4 = IL3 +
-  assumes comm: "x + y = y + x"
-
-lemma (in IL4) lcomm: "x + (y + z) = y + (x + z)"
-proof -
-  have "x + (y + z) = (x + y) + z" by (simp add: assoc)
-  also have "... = (y + x) + z" by (simp add: comm)
-  also have "... = y + (x + z)" by (simp add: assoc)
-  finally show ?thesis .
-qed
-
-lemmas (in IL4) AC = comm assoc lcomm
-
-lemma "(x::i) # y # z # w = y # x # w # z"
-proof -
-  interpret my: IL4 ["op #"]
-    by (auto intro: IL4.intro IL3.intro IL4_axioms.intro i_assoc i_comm)
-  show ?thesis by (simp only: my.OP.AC)  (* or simply AC *)
-qed
-
-text {* Locale with definition *}
-
-text {* This example is admittedly not very creative :-) *}
-
-locale IL5 = IL4 + var A +
-  defines A_def: "A == True"
-
-lemma (in IL5) lem: A
-  by (unfold A_def) rule
-
-lemma "IL5(op #) ==> True"
-proof -
-  assume "IL5(op #)"
-  then interpret IL5 ["op #"] by (auto intro: IL5.axioms)
-  show ?thesis by (rule lem)  (* lem instantiated to True *)
-qed
-
-text {* Interpretation in a context with target *}
-
-lemma (in IL4)
-  fixes A (infixl "$" 60)
-  assumes A: "IL4(A)"
-  shows "(x::i) $ y $ z $ w = y $ x $ w $ z"
-proof -
-  from A interpret A: IL4 ["A"] by (auto intro: IL4.axioms)
-  show ?thesis by (simp only: A.OP.AC)
-qed
-
-
-section {* Interpretation in Locales *}
-
-text {* Naming convention for global objects: prefixes R and r *}
-
-(* locale with many parameters ---
-   interpretations generate alternating group A5 *)
-
-locale RA5 = var A + var B + var C + var D + var E +
-  assumes eq: "A <-> B <-> C <-> D <-> E"
-
-interpretation RA5 < RA5 _ _ D E C
-print_facts
-print_interps RA5
-  using A_B_C_D_E.eq apply (blast intro: RA5.intro) done
-
-interpretation RA5 < RA5 C _ E _ A
-print_facts
-print_interps RA5
-  using A_B_C_D_E.eq apply (blast intro: RA5.intro) done
-
-interpretation RA5 < RA5 B C A _ _
-print_facts
-print_interps RA5
-  using A_B_C_D_E.eq apply (blast intro: RA5.intro) done
-
-interpretation RA5 < RA5 _ C D B _ .
-  (* Any even permutation of parameters is subsumed by the above. *)
-
-(* circle of three locales, forward direction *)
-
-locale RA1 = var A + var B + assumes p: "A <-> B"
-locale RA2 = var A + var B + assumes q: "A & B | ~ A & ~ B"
-locale RA3 = var A + var B + assumes r: "(A --> B) & (B --> A)"
-
-interpretation RA1 < RA2
-  print_facts
-  using p apply unfold_locales apply fast done
-interpretation RA2 < RA3
-  print_facts
-  using q apply unfold_locales apply fast done
-interpretation RA3 < RA1
-  print_facts
-  using r apply unfold_locales apply fast done
-
-(* circle of three locales, backward direction *)
-
-locale RB1 = var A + var B + assumes p: "A <-> B"
-locale RB2 = var A + var B + assumes q: "A & B | ~ A & ~ B"
-locale RB3 = var A + var B + assumes r: "(A --> B) & (B --> A)"
-
-interpretation RB1 < RB2
-  print_facts
-  using p apply unfold_locales apply fast done
-interpretation RB3 < RB1
-  print_facts
-  using r apply unfold_locales apply fast done
-interpretation RB2 < RB3
-  print_facts
-  using q apply unfold_locales apply fast done
-
-lemma (in RB1) True
-  print_facts
-  ..
-
-
-(* Group example *)
-
-locale Rsemi = var prod (infixl "**" 65) +
-  assumes assoc: "(x ** y) ** z = x ** (y ** z)"
-
-locale Rlgrp = Rsemi + var one + var inv +
-  assumes lone: "one ** x = x"
-    and linv: "inv(x) ** x = one"
-
-lemma (in Rlgrp) lcancel:
-  "x ** y = x ** z <-> y = z"
-proof
-  assume "x ** y = x ** z"
-  then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc)
-  then show "y = z" by (simp add: lone linv)
-qed simp
-
-locale Rrgrp = Rsemi + var one + var inv +
-  assumes rone: "x ** one = x"
-    and rinv: "x ** inv(x) = one"
-
-lemma (in Rrgrp) rcancel:
-  "y ** x = z ** x <-> y = z"
-proof
-  assume "y ** x = z ** x"
-  then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
-    by (simp add: assoc [symmetric])
-  then show "y = z" by (simp add: rone rinv)
-qed simp
-
-interpretation Rlgrp < Rrgrp
-  proof unfold_locales
-    {
-      fix x
-      have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone)
-      then show "x ** one = x" by (simp add: assoc lcancel)
-    }
-    note rone = this
-    {
-      fix x
-      have "inv(x) ** x ** inv(x) = inv(x) ** one"
-	by (simp add: linv lone rone)
-      then show "x ** inv(x) = one" by (simp add: assoc lcancel)
-    }
-  qed
-
-(* effect on printed locale *)
-
-print_locale! Rlgrp
-
-(* use of derived theorem *)
-
-lemma (in Rlgrp)
-  "y ** x = z ** x <-> y = z"
-  apply (rule rcancel)
-  print_interps Rrgrp thm lcancel rcancel
-  done
-
-(* circular interpretation *)
-
-interpretation Rrgrp < Rlgrp
-  proof unfold_locales
-    {
-      fix x
-      have "one ** (x ** inv(x)) = x ** inv(x)" by (simp add: rinv rone)
-      then show "one ** x = x" by (simp add: assoc [symmetric] rcancel)
-    }
-    note lone = this
-    {
-      fix x
-      have "inv(x) ** (x ** inv(x)) = one ** inv(x)"
-	by (simp add: rinv lone rone)
-      then show "inv(x) ** x = one" by (simp add: assoc [symmetric] rcancel)
-    }
-  qed
-
-(* effect on printed locale *)
-
-print_locale! Rrgrp
-print_locale! Rlgrp
-
-subsection {* Interaction of Interpretation in Theories and Locales:
-  in Locale, then in Theory *}
-
-consts
-  rone :: i
-  rinv :: "i => i"
-
-axioms
-  r_one : "rone # x = x"
-  r_inv : "rinv(x) # x = rone"
-
-interpretation Rbool: Rlgrp ["op #" "rone" "rinv"]
-proof
-  fix x y z
-  {
-    show "(x # y) # z = x # (y # z)" by (rule i_assoc)
-  next
-    show "rone # x = x" by (rule r_one)
-  next
-    show "rinv(x) # x = rone" by (rule r_inv)
-  }
-qed
-
-(* derived elements *)
-
-print_interps Rrgrp
-print_interps Rlgrp
-
-lemma "y # x = z # x <-> y = z" by (rule Rbool.rcancel)
-
-(* adding lemma to derived element *)
-
-lemma (in Rrgrp) new_cancel:
-  "b ** a = c ** a <-> b = c"
-  by (rule rcancel)
-
-thm Rbool.new_cancel (* additional prems discharged!! *)
-
-ML {* check_thm "Rbool.new_cancel" *}
-
-lemma "b # a = c # a <-> b = c" by (rule Rbool.new_cancel)
-
-
-subsection {* Interaction of Interpretation in Theories and Locales:
-  in Theory, then in Locale *}
-
-(* Another copy of the group example *)
-
-locale Rqsemi = var prod (infixl "**" 65) +
-  assumes assoc: "(x ** y) ** z = x ** (y ** z)"
-
-locale Rqlgrp = Rqsemi + var one + var inv +
-  assumes lone: "one ** x = x"
-    and linv: "inv(x) ** x = one"
-
-lemma (in Rqlgrp) lcancel:
-  "x ** y = x ** z <-> y = z"
-proof
-  assume "x ** y = x ** z"
-  then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc)
-  then show "y = z" by (simp add: lone linv)
-qed simp
-
-locale Rqrgrp = Rqsemi + var one + var inv +
-  assumes rone: "x ** one = x"
-    and rinv: "x ** inv(x) = one"
-
-lemma (in Rqrgrp) rcancel:
-  "y ** x = z ** x <-> y = z"
-proof
-  assume "y ** x = z ** x"
-  then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
-    by (simp add: assoc [symmetric])
-  then show "y = z" by (simp add: rone rinv)
-qed simp
-
-interpretation Rqrgrp < Rrgrp
-  apply unfold_locales
-  apply (rule assoc)
-  apply (rule rone)
-  apply (rule rinv)
-  done
-
-interpretation R2: Rqlgrp ["op #" "rone" "rinv"] 
-  apply unfold_locales  (* FIXME: unfold_locales is too eager and shouldn't
-                          solve this. *)
-  apply (rule i_assoc)
-  apply (rule r_one)
-  apply (rule r_inv)
-  done
-
-print_interps Rqsemi
-print_interps Rqlgrp
-print_interps Rlgrp  (* no interpretations yet *)
-
-
-interpretation Rqlgrp < Rqrgrp
-  proof unfold_locales
-    {
-      fix x
-      have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone)
-      then show "x ** one = x" by (simp add: assoc lcancel)
-    }
-    note rone = this
-    {
-      fix x
-      have "inv(x) ** x ** inv(x) = inv(x) ** one"
-	by (simp add: linv lone rone)
-      then show "x ** inv(x) = one" by (simp add: assoc lcancel)
-    }
-  qed
-
-print_interps! Rqrgrp
-print_interps! Rsemi  (* witness must not have meta hyps *)
-print_interps! Rrgrp  (* witness must not have meta hyps *)
-print_interps! Rlgrp  (* witness must not have meta hyps *)
-thm R2.rcancel
-thm R2.lcancel
-
-ML {* check_thm "R2.rcancel"; check_thm "R2.lcancel" *}
-
-
-subsection {* Generation of Witness Theorems for Transitive Interpretations *}
-
-locale Rtriv = var x +
-  assumes x: "x = x"
-
-locale Rtriv2 = var x + var y +
-  assumes x: "x = x" and y: "y = y"
-
-interpretation Rtriv2 < Rtriv x
-  apply unfold_locales
-  apply (rule x)
-  done
-
-interpretation Rtriv2 < Rtriv y
-  apply unfold_locales
-  apply (rule y)
-  done
-
-print_locale Rtriv2
-
-locale Rtriv3 = var x + var y + var z +
-  assumes x: "x = x" and y: "y = y" and z: "z = z"
-
-interpretation Rtriv3 < Rtriv2 x y
-  apply unfold_locales
-  apply (rule x)
-  apply (rule y)
-  done
-
-print_locale Rtriv3
-
-interpretation Rtriv3 < Rtriv2 x z
-  apply unfold_locales
-  apply (rule x_y_z.x)
-  apply (rule z)
-  done
-
-ML {* set show_types *}
-
-print_locale Rtriv3
-
-
-subsection {* Normalisation Replaces Assumed Element by Derived Element *}
-
-typedecl ('a, 'b) pair
-arities pair :: ("term", "term") "term"
-
-consts
-  pair :: "['a, 'b] => ('a, 'b) pair"
-  fst :: "('a, 'b) pair => 'a"
-  snd :: "('a, 'b) pair => 'b"
-
-axioms
-  fst [simp]: "fst(pair(x, y)) = x"
-  snd [simp]: "snd(pair(x, y)) = y"
-
-locale Rpair = var prod (infixl "**" 65) + var prodP (infixl "***" 65) +
-  defines P_def: "x *** y == pair(fst(x) ** fst(y), snd(x) ** snd(y))"
-
-locale Rpair_semi = Rpair + Rsemi
-
-interpretation Rpair_semi < Rsemi prodP (infixl "***" 65)
-proof unfold_locales
-  fix x y z
-  show "(x *** y) *** z = x *** (y *** z)"
-    apply (simp only: P_def) apply (simp add: assoc) (* FIXME: unfold P_def fails *)
-    done
-qed
-
-locale Rsemi_rev = Rsemi + var rprod (infixl "++" 65) +
-  defines r_def: "x ++ y == y ** x"
-
-lemma (in Rsemi_rev) r_assoc:
-  "(x ++ y) ++ z = x ++ (y ++ z)"
-  by (simp add: r_def assoc)
-
-
-subsection {* Import of Locales with Predicates as Interpretation *}
-
-locale Ra =
-  assumes Ra: "True"
-
-locale Rb = Ra +
-  assumes Rb: "True"
-
-locale Rc = Rb +
-  assumes Rc: "True"
-
-print_locale! Rc
-
-
-section {* Interpretation of Defined Concepts *}
-
-text {* Naming convention for global objects: prefixes D and d *}
-
-
-subsection {* Simple examples *}
-
-locale Da = fixes a :: o
-  assumes true: a
-
-text {* In the following examples, @{term "~ a"} is the defined concept. *}
-
-lemma (in Da) not_false: "~ a <-> False"
-  apply simp apply (rule true) done
-
-interpretation D1: Da ["True"]
-  where "~ True == False"
-  apply -
-  apply unfold_locales [1] apply rule
-  by simp
-
-thm D1.not_false
-lemma "False <-> False" apply (rule D1.not_false) done
-
-interpretation D2: Da ["x | ~ x"]
-  where "~ (x | ~ x) <-> ~ x & x"
-  apply -
-  apply unfold_locales [1] apply fast
-  by simp
-
-thm D2.not_false
-lemma "~ x & x <-> False" apply (rule D2.not_false) done
-
-print_interps! Da
-
-(* Subscriptions of interpretations *)
-
-lemma (in Da) not_false2: "~a <-> False"
-  apply simp apply (rule true) done
-
-thm D1.not_false2 D2.not_false2
-lemma "False <-> False" apply (rule D1.not_false2) done
-lemma "~x & x <-> False" apply (rule D2.not_false2) done
-
-(* Unfolding in attributes *)
-
-locale Db = Da +
-  fixes b :: o
-  assumes a_iff_b: "~a <-> b"
-
-lemmas (in Db) not_false_b = not_false [unfolded a_iff_b]
-
-interpretation D2: Db ["x | ~ x" "~ (x <-> x)"]
-  apply unfold_locales apply fast done
-
-thm D2.not_false_b
-lemma "~(x <-> x) <-> False" apply (rule D2.not_false_b) done
-
-(* Subscription and attributes *)
-
-lemmas (in Db) not_false_b2 = not_false [unfolded a_iff_b]
-
-thm D2.not_false_b2
-lemma "~(x <-> x) <-> False" apply (rule D2.not_false_b2) done
-
-end
--- a/src/FOL/ex/NewLocaleSetup.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/FOL/ex/NewLocaleSetup.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -44,9 +44,10 @@
 val _ =
   OuterSyntax.command "interpretation"
     "prove interpretation of locale expression in theory" K.thy_goal
-    (P.!!! SpecParse.locale_expression
-        >> (fn expr => Toplevel.print o
-            Toplevel.theory_to_proof (Expression.interpretation_cmd expr)));
+    (P.!!! SpecParse.locale_expression --
+      Scan.optional (P.$$$ "where" |-- P.and_list1 (SpecParse.opt_thm_name ":" -- P.prop)) []
+        >> (fn (expr, mixin) => Toplevel.print o
+            Toplevel.theory_to_proof (Expression.interpretation_cmd expr mixin)));
 
 val _ =
   OuterSyntax.command "interpret"
--- a/src/FOL/ex/NewLocaleTest.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/FOL/ex/NewLocaleTest.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -8,9 +8,7 @@
 imports NewLocaleSetup
 begin
 
-ML_val {* set new_locales *}
 ML_val {* set Toplevel.debug *}
-ML_val {* set show_hyps *}
 
 
 typedecl int arities int :: "term"
@@ -24,7 +22,7 @@
   int_minus: "(-x) + x = 0"
   int_minus2: "-(-x) = x"
 
-text {* Inference of parameter types *}
+section {* Inference of parameter types *}
 
 locale param1 = fixes p
 print_locale! param1
@@ -44,7 +42,7 @@
 print_locale! param4
 
 
-text {* Incremental type constraints *}
+subsection {* Incremental type constraints *}
 
 locale constraint1 =
   fixes  prod (infixl "**" 65)
@@ -58,7 +56,7 @@
 print_locale! constraint2
 
 
-text {* Inheritance *}
+section {* Inheritance *}
 
 locale semi =
   fixes prod (infixl "**" 65)
@@ -94,7 +92,7 @@
 print_locale! pert_hom' thm pert_hom'_def
 
 
-text {* Syntax declarations *}
+section {* Syntax declarations *}
 
 locale logic =
   fixes land (infixl "&&" 55)
@@ -112,14 +110,30 @@
 locale use_decl = logic + semi "op ||"
 print_locale! use_decl thm use_decl_def
 
+locale extra_type =
+  fixes a :: 'a
+    and P :: "'a => 'b => o"
+begin
 
-text {* Foundational versions of theorems *}
+definition test :: "'a => o" where
+  "test(x) <-> (ALL b. P(x, b))"
+
+end
+
+term extra_type.test thm extra_type.test_def
+
+interpretation var: extra_type "0" "%x y. x = 0" .
+
+thm var.test_def
+
+
+section {* Foundational versions of theorems *}
 
 thm logic.assoc
 thm logic.lor_def
 
 
-text {* Defines *}
+section {* Defines *}
 
 locale logic_def =
   fixes land (infixl "&&" 55)
@@ -149,7 +163,7 @@
 end
 
 
-text {* Notes *}
+section {* Notes *}
 
 (* A somewhat arcane homomorphism example *)
 
@@ -165,11 +179,21 @@
   assumes semi_homh: "semi_hom(prod, sum, h)"
   notes semi_hom_mult = semi_hom_mult [OF semi_homh]
 
+thm semi_hom_loc.semi_hom_mult
+(* unspecified, attribute not applied in backgroud theory !!! *)
+
 lemma (in semi_hom_loc) "h(prod(x, y)) = sum(h(x), h(y))"
   by (rule semi_hom_mult)
 
+(* Referring to facts from within a context specification *)
 
-text {* Theorem statements *}
+lemma
+  assumes x: "P <-> P"
+  notes y = x
+  shows True ..
+
+
+section {* Theorem statements *}
 
 lemma (in lgrp) lcancel:
   "x ** y = x ** z <-> y = z"
@@ -200,7 +224,7 @@
 print_locale! rgrp
 
 
-text {* Patterns *}
+subsection {* Patterns *}
 
 lemma (in rgrp)
   assumes "y ** x = z ** x" (is ?a)
@@ -218,7 +242,7 @@
 qed
 
 
-text {* Interpretation between locales: sublocales *}
+section {* Interpretation between locales: sublocales *}
 
 sublocale lgrp < right: rgrp
 print_facts
@@ -305,8 +329,7 @@
   done
 
 print_locale! order_with_def
-(* Note that decls come after theorems that make use of them.
-  Appears to be harmless at least in this example. *)
+(* Note that decls come after theorems that make use of them. *)
 
 
 (* locale with many parameters ---
@@ -359,7 +382,7 @@
 print_locale! trivial  (* No instance for y created (subsumed). *)
 
 
-text {* Sublocale, then interpretation in theory *}
+subsection {* Sublocale, then interpretation in theory *}
 
 interpretation int: lgrp "op +" "0" "minus"
 proof unfold_locales
@@ -374,7 +397,7 @@
   (* the latter comes through the sublocale relation *)
 
 
-text {* Interpretation in theory, then sublocale *}
+subsection {* Interpretation in theory, then sublocale *}
 
 interpretation (* fol: *) logic "op +" "minus"
 (* FIXME declaration of qualified names *)
@@ -386,10 +409,12 @@
   assumes assoc: "(x && y) && z = x && (y && z)"
     and notnot: "-- (-- x) = x"
 begin
-(* FIXME
+
+(* FIXME interpretation below fails
 definition lor (infixl "||" 50) where
   "x || y = --(-- x && -- y)"
 *)
+
 end
 
 sublocale logic < two: logic2
@@ -398,7 +423,48 @@
 thm two.assoc
 
 
-text {* Interpretation in proofs *}
+subsection {* Declarations and sublocale *}
+
+locale logic_a = logic
+locale logic_b = logic
+
+sublocale logic_a < logic_b
+  by unfold_locales
+
+
+subsection {* Equations *}
+
+locale logic_o =
+  fixes land (infixl "&&" 55)
+    and lnot ("-- _" [60] 60)
+  assumes assoc_o: "(x && y) && z <-> x && (y && z)"
+    and notnot_o: "-- (-- x) <-> x"
+begin
+
+definition lor_o (infixl "||" 50) where
+  "x || y <-> --(-- x && -- y)"
+
+end
+
+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+
+  show "logic_o.lor_o(op &, Not, x, y) <-> x | y"
+    by (unfold logic_o.lor_o_def [OF bool_logic_o]) fast
+qed
+
+thm x.lor_o_def bool_logic_o
+
+lemma lor_triv: "z <-> z" ..
+
+lemma (in logic_o) lor_triv: "x || y <-> x || y" by fast
+
+thm lor_triv [where z = True] (* Check strict prefix. *)
+  x.lor_triv
+
+
+subsection {* Interpretation in proofs *}
 
 lemma True
 proof
--- a/src/FOL/ex/ROOT.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/FOL/ex/ROOT.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -29,6 +29,5 @@
 ];
 
 (*regression test for locales -- sets several global flags!*)
-no_document use_thy "LocaleTest";
 no_document use_thy "NewLocaleTest";
 
--- a/src/FOLP/IFOLP.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/FOLP/IFOLP.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      FOLP/IFOLP.thy
-    ID:         $Id$
     Author:     Martin D Coen, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 *)
@@ -247,7 +246,7 @@
           and concl = discard_proof (Logic.strip_assums_concl prem)
       in
           if exists (fn hyp => hyp aconv concl) hyps
-          then case distinct (op =) (filter (fn hyp => could_unify (hyp, concl)) hyps) of
+          then case distinct (op =) (filter (fn hyp => Term.could_unify (hyp, concl)) hyps) of
                    [_] => assume_tac i
                  |  _  => no_tac
           else no_tac
@@ -340,6 +339,7 @@
   shows "?a : R"
   apply (insert assms(1) [unfolded ex1_def])
   apply (erule exE conjE | assumption | rule assms(1))+
+  apply (erule assms(2), assumption)
   done
 
 
--- a/src/FOLP/simp.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/FOLP/simp.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
-(*  Title:      FOLP/simp
-    ID:         $Id$
+(*  Title:      FOLP/simp.ML
     Author:     Tobias Nipkow
     Copyright   1993  University of Cambridge
 
@@ -156,21 +155,21 @@
 (*ccs contains the names of the constants possessing congruence rules*)
 fun add_hidden_vars ccs =
   let fun add_hvars tm hvars = case tm of
-              Abs(_,_,body) => add_term_vars(body,hvars)
+              Abs(_,_,body) => OldTerm.add_term_vars(body,hvars)
             | _$_ => let val (f,args) = strip_comb tm 
                      in case f of
                             Const(c,T) => 
                                 if member (op =) ccs c
                                 then fold_rev add_hvars args hvars
-                                else add_term_vars (tm, hvars)
-                          | _ => add_term_vars (tm, hvars)
+                                else OldTerm.add_term_vars (tm, hvars)
+                          | _ => OldTerm.add_term_vars (tm, hvars)
                      end
             | _ => hvars;
   in add_hvars end;
 
 fun add_new_asm_vars new_asms =
     let fun itf (tm, at) vars =
-                if at then vars else add_term_vars(tm,vars)
+                if at then vars else OldTerm.add_term_vars(tm,vars)
         fun add_list(tm,al,vars) = let val (_,tml) = strip_comb tm
                 in if length(tml)=length(al)
                    then fold_rev itf (tml ~~ al) vars
@@ -198,7 +197,7 @@
     val hvars = add_new_asm_vars new_asms (rhs,hvars)
     fun it_asms asm hvars =
         if atomic asm then add_new_asm_vars new_asms (asm,hvars)
-        else add_term_frees(asm,hvars)
+        else OldTerm.add_term_frees(asm,hvars)
     val hvars = fold_rev it_asms asms hvars
     val hvs = map (#1 o dest_Var) hvars
     val refl1_tac = refl_tac 1
@@ -542,7 +541,7 @@
 fun find_subst sg T =
 let fun find (thm::thms) =
         let val (Const(_,cT), va, vb) = dest_red(hd(prems_of thm));
-            val [P] = add_term_vars(concl_of thm,[]) \\ [va,vb]
+            val [P] = OldTerm.add_term_vars(concl_of thm,[]) \\ [va,vb]
             val eqT::_ = binder_types cT
         in if Sign.typ_instance sg (T,eqT) then SOME(thm,va,vb,P)
            else find thms
--- a/src/HOL/Algebra/AbelCoset.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/AbelCoset.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     HOL/Algebra/AbelCoset.thy
-  Id:        $Id$
   Author:    Stephan Hohe, TU Muenchen
 *)
 
@@ -52,7 +51,9 @@
   "a_kernel G H h \<equiv> kernel \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>
                               \<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr> h"
 
-locale abelian_group_hom = abelian_group G + abelian_group H + var h +
+locale abelian_group_hom = G: abelian_group G + H: abelian_group H
+    for G (structure) and H (structure) +
+  fixes h
   assumes a_group_hom: "group_hom (| carrier = carrier G, mult = add G, one = zero G |)
                                   (| carrier = carrier H, mult = add H, one = zero H |) h"
 
@@ -180,7 +181,8 @@
 
 subsubsection {* Subgroups *}
 
-locale additive_subgroup = var H + struct G +
+locale additive_subgroup =
+  fixes H and G (structure)
   assumes a_subgroup: "subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
 
 lemma (in additive_subgroup) is_additive_subgroup:
@@ -218,7 +220,7 @@
 
 text {* Every subgroup of an @{text "abelian_group"} is normal *}
 
-locale abelian_subgroup = additive_subgroup H G + abelian_group G +
+locale abelian_subgroup = additive_subgroup + abelian_group G +
   assumes a_normal: "normal H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
 
 lemma (in abelian_subgroup) is_abelian_subgroup:
@@ -230,7 +232,7 @@
       and a_comm: "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<oplus>\<^bsub>G\<^esub> y = y \<oplus>\<^bsub>G\<^esub> x"
   shows "abelian_subgroup H G"
 proof -
-  interpret normal ["H" "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"]
+  interpret normal "H" "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
   by (rule a_normal)
 
   show "abelian_subgroup H G"
@@ -243,9 +245,9 @@
       and a_subgroup: "subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
   shows "abelian_subgroup H G"
 proof -
-  interpret comm_group ["\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"]
+  interpret comm_group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
   by (rule a_comm_group)
-  interpret subgroup ["H" "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"]
+  interpret subgroup "H" "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
   by (rule a_subgroup)
 
   show "abelian_subgroup H G"
@@ -538,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
@@ -690,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/Congruence.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/Congruence.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:  Algebra/Congruence.thy
-  Id:     $Id$
   Author: Clemens Ballarin, started 3 January 2008
   Copyright: Clemens Ballarin
 *)
--- a/src/HOL/Algebra/Coset.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/Coset.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Algebra/Coset.thy
-    ID:         $Id$
     Author:     Florian Kammueller, with new proofs by L C Paulson, and
                 Stephan Hohe
 *)
@@ -114,7 +113,7 @@
       and repr:  "H #> x = H #> y"
   shows "y \<in> H #> x"
 proof -
-  interpret subgroup [H G] by fact
+  interpret subgroup H G by fact
   show ?thesis  apply (subst repr)
   apply (intro rcos_self)
    apply (rule ycarr)
@@ -129,7 +128,7 @@
     and a': "a' \<in> H #> a"
   shows "a' \<in> carrier G"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   from subset and acarr
   have "H #> a \<subseteq> carrier G" by (rule r_coset_subset_G)
   from this and a'
@@ -142,7 +141,7 @@
   assumes hH: "h \<in> H"
   shows "H #> h = H"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis apply (unfold r_coset_def)
     apply rule
     apply rule
@@ -173,7 +172,7 @@
       and x'cos: "x' \<in> H #> x"
   shows "(x' \<otimes> inv x) \<in> H"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   from xcarr x'cos
       have x'carr: "x' \<in> carrier G"
       by (rule elemrcos_carrier[OF is_group])
@@ -213,7 +212,7 @@
       and xixH: "(x' \<otimes> inv x) \<in> H"
   shows "x' \<in> H #> x"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   from xixH
       have "\<exists>h\<in>H. x' \<otimes> (inv x) = h" by fast
   from this
@@ -244,7 +243,7 @@
   assumes carr: "x \<in> carrier G" "x' \<in> carrier G"
   shows "(x' \<in> H #> x) = (x' \<otimes> inv x \<in> H)"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis proof  assume "x' \<in> H #> x"
     from this and carr
     show "x' \<otimes> inv x \<in> H"
@@ -263,7 +262,7 @@
   assumes XH: "X \<in> rcosets H"
   shows "X \<subseteq> carrier G"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   from XH have "\<exists>x\<in> carrier G. X = H #> x"
       unfolding RCOSETS_def
       by fast
@@ -348,7 +347,7 @@
       and xixH: "(inv x \<otimes> x') \<in> H"
   shows "x' \<in> x <# H"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   from xixH
       have "\<exists>h\<in>H. (inv x) \<otimes> x' = h" by fast
   from this
@@ -600,7 +599,7 @@
    assumes "group G"
    shows "equiv (carrier G) (rcong H)"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis
   proof (intro equiv.intro)
     show "refl (carrier G) (rcong H)"
@@ -647,7 +646,7 @@
   assumes a: "a \<in> carrier G"
   shows "a <# H = rcong H `` {a}"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a ) 
 qed
 
@@ -658,7 +657,7 @@
   assumes p: "ha \<otimes> a = h \<otimes> b" "a \<in> carrier G" "b \<in> carrier G" "h \<in> H" "ha \<in> H" "hb \<in> H"
   shows "hb \<otimes> a \<in> (\<Union>h\<in>H. {h \<otimes> b})"
 proof -
-  interpret subgroup [H G] by fact
+  interpret subgroup H G by fact
   from p show ?thesis apply (rule_tac UN_I [of "hb \<otimes> ((inv ha) \<otimes> h)"])
     apply (simp add: )
     apply (simp add: m_assoc transpose_inv)
@@ -670,7 +669,7 @@
   assumes p: "a \<in> rcosets H" "b \<in> rcosets H" "a\<noteq>b"
   shows "a \<inter> b = {}"
 proof -
-  interpret subgroup [H G] by fact
+  interpret subgroup H G by fact
   from p show ?thesis apply (simp add: RCOSETS_def r_coset_def)
     apply (blast intro: rcos_equation prems sym)
     done
@@ -760,7 +759,7 @@
   assumes "subgroup H G"
   shows "\<Union>(rcosets H) = carrier G"
 proof -
-  interpret subgroup [H G] by fact
+  interpret subgroup H G by fact
   show ?thesis
     apply (rule equalityI)
     apply (force simp add: RCOSETS_def r_coset_def)
@@ -847,7 +846,7 @@
   assumes "group G"
   shows "H \<in> rcosets H"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   from _ subgroup_axioms have "H #> \<one> = H"
     by (rule coset_join2) auto
   then show ?thesis
--- a/src/HOL/Algebra/Divisibility.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/Divisibility.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     Divisibility in monoids and rings
-  Id:        $Id$
   Author:    Clemens Ballarin, started 18 July 2008
 
 Based on work by Stephan Hohe.
@@ -32,7 +31,7 @@
   "monoid_cancel G"
   ..
 
-interpretation group \<subseteq> monoid_cancel
+sublocale group \<subseteq> monoid_cancel
   proof qed simp+
 
 
@@ -45,7 +44,7 @@
           "\<And>a b c. \<lbrakk>a \<otimes> c = b \<otimes> c; a \<in> carrier G; b \<in> carrier G; c \<in> carrier G\<rbrakk> \<Longrightarrow> a = b"
   shows "comm_monoid_cancel G"
 proof -
-  interpret comm_monoid [G] by fact
+  interpret comm_monoid G by fact
   show "comm_monoid_cancel G"
     apply unfold_locales
     apply (subgoal_tac "a \<otimes> c = b \<otimes> c")
@@ -59,7 +58,7 @@
   "comm_monoid_cancel G"
   by intro_locales
 
-interpretation comm_group \<subseteq> comm_monoid_cancel
+sublocale comm_group \<subseteq> comm_monoid_cancel
   ..
 
 
@@ -755,7 +754,7 @@
 using pf
 unfolding properfactor_lless
 proof -
-  interpret weak_partial_order ["division_rel G"] ..
+  interpret weak_partial_order "division_rel G" ..
   from x'x
        have "x' .=\<^bsub>division_rel G\<^esub> x" by simp
   also assume "x \<sqsubset>\<^bsub>division_rel G\<^esub> y"
@@ -771,7 +770,7 @@
 using pf
 unfolding properfactor_lless
 proof -
-  interpret weak_partial_order ["division_rel G"] ..
+  interpret weak_partial_order "division_rel G" ..
   assume "x \<sqsubset>\<^bsub>division_rel G\<^esub> y"
   also from yy'
        have "y .=\<^bsub>division_rel G\<^esub> y'" by simp
@@ -2916,7 +2915,7 @@
 lemma (in gcd_condition_monoid) division_weak_lower_semilattice [simp]:
   shows "weak_lower_semilattice (division_rel G)"
 proof -
-  interpret weak_partial_order ["division_rel G"] ..
+  interpret weak_partial_order "division_rel G" ..
   show ?thesis
   apply (unfold_locales, simp_all)
   proof -
@@ -2941,7 +2940,7 @@
   shows "a' gcdof b c"
 proof -
   note carr = a'carr carr'
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   have "a' \<in> carrier G \<and> a' gcdof b c"
     apply (simp add: gcdof_greatestLower carr')
     apply (subst greatest_Lower_cong_l[of _ a])
@@ -2958,7 +2957,7 @@
   assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
   shows "somegcd G a b \<in> carrier G"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (simp add: somegcd_meet[OF carr])
     apply (rule meet_closed[simplified], fact+)
@@ -2969,7 +2968,7 @@
   assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
   shows "(somegcd G a b) gcdof a b"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   from carr
   have "somegcd G a b \<in> carrier G \<and> (somegcd G a b) gcdof a b"
     apply (subst gcdof_greatestLower, simp, simp)
@@ -2983,7 +2982,7 @@
   assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
   shows "\<exists>x\<in>carrier G. x = somegcd G a b"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (simp add: somegcd_meet[OF carr])
     apply (rule meet_closed[simplified], fact+)
@@ -2994,7 +2993,7 @@
   assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
   shows "(somegcd G a b) divides a"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (simp add: somegcd_meet[OF carr])
     apply (rule meet_left[simplified], fact+)
@@ -3005,7 +3004,7 @@
   assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
   shows "(somegcd G a b) divides b"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (simp add: somegcd_meet[OF carr])
     apply (rule meet_right[simplified], fact+)
@@ -3017,7 +3016,7 @@
     and L: "x \<in> carrier G"  "y \<in> carrier G"  "z \<in> carrier G"
   shows "z divides (somegcd G x y)"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (simp add: somegcd_meet L)
     apply (rule meet_le[simplified], fact+)
@@ -3029,7 +3028,7 @@
     and carr: "x \<in> carrier G"  "x' \<in> carrier G"  "y \<in> carrier G"
   shows "somegcd G x y \<sim> somegcd G x' y"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (simp add: somegcd_meet carr)
     apply (rule meet_cong_l[simplified], fact+)
@@ -3041,7 +3040,7 @@
     and yy': "y \<sim> y'"
   shows "somegcd G x y \<sim> somegcd G x y'"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (simp add: somegcd_meet carr)
     apply (rule meet_cong_r[simplified], fact+)
@@ -3092,7 +3091,7 @@
   assumes "finite A"  "A \<subseteq> carrier G"  "A \<noteq> {}"
   shows "\<exists>x\<in> carrier G. x = SomeGcd G A"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (simp add: SomeGcd_def)
     apply (rule finite_inf_closed[simplified], fact+)
@@ -3103,7 +3102,7 @@
   assumes carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
   shows "somegcd G (somegcd G a b) c \<sim> somegcd G a (somegcd G b c)"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
   show ?thesis
     apply (subst (2 3) somegcd_meet, (simp add: carr)+)
     apply (simp add: somegcd_meet carr)
@@ -3313,7 +3312,7 @@
   qed
 qed
 
-interpretation gcd_condition_monoid \<subseteq> primeness_condition_monoid
+sublocale gcd_condition_monoid \<subseteq> primeness_condition_monoid
   by (rule primeness_condition)
 
 
@@ -3832,7 +3831,7 @@
   with fca fcb show ?thesis by simp
 qed
 
-interpretation factorial_monoid \<subseteq> divisor_chain_condition_monoid
+sublocale factorial_monoid \<subseteq> divisor_chain_condition_monoid
 apply unfold_locales
 apply (rule wfUNIVI)
 apply (rule measure_induct[of "factorcount G"])
@@ -3854,7 +3853,7 @@
   done
 qed
 
-interpretation factorial_monoid \<subseteq> primeness_condition_monoid
+sublocale factorial_monoid \<subseteq> primeness_condition_monoid
   proof qed (rule irreducible_is_prime)
 
 
@@ -3866,13 +3865,13 @@
   shows "gcd_condition_monoid G"
   proof qed (rule gcdof_exists)
 
-interpretation factorial_monoid \<subseteq> gcd_condition_monoid
+sublocale factorial_monoid \<subseteq> gcd_condition_monoid
   proof qed (rule gcdof_exists)
 
 lemma (in factorial_monoid) division_weak_lattice [simp]:
   shows "weak_lattice (division_rel G)"
 proof -
-  interpret weak_lower_semilattice ["division_rel G"] by simp
+  interpret weak_lower_semilattice "division_rel G" by simp
 
   show "weak_lattice (division_rel G)"
   apply (unfold_locales, simp_all)
@@ -3902,14 +3901,14 @@
 proof clarify
   assume dcc: "divisor_chain_condition_monoid G"
      and pc: "primeness_condition_monoid G"
-  interpret divisor_chain_condition_monoid ["G"] by (rule dcc)
-  interpret primeness_condition_monoid ["G"] by (rule pc)
+  interpret divisor_chain_condition_monoid "G" by (rule dcc)
+  interpret primeness_condition_monoid "G" by (rule pc)
 
   show "factorial_monoid G"
       by (fast intro: factorial_monoidI wfactors_exist wfactors_unique)
 next
   assume fm: "factorial_monoid G"
-  interpret factorial_monoid ["G"] by (rule fm)
+  interpret factorial_monoid "G" by (rule fm)
   show "divisor_chain_condition_monoid G \<and> primeness_condition_monoid G"
       by rule unfold_locales
 qed
@@ -3920,13 +3919,13 @@
 proof clarify
     assume dcc: "divisor_chain_condition_monoid G"
      and gc: "gcd_condition_monoid G"
-  interpret divisor_chain_condition_monoid ["G"] by (rule dcc)
-  interpret gcd_condition_monoid ["G"] by (rule gc)
+  interpret divisor_chain_condition_monoid "G" by (rule dcc)
+  interpret gcd_condition_monoid "G" by (rule gc)
   show "factorial_monoid G"
       by (simp add: factorial_condition_one[symmetric], rule, unfold_locales)
 next
   assume fm: "factorial_monoid G"
-  interpret factorial_monoid ["G"] by (rule fm)
+  interpret factorial_monoid "G" by (rule fm)
   show "divisor_chain_condition_monoid G \<and> gcd_condition_monoid G"
       by rule unfold_locales
 qed
--- a/src/HOL/Algebra/FiniteProduct.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/FiniteProduct.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Algebra/FiniteProduct.thy
-    ID:         $Id$
     Author:     Clemens Ballarin, started 19 November 2002
 
 This file is largely based on HOL/Finite_Set.thy.
@@ -519,9 +518,9 @@
 lemma finprod_singleton:
   assumes i_in_A: "i \<in> A" and fin_A: "finite A" and f_Pi: "f \<in> A \<rightarrow> carrier G"
   shows "(\<Otimes>j\<in>A. if i = j then f j else \<one>) = f i"
-  using i_in_A G.finprod_insert [of "A - {i}" i "(\<lambda>j. if i = j then f j else \<one>)"]
-    fin_A f_Pi G.finprod_one [of "A - {i}"]
-    G.finprod_cong [of "A - {i}" "A - {i}" "(\<lambda>j. if i = j then f j else \<one>)" "(\<lambda>i. \<one>)"] 
+  using i_in_A finprod_insert [of "A - {i}" i "(\<lambda>j. if i = j then f j else \<one>)"]
+    fin_A f_Pi finprod_one [of "A - {i}"]
+    finprod_cong [of "A - {i}" "A - {i}" "(\<lambda>j. if i = j then f j else \<one>)" "(\<lambda>i. \<one>)"] 
   unfolding Pi_def simp_implies_def by (force simp add: insert_absorb)
 
 end
--- a/src/HOL/Algebra/Group.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/Group.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:  HOL/Algebra/Group.thy
-  Id:     $Id$
   Author: Clemens Ballarin, started 4 February 2003
 
 Based on work by Florian Kammueller, L C Paulson and Markus Wenzel.
@@ -425,7 +424,7 @@
   assumes "group G"
   shows "group (G\<lparr>carrier := H\<rparr>)"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis
     apply (rule monoid.group_l_invI)
     apply (unfold_locales) [1]
@@ -489,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
@@ -501,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)
@@ -526,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
@@ -542,15 +541,6 @@
     {h. h \<in> carrier G -> carrier H &
       (\<forall>x \<in> carrier G. \<forall>y \<in> carrier G. h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y)}"
 
-lemma hom_mult:
-  "[| h \<in> hom G H; x \<in> carrier G; y \<in> carrier G |]
-   ==> h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y"
-  by (simp add: hom_def)
-
-lemma hom_closed:
-  "[| h \<in> hom G H; x \<in> carrier G |] ==> h x \<in> carrier H"
-  by (auto simp add: hom_def funcset_mem)
-
 lemma (in group) hom_compose:
      "[|h \<in> hom G H; i \<in> hom H I|] ==> compose (carrier G) i h \<in> hom G I"
 apply (auto simp add: hom_def funcset_compose) 
@@ -587,10 +577,23 @@
 
 text{*Basis for homomorphism proofs: we assume two groups @{term G} and
   @{term H}, with a homomorphism @{term h} between them*}
-locale group_hom = group G + group H + var h +
+locale group_hom = G: group G + H: group H for G (structure) and H (structure) +
+  fixes h
   assumes homh: "h \<in> hom G H"
-  notes hom_mult [simp] = hom_mult [OF homh]
-    and hom_closed [simp] = hom_closed [OF homh]
+
+lemma (in group_hom) hom_mult [simp]:
+  "[| x \<in> carrier G; y \<in> carrier G |] ==> h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y"
+proof -
+  assume "x \<in> carrier G" "y \<in> carrier G"
+  with homh [unfolded hom_def] show ?thesis by simp
+qed
+
+lemma (in group_hom) hom_closed [simp]:
+  "x \<in> carrier G ==> h x \<in> carrier H"
+proof -
+  assume "x \<in> carrier G"
+  with homh [unfolded hom_def] show ?thesis by (auto simp add: funcset_mem)
+qed
 
 lemma (in group_hom) one_closed [simp]:
   "h \<one> \<in> carrier H"
--- a/src/HOL/Algebra/Ideal.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/Ideal.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     HOL/Algebra/CIdeal.thy
-  Id:        $Id$
   Author:    Stephan Hohe, TU Muenchen
 *)
 
@@ -14,11 +13,11 @@
 
 subsubsection {* General definition *}
 
-locale ideal = additive_subgroup I R + ring R +
+locale ideal = additive_subgroup I R + ring R for I and R (structure) +
   assumes I_l_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
       and I_r_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
 
-interpretation ideal \<subseteq> abelian_subgroup I R
+sublocale ideal \<subseteq> abelian_subgroup I R
 apply (intro abelian_subgroupI3 abelian_group.intro)
   apply (rule ideal.axioms, rule ideal_axioms)
  apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms)
@@ -37,7 +36,7 @@
       and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
   shows "ideal I R"
 proof -
-  interpret ring [R] by fact
+  interpret ring R by fact
   show ?thesis  apply (intro ideal.intro ideal_axioms.intro additive_subgroupI)
      apply (rule a_subgroup)
     apply (rule is_ring)
@@ -68,7 +67,7 @@
   assumes generate: "\<exists>i \<in> carrier R. I = Idl {i}"
   shows "principalideal I R"
 proof -
-  interpret ideal [I R] by fact
+  interpret ideal I R by fact
   show ?thesis  by (intro principalideal.intro principalideal_axioms.intro) (rule is_ideal, rule generate)
 qed
 
@@ -89,7 +88,7 @@
      and I_maximal: "\<And>J. \<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R"
   shows "maximalideal I R"
 proof -
-  interpret ideal [I R] by fact
+  interpret ideal I R by fact
   show ?thesis by (intro maximalideal.intro maximalideal_axioms.intro)
     (rule is_ideal, rule I_notcarr, rule I_maximal)
 qed
@@ -112,8 +111,8 @@
       and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
   shows "primeideal I R"
 proof -
-  interpret ideal [I R] by fact
-  interpret cring [R] by fact
+  interpret ideal I R by fact
+  interpret cring R by fact
   show ?thesis by (intro primeideal.intro primeideal_axioms.intro)
     (rule is_ideal, rule is_cring, rule I_notcarr, rule I_prime)
 qed
@@ -128,8 +127,8 @@
       and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
   shows "primeideal I R"
 proof -
-  interpret additive_subgroup [I R] by fact
-  interpret cring [R] by fact
+  interpret additive_subgroup I R by fact
+  interpret cring R by fact
   show ?thesis apply (intro_locales)
     apply (intro ideal_axioms.intro)
     apply (erule (1) I_l_closed)
@@ -207,8 +206,8 @@
   assumes "ideal J R"
   shows "ideal (I \<inter> J) R"
 proof -
-  interpret ideal [I R] by fact
-  interpret ideal [J R] by fact
+  interpret ideal I R by fact
+  interpret ideal J R by fact
   show ?thesis
 apply (intro idealI subgroup.intro)
       apply (rule is_ring)
@@ -245,7 +244,7 @@
   from notempty have "\<exists>I0. I0 \<in> S" by blast
   from this obtain I0 where I0S: "I0 \<in> S" by auto
 
-  interpret ideal ["I0" "R"] by (rule Sideals[OF I0S])
+  interpret ideal I0 R by (rule Sideals[OF I0S])
 
   from xI[OF I0S] have "x \<in> I0" .
   from this and a_subset show "x \<in> carrier R" by fast
@@ -258,13 +257,13 @@
 
   fix J
   assume JS: "J \<in> S"
-  interpret ideal ["J" "R"] by (rule Sideals[OF JS])
+  interpret ideal J R by (rule Sideals[OF JS])
   from xI[OF JS] and yI[OF JS]
       show "x \<oplus> y \<in> J" by (rule a_closed)
 next
   fix J
   assume JS: "J \<in> S"
-  interpret ideal ["J" "R"] by (rule Sideals[OF JS])
+  interpret ideal J R by (rule Sideals[OF JS])
   show "\<zero> \<in> J" by simp
 next
   fix x
@@ -273,7 +272,7 @@
 
   fix J
   assume JS: "J \<in> S"
-  interpret ideal ["J" "R"] by (rule Sideals[OF JS])
+  interpret ideal J R by (rule Sideals[OF JS])
 
   from xI[OF JS]
       show "\<ominus> x \<in> J" by (rule a_inv_closed)
@@ -285,7 +284,7 @@
 
   fix J
   assume JS: "J \<in> S"
-  interpret ideal ["J" "R"] by (rule Sideals[OF JS])
+  interpret ideal J R by (rule Sideals[OF JS])
 
   from xI[OF JS] and ycarr
       show "y \<otimes> x \<in> J" by (rule I_l_closed)
@@ -297,7 +296,7 @@
 
   fix J
   assume JS: "J \<in> S"
-  interpret ideal ["J" "R"] by (rule Sideals[OF JS])
+  interpret ideal J R by (rule Sideals[OF JS])
 
   from xI[OF JS] and ycarr
       show "x \<otimes> y \<in> J" by (rule I_r_closed)
@@ -443,7 +442,7 @@
 lemma (in ring) genideal_one:
   "Idl {\<one>} = carrier R"
 proof -
-  interpret ideal ["Idl {\<one>}" "R"] by (rule genideal_ideal, fast intro: one_closed)
+  interpret ideal "Idl {\<one>}" "R" by (rule genideal_ideal, fast intro: one_closed)
   show "Idl {\<one>} = carrier R"
   apply (rule, rule a_subset)
   apply (simp add: one_imp_carrier genideal_self')
@@ -533,7 +532,7 @@
   assumes aJ: "a \<in> J"
   shows "PIdl a \<subseteq> J"
 proof -
-  interpret ideal [J R] by fact
+  interpret ideal J R by fact
   show ?thesis unfolding cgenideal_def
     apply rule
     apply clarify
@@ -580,7 +579,7 @@
   shows "Idl (I \<union> J) = I <+> J"
 apply rule
  apply (rule ring.genideal_minimal)
-   apply (rule R.is_ring)
+   apply (rule is_ring)
   apply (rule add_ideals[OF idealI idealJ])
  apply (rule)
  apply (simp add: set_add_defs) apply (elim disjE) defer 1 defer 1
@@ -660,7 +659,7 @@
   assumes "cring R"
   shows "\<exists>x\<in>I. I = carrier R #> x"
 proof -
-  interpret cring [R] by fact
+  interpret cring R by fact
   from generate
       obtain i
         where icarr: "i \<in> carrier R"
@@ -693,7 +692,7 @@
   assumes notprime: "\<not> primeideal I R"
   shows "carrier R = I \<or> (\<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 (rule ccontr, clarsimp)
-  interpret cring [R] by fact
+  interpret cring R by fact
   assume InR: "carrier R \<noteq> I"
      and "\<forall>a. a \<in> carrier R \<longrightarrow> (\<forall>b. a \<otimes> b \<in> I \<longrightarrow> b \<in> carrier R \<longrightarrow> a \<in> I \<or> b \<in> I)"
   hence I_prime: "\<And> a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I" by simp
@@ -713,7 +712,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
@@ -726,13 +725,13 @@
 apply (rule domain.intro, rule is_cring)
 apply (rule domain_axioms.intro)
 proof (rule ccontr, simp)
-  interpret primeideal ["{\<zero>}" "R"] by (rule pi)
+  interpret primeideal "{\<zero>}" "R" by (rule pi)
   assume "\<one> = \<zero>"
   hence "carrier R = {\<zero>}" by (rule one_zeroD)
   from this[symmetric] and I_notcarr
       show "False" by simp
 next
-  interpret primeideal ["{\<zero>}" "R"] by (rule pi)
+  interpret primeideal "{\<zero>}" "R" by (rule pi)
   fix a b
   assume ab: "a \<otimes> b = \<zero>"
      and carr: "a \<in> carrier R" "b \<in> carrier R"
@@ -771,7 +770,7 @@
   assumes acarr: "a \<in> carrier R"
   shows "ideal {x\<in>carrier R. a \<otimes> x \<in> I} R"
 proof -
-  interpret cring [R] by fact
+  interpret cring R by fact
   show ?thesis apply (rule idealI)
     apply (rule cring.axioms[OF is_cring])
     apply (rule subgroup.intro)
@@ -812,7 +811,7 @@
   assumes "maximalideal I R"
   shows "primeideal I R"
 proof -
-  interpret maximalideal [I R] by fact
+  interpret maximalideal I R by fact
   show ?thesis apply (rule ccontr)
     apply (rule primeidealCE)
     apply (rule is_cring)
@@ -830,7 +829,7 @@
       by fast
     def J \<equiv> "{x\<in>carrier R. a \<otimes> x \<in> I}"
     
-    from R.is_cring and acarr
+    from is_cring and acarr
     have idealJ: "ideal J R" unfolding J_def by (rule helper_max_prime)
     
     have IsubJ: "I \<subseteq> J"
@@ -855,7 +854,7 @@
     have "\<one> \<notin> J" unfolding J_def by fast
     hence Jncarr: "J \<noteq> carrier R" by fast
     
-    interpret ideal ["J" "R"] by (rule idealJ)
+    interpret ideal J R by (rule idealJ)
     
     have "J = I \<or> J = carrier R"
       apply (intro I_maximal)
@@ -932,7 +931,7 @@
   fix I
   assume a: "I \<in> {I. ideal I R}"
   with this
-      interpret ideal ["I" "R"] by simp
+      interpret ideal I R by simp
 
   show "I \<in> {{\<zero>}, carrier R}"
   proof (cases "\<exists>a. a \<in> I - {\<zero>}")
@@ -1019,7 +1018,7 @@
   fix J
   assume Jn0: "J \<noteq> {\<zero>}"
      and idealJ: "ideal J R"
-  interpret ideal ["J" "R"] by (rule idealJ)
+  interpret ideal J R by (rule idealJ)
   have "{\<zero>} \<subseteq> J" by (rule ccontr, simp)
   from zeromax and idealJ and this and a_subset
       have "J = {\<zero>} \<or> J = carrier R" by (rule maximalideal.I_maximal)
--- a/src/HOL/Algebra/IntRing.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/IntRing.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     HOL/Algebra/IntRing.thy
-  Id:        $Id$
   Author:    Stephan Hohe, TU Muenchen
 *)
 
@@ -97,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"
@@ -105,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)
@@ -117,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) }
@@ -140,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) }
@@ -165,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 -
@@ -175,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) }
@@ -188,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)
 
 
@@ -204,8 +203,8 @@
   "(True ==> PROP R) == PROP R"
   by simp_all
 
-interpretation int (* FIXME [unfolded UNIV] *) :
-  partial_order ["(| carrier = UNIV::int set, eq = op =, le = op \<le> |)"]
+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)"
     and "lless (| carrier = UNIV::int set, eq = op =, le = op \<le> |) x y = (x < y)"
@@ -220,8 +219,8 @@
     by (simp add: lless_def) auto
 qed
 
-interpretation int (* FIXME [unfolded UNIV] *) :
-  lattice ["(| carrier = UNIV::int set, eq = op =, le = op \<le> |)"]
+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"
 proof -
@@ -233,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)
@@ -246,8 +245,8 @@
     done
 qed
 
-interpretation int (* [unfolded UNIV] *) :
-  total_order ["(| carrier = UNIV::int set, eq = op =, le = op \<le> |)"]
+interpretation int! (* [unfolded UNIV] *) :
+  total_order "(| carrier = UNIV::int set, eq = op =, le = op \<le> |)"
   proof qed clarsimp
 
 
@@ -329,7 +328,7 @@
 next
   assume "UNIV = {uu. EX x. uu = x * p}"
   from this obtain x 
-      where "1 = x * p" by fast
+      where "1 = x * p" by best
   from this [symmetric]
       have "p * x = 1" by (subst zmult_commute)
   hence "\<bar>p * x\<bar> = 1" by simp
@@ -404,7 +403,7 @@
   assumes zmods: "ZMod m a = ZMod m b"
   shows "a mod m = b mod m"
 proof -
-  interpret ideal ["Idl\<^bsub>\<Z>\<^esub> {m}" \<Z>] by (rule int.genideal_ideal, fast)
+  interpret ideal "Idl\<^bsub>\<Z>\<^esub> {m}" \<Z> by (rule int.genideal_ideal, fast)
   from zmods
       have "b \<in> ZMod m a"
       unfolding ZMod_def
@@ -428,7 +427,7 @@
 lemma ZMod_mod:
   shows "ZMod m a = ZMod m (a mod m)"
 proof -
-  interpret ideal ["Idl\<^bsub>\<Z>\<^esub> {m}" \<Z>] by (rule int.genideal_ideal, fast)
+  interpret ideal "Idl\<^bsub>\<Z>\<^esub> {m}" \<Z> by (rule int.genideal_ideal, fast)
   show ?thesis
       unfolding ZMod_def
   apply (rule a_repr_independence'[symmetric])
--- a/src/HOL/Algebra/Lattice.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/Lattice.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     HOL/Algebra/Lattice.thy
-  Id:        $Id$
   Author:    Clemens Ballarin, started 7 November 2003
   Copyright: Clemens Ballarin
 
@@ -16,7 +15,7 @@
 record 'a gorder = "'a eq_object" +
   le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>\<index>" 50)
 
-locale weak_partial_order = equivalence L +
+locale weak_partial_order = equivalence L for L (structure) +
   assumes le_refl [intro, simp]:
       "x \<in> carrier L ==> x \<sqsubseteq> x"
     and weak_le_anti_sym [intro]:
@@ -920,7 +919,7 @@
 
 text {* Total orders are lattices. *}
 
-interpretation weak_total_order < weak_lattice
+sublocale weak_total_order < weak: weak_lattice
 proof
   fix x y
   assume L: "x \<in> carrier L"  "y \<in> carrier L"
@@ -1126,14 +1125,14 @@
   assumes sup_of_two_exists:
     "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. least L s (Upper L {x, y})"
 
-interpretation upper_semilattice < weak_upper_semilattice
+sublocale upper_semilattice < weak: weak_upper_semilattice
   proof qed (rule sup_of_two_exists)
 
 locale lower_semilattice = partial_order +
   assumes inf_of_two_exists:
     "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. greatest L s (Lower L {x, y})"
 
-interpretation lower_semilattice < weak_lower_semilattice
+sublocale lower_semilattice < weak: weak_lower_semilattice
   proof qed (rule inf_of_two_exists)
 
 locale lattice = upper_semilattice + lower_semilattice
@@ -1184,7 +1183,7 @@
 locale total_order = partial_order +
   assumes total_order_total: "[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
 
-interpretation total_order < weak_total_order
+sublocale total_order < weak: weak_total_order
   proof qed (rule total_order_total)
 
 text {* Introduction rule: the usual definition of total order *}
@@ -1196,7 +1195,7 @@
 
 text {* Total orders are lattices. *}
 
-interpretation total_order < lattice
+sublocale total_order < weak: lattice
   proof qed (auto intro: sup_of_two_exists inf_of_two_exists)
 
 
@@ -1208,7 +1207,7 @@
     and inf_exists:
     "[| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
 
-interpretation complete_lattice < weak_complete_lattice
+sublocale complete_lattice < weak: weak_complete_lattice
   proof qed (auto intro: sup_exists inf_exists)
 
 text {* Introduction rule: the usual definition of complete lattice *}
--- a/src/HOL/Algebra/Module.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/Module.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Algebra/Module.thy
-    ID:         $Id$
     Author:     Clemens Ballarin, started 15 April 2003
     Copyright:  Clemens Ballarin
 *)
@@ -14,7 +13,7 @@
 record ('a, 'b) module = "'b ring" +
   smult :: "['a, 'b] => 'b" (infixl "\<odot>\<index>" 70)
 
-locale module = cring R + abelian_group M +
+locale module = R: cring + M: abelian_group M for M (structure) +
   assumes smult_closed [simp, intro]:
       "[| a \<in> carrier R; x \<in> carrier M |] ==> a \<odot>\<^bsub>M\<^esub> x \<in> carrier M"
     and smult_l_distr:
@@ -29,7 +28,7 @@
     and smult_one [simp]:
       "x \<in> carrier M ==> \<one> \<odot>\<^bsub>M\<^esub> x = x"
 
-locale algebra = module R M + cring M +
+locale algebra = module + cring M +
   assumes smult_assoc2:
       "[| a \<in> carrier R; x \<in> carrier M; y \<in> carrier M |] ==>
       (a \<odot>\<^bsub>M\<^esub> x) \<otimes>\<^bsub>M\<^esub> y = a \<odot>\<^bsub>M\<^esub> (x \<otimes>\<^bsub>M\<^esub> y)"
--- a/src/HOL/Algebra/QuotRing.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/QuotRing.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     HOL/Algebra/QuotRing.thy
-  Id:        $Id$
   Author:    Stephan Hohe
 *)
 
@@ -158,7 +157,7 @@
   assumes "cring R"
   shows "cring (R Quot I)"
 proof -
-  interpret cring [R] by fact
+  interpret cring R by fact
   show ?thesis apply (intro cring.intro comm_monoid.intro comm_monoid_axioms.intro)
   apply (rule quotient_is_ring)
  apply (rule ring.axioms[OF quotient_is_ring])
@@ -173,12 +172,12 @@
   assumes "cring R"
   shows "ring_hom_cring R (R Quot I) (op +> I)"
 proof -
-  interpret cring [R] by fact
+  interpret cring R by fact
   show ?thesis apply (rule ring_hom_cringI)
   apply (rule rcos_ring_hom_ring)
- apply (rule R.is_cring)
+ apply (rule is_cring)
 apply (rule quotient_is_cring)
-apply (rule R.is_cring)
+apply (rule is_cring)
 done
 qed
 
@@ -244,7 +243,7 @@
   assumes "cring R"
   shows "field (R Quot I)"
 proof -
-  interpret cring [R] by fact
+  interpret cring R by fact
   show ?thesis apply (intro cring.cring_fieldI2)
   apply (rule quotient_is_cring, rule is_cring)
  defer 1
--- a/src/HOL/Algebra/Ring.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/Ring.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     The algebraic hierarchy of rings
-  Id:        $Id$
   Author:    Clemens Ballarin, started 9 December 1996
   Copyright: Clemens Ballarin
 *)
@@ -187,7 +186,7 @@
   assumes cg: "comm_group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
   shows "abelian_group G"
 proof -
-  interpret comm_group ["\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"]
+  interpret comm_group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
     by (rule cg)
   show "abelian_group G" ..
 qed
@@ -360,7 +359,7 @@
 
 subsection {* Rings: Basic Definitions *}
 
-locale ring = abelian_group R + monoid R +
+locale ring = abelian_group R + monoid R for R (structure) +
   assumes l_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
       ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
     and r_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
@@ -585,8 +584,8 @@
   assumes RS: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier S" "d \<in> carrier S"
   shows "a \<oplus> \<ominus> (a \<oplus> \<ominus> b) = b & c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c"
 proof -
-  interpret ring [R] by fact
-  interpret cring [S] by fact
+  interpret ring R by fact
+  interpret cring S by fact
 ML_val {* Algebra.print_structures @{context} *}
   from RS show ?thesis by algebra
 qed
@@ -597,7 +596,7 @@
   assumes R: "a \<in> carrier R" "b \<in> carrier R"
   shows "a \<ominus> (a \<ominus> b) = b"
 proof -
-  interpret ring [R] by fact
+  interpret ring R by fact
   from R show ?thesis by algebra
 qed
 
@@ -771,7 +770,8 @@
   shows "h \<in> ring_hom R S ==> h \<one> = \<one>\<^bsub>S\<^esub>"
   by (simp add: ring_hom_def)
 
-locale ring_hom_cring = cring R + cring S +
+locale ring_hom_cring = R: cring R + S: cring S
+    for R (structure) and S (structure) +
   fixes h
   assumes homh [simp, intro]: "h \<in> ring_hom R S"
   notes hom_closed [simp, intro] = ring_hom_closed [OF homh]
--- a/src/HOL/Algebra/RingHom.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/RingHom.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     HOL/Algebra/RingHom.thy
-  Id:        $Id$
   Author:    Stephan Hohe, TU Muenchen
 *)
 
@@ -11,15 +10,17 @@
 section {* Homomorphisms of Non-Commutative Rings *}
 
 text {* Lifting existing lemmas in a @{text ring_hom_ring} locale *}
-locale ring_hom_ring = ring R + ring S + var h +
+locale ring_hom_ring = R: ring R + S: ring S
+    for R (structure) and S (structure) +
+  fixes h
   assumes homh: "h \<in> ring_hom R S"
   notes hom_mult [simp] = ring_hom_mult [OF homh]
     and hom_one [simp] = ring_hom_one [OF homh]
 
-interpretation ring_hom_cring \<subseteq> ring_hom_ring
+sublocale ring_hom_cring \<subseteq> ring: ring_hom_ring
   proof qed (rule homh)
 
-interpretation ring_hom_ring \<subseteq> abelian_group_hom R S
+sublocale ring_hom_ring \<subseteq> abelian_group: abelian_group_hom R S
 apply (rule abelian_group_homI)
   apply (rule R.is_abelian_group)
  apply (rule S.is_abelian_group)
@@ -44,8 +45,8 @@
       and compatible_one: "h \<one> = \<one>\<^bsub>S\<^esub>"
   shows "ring_hom_ring R S h"
 proof -
-  interpret ring [R] by fact
-  interpret ring [S] by fact
+  interpret ring R by fact
+  interpret ring S by fact
   show ?thesis apply unfold_locales
 apply (unfold ring_hom_def, safe)
    apply (simp add: hom_closed Pi_def)
@@ -60,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)
@@ -76,9 +77,9 @@
       and compatible_one: "h \<one> = \<one>\<^bsub>S\<^esub>"
   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 abelian_group_hom R S h 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)
@@ -92,9 +93,9 @@
   assumes "ring_hom_ring R S h" "cring R" "cring S"
   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 ring_hom_ring R S h 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
@@ -124,7 +125,7 @@
       and xrcos: "x \<in> a_kernel R S h +> a"
   shows "h x = h a"
 proof -
-  interpret ideal ["a_kernel R S h" "R"] by (rule kernel_is_ideal)
+  interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal)
 
   from xrcos
       have "\<exists>i \<in> a_kernel R S h. x = i \<oplus> a" by (simp add: a_r_coset_defs)
@@ -152,7 +153,7 @@
       and hx: "h x = h a"
   shows "x \<in> a_kernel R S h +> a"
 proof -
-  interpret ideal ["a_kernel R S h" "R"] by (rule kernel_is_ideal)
+  interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal)
  
   note carr = acarr xcarr
   note hcarr = acarr[THEN hom_closed] xcarr[THEN hom_closed]
@@ -180,7 +181,7 @@
 apply rule defer 1
 apply clarsimp defer 1
 proof
-  interpret ideal ["a_kernel R S h" "R"] by (rule kernel_is_ideal)
+  interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal)
 
   fix x
   assume xrcos: "x \<in> a_kernel R S h +> a"
@@ -193,7 +194,7 @@
   from xcarr and this
       show "x \<in> {x \<in> carrier R. h x = h a}" by fast
 next
-  interpret ideal ["a_kernel R S h" "R"] by (rule kernel_is_ideal)
+  interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal)
 
   fix x
   assume xcarr: "x \<in> carrier R"
--- a/src/HOL/Algebra/UnivPoly.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/UnivPoly.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*
   Title:     HOL/Algebra/UnivPoly.thy
-  Id:        $Id$
   Author:    Clemens Ballarin, started 9 December 1996
   Copyright: Clemens Ballarin
 
@@ -176,17 +175,17 @@
   fixes R (structure) and P (structure)
   defines P_def: "P == UP R"
 
-locale UP_ring = UP + ring R
+locale UP_ring = UP + R: ring R
 
-locale UP_cring = UP + cring R
+locale UP_cring = UP + R: cring R
 
-interpretation UP_cring < UP_ring
-  by (rule P_def) intro_locales
+sublocale UP_cring < UP_ring
+  by intro_locales [1] (rule P_def)
 
-locale UP_domain = UP + "domain" R
+locale UP_domain = UP + R: "domain" R
 
-interpretation UP_domain < UP_cring
-  by (rule P_def) intro_locales
+sublocale UP_domain < UP_cring
+  by intro_locales [1] (rule P_def)
 
 context UP
 begin
@@ -458,8 +457,8 @@
 
 end
 
-interpretation UP_ring < ring P using UP_ring .
-interpretation UP_cring < cring P using UP_cring .
+sublocale UP_ring < P: ring P using UP_ring .
+sublocale UP_cring < P: cring P using UP_cring .
 
 
 subsection {* Polynomials Form an Algebra *}
@@ -508,7 +507,7 @@
   "algebra R P" by (auto intro!: algebraI R.cring UP_cring UP_smult_l_distr UP_smult_r_distr
     UP_smult_assoc1 UP_smult_assoc2)
 
-interpretation UP_cring < algebra R P using UP_algebra .
+sublocale UP_cring < algebra R P using UP_algebra .
 
 
 subsection {* Further Lemmas Involving Monomials *}
@@ -1085,7 +1084,7 @@
   Interpretation of theorems from @{term domain}.
 *}
 
-interpretation UP_domain < "domain" P
+sublocale UP_domain < "domain" P
   by intro_locales (rule domain.axioms UP_domain)+
 
 
@@ -1204,7 +1203,9 @@
 
 text {* The universal property of the polynomial ring *}
 
-locale UP_pre_univ_prop = ring_hom_cring R S h + UP_cring R P
+locale UP_pre_univ_prop = ring_hom_cring + UP_cring
+
+(* FIXME print_locale ring_hom_cring fails *)
 
 locale UP_univ_prop = UP_pre_univ_prop +
   fixes s and Eval
@@ -1350,7 +1351,7 @@
 
 text {* Interpretation of ring homomorphism lemmas. *}
 
-interpretation UP_univ_prop < ring_hom_cring P S Eval
+sublocale UP_univ_prop < ring_hom_cring P S Eval
   apply (unfold Eval_def)
   apply intro_locales
   apply (rule ring_hom_cring.axioms)
@@ -1391,7 +1392,7 @@
   assumes R: "r \<in> carrier R" and S: "s \<in> carrier S"
   shows "eval R S h s (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
 proof -
-  interpret UP_univ_prop [R S h P s _]
+  interpret UP_univ_prop R S h P s "eval R S h s"
     using UP_pre_univ_prop_axioms P_def R S
     by (auto intro: UP_univ_prop.intro UP_univ_prop_axioms.intro)
   from R
@@ -1428,8 +1429,8 @@
     and P: "p \<in> carrier P" and S: "s \<in> carrier S"
   shows "Phi p = Psi p"
 proof -
-  interpret ring_hom_cring [P S Phi] by fact
-  interpret ring_hom_cring [P S Psi] by fact
+  interpret ring_hom_cring P S Phi by fact
+  interpret ring_hom_cring P S Psi by fact
   have "Phi p =
       Phi (\<Oplus>\<^bsub>P \<^esub>i \<in> {..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 (^)\<^bsub>P\<^esub> i)"
     by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)
@@ -1772,17 +1773,17 @@
   shows "eval R R id a (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = \<zero>"
   (is "eval R R id a ?g = _")
 proof -
-  interpret UP_pre_univ_prop [R R id P] proof qed simp
+  interpret UP_pre_univ_prop R R id proof qed simp
   have eval_ring_hom: "eval R R id a \<in> ring_hom P R" using eval_ring_hom [OF a] by simp
-  interpret ring_hom_cring [P R "eval R R id a"] proof qed (simp add: eval_ring_hom)
+  interpret ring_hom_cring P R "eval R R id a" proof qed (simp add: eval_ring_hom)
   have mon1_closed: "monom P \<one>\<^bsub>R\<^esub> 1 \<in> carrier P" 
     and mon0_closed: "monom P a 0 \<in> carrier P" 
     and min_mon0_closed: "\<ominus>\<^bsub>P\<^esub> monom P a 0 \<in> carrier P"
     using a R.a_inv_closed by auto
   have "eval R R id a ?g = eval R R id a (monom P \<one> 1) \<ominus> eval R R id a (monom P a 0)"
     unfolding P.minus_eq [OF mon1_closed mon0_closed]
-    unfolding R_S_h.hom_add [OF mon1_closed min_mon0_closed]
-    unfolding R_S_h.hom_a_inv [OF mon0_closed] 
+    unfolding hom_add [OF mon1_closed min_mon0_closed]
+    unfolding hom_a_inv [OF mon0_closed] 
     using R.minus_eq [symmetric] mon1_closed mon0_closed by auto
   also have "\<dots> = a \<ominus> a"
     using eval_monom [OF R.one_closed a, of 1] using eval_monom [OF a a, of 0] using a by simp
@@ -1819,7 +1820,7 @@
     and deg_r_0: "deg R r = 0"
     shows "r = monom P (eval R R id a f) 0"
 proof -
-  interpret UP_pre_univ_prop [R R id P] proof qed simp
+  interpret UP_pre_univ_prop R R id P proof qed simp
   have eval_ring_hom: "eval R R id a \<in> ring_hom P R"
     using eval_ring_hom [OF a] by simp
   have "eval R R id a f = eval R R id a ?gq \<oplus>\<^bsub>R\<^esub> eval R R id a r"
@@ -1885,7 +1886,7 @@
   "UP INTEG"} globally.
 *}
 
-interpretation INTEG: UP_pre_univ_prop [INTEG INTEG id]
+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/abstract/Ring2.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/abstract/Ring2.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,8 +1,7 @@
-(*
-  Title:     The algebraic hierarchy of rings as axiomatic classes
-  Id:        $Id$
-  Author:    Clemens Ballarin, started 9 December 1996
-  Copyright: Clemens Ballarin
+(*  Title:     HOL/Algebra/abstract/Ring2.thy
+    Author:    Clemens Ballarin
+
+The algebraic hierarchy of rings as axiomatic classes.
 *)
 
 header {* The algebraic hierarchy of rings as type classes *}
@@ -211,7 +210,7 @@
         @{const_name HOL.minus}, @{const_name HOL.one}, @{const_name HOL.times}]
   | ring_ord _ = ~1;
 
-fun termless_ring (a, b) = (Term.term_lpo ring_ord (a, b) = LESS);
+fun termless_ring (a, b) = (TermOrd.term_lpo ring_ord (a, b) = LESS);
 
 val ring_ss = HOL_basic_ss settermless termless_ring addsimps
   [thm "a_assoc", thm "l_zero", thm "l_neg", thm "a_comm", thm "m_assoc",
--- a/src/HOL/Algebra/ringsimp.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Algebra/ringsimp.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,4 @@
-(*
-  Id:        $Id$
-  Author:    Clemens Ballarin
+(*  Author:    Clemens Ballarin
 
 Normalisation method for locales ring and cring.
 *)
@@ -48,7 +46,7 @@
     val ops = map (fst o Term.strip_comb) ts;
     fun ord (Const (a, _)) = find_index (fn (Const (b, _)) => a=b | _ => false) ops
       | ord (Free (a, _)) = find_index (fn (Free (b, _)) => a=b | _ => false) ops;
-    fun less (a, b) = (Term.term_lpo ord (a, b) = LESS);
+    fun less (a, b) = (TermOrd.term_lpo ord (a, b) = LESS);
   in asm_full_simp_tac (HOL_ss settermless less addsimps simps) end;
 
 fun algebra_tac ctxt =
--- a/src/HOL/Bali/AxExample.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Bali/AxExample.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -41,7 +41,7 @@
 
 ML {*
 fun inst1_tac ctxt s t st =
-  case AList.lookup (op =) (rev (Term.add_varnames (Thm.prop_of st) [])) s of
+  case AList.lookup (op =) (rev (Term.add_var_names (Thm.prop_of st) [])) s of
   SOME i => instantiate_tac ctxt [((s, i), t)] st | NONE => Seq.empty;
 
 val ax_tac =
--- a/src/HOL/Complex.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Complex.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -345,13 +345,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)
@@ -513,7 +513,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/Complex/Fundamental_Theorem_Algebra.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1329 +0,0 @@
-(*  Title:       Fundamental_Theorem_Algebra.thy
-    Author:      Amine Chaieb
-*)
-
-header{*Fundamental Theorem of Algebra*}
-
-theory Fundamental_Theorem_Algebra
-imports "~~/src/HOL/Univ_Poly" "~~/src/HOL/Library/Dense_Linear_Order" "~~/src/HOL/Complex"
-begin
-
-subsection {* Square root of complex numbers *}
-definition csqrt :: "complex \<Rightarrow> complex" where
-"csqrt z = (if Im z = 0 then
-            if 0 \<le> Re z then Complex (sqrt(Re z)) 0
-            else Complex 0 (sqrt(- Re z))
-           else Complex (sqrt((cmod z + Re z) /2))
-                        ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
-
-lemma csqrt[algebra]: "csqrt z ^ 2 = z"
-proof-
-  obtain x y where xy: "z = Complex x y" by (cases z, simp_all)
-  {assume y0: "y = 0"
-    {assume x0: "x \<ge> 0" 
-      then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
-	by (simp add: csqrt_def power2_eq_square)}
-    moreover
-    {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
-      then have ?thesis using y0 xy real_sqrt_pow2[OF x0] 
-	by (simp add: csqrt_def power2_eq_square) }
-    ultimately have ?thesis by blast}
-  moreover
-  {assume y0: "y\<noteq>0"
-    {fix x y
-      let ?z = "Complex x y"
-      from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
-      hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+ 
-      hence "(sqrt (x * x + y * y) + x) / 2 \<ge> 0" "(sqrt (x * x + y * y) - x) / 2 \<ge> 0" by (simp_all add: power2_eq_square) }
-    note th = this
-    have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2" 
-      by (simp add: power2_eq_square) 
-    from th[of x y]
-    have sq4': "sqrt (((sqrt (x * x + y * y) + x)^2 / 4)) = (sqrt (x * x + y * y) + x) / 2" "sqrt (((sqrt (x * x + y * y) - x)^2 / 4)) = (sqrt (x * x + y * y) - x) / 2" unfolding sq4 by simp_all
-    then have th1: "sqrt ((sqrt (x * x + y * y) + x) * (sqrt (x * x + y * y) + x) / 4) - sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) - x) / 4) = x"
-      unfolding power2_eq_square by simp 
-    have "sqrt 4 = sqrt (2^2)" by simp 
-    hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
-    have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
-      using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
-      unfolding power2_eq_square 
-      by (simp add: ring_simps real_sqrt_divide sqrt4)
-     from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
-       apply (simp add: real_sqrt_sum_squares_mult_ge_zero[of x y] real_sqrt_pow2[OF th(1)[of x y], unfolded power2_eq_square] real_sqrt_pow2[OF th(2)[of x y], unfolded power2_eq_square] real_sqrt_mult[symmetric])
-      using th1 th2  ..}
-  ultimately show ?thesis by blast
-qed
-
-
-subsection{* More lemmas about module of complex numbers *}
-
-lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
-  by (rule of_real_power [symmetric])
-
-lemma real_down2: "(0::real) < d1 \<Longrightarrow> 0 < d2 ==> EX e. 0 < e & e < d1 & e < d2"
-  apply ferrack apply arith done
-
-text{* The triangle inequality for cmod *}
-lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
-  using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
-
-subsection{* Basic lemmas about complex polynomials *}
-
-lemma poly_bound_exists:
-  shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
-proof(induct p)
-  case Nil thus ?case by (rule exI[where x=1], simp) 
-next
-  case (Cons c cs)
-  from Cons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
-    by blast
-  let ?k = " 1 + cmod c + \<bar>r * m\<bar>"
-  have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
-  {fix z
-    assume H: "cmod z \<le> r"
-    from m H have th: "cmod (poly cs z) \<le> m" by blast
-    from H have rp: "r \<ge> 0" using norm_ge_zero[of z] by arith
-    have "cmod (poly (c # cs) z) \<le> cmod c + cmod (z* poly cs z)"
-      using norm_triangle_ineq[of c "z* poly cs z"] by simp
-    also have "\<dots> \<le> cmod c + r*m" using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]] by (simp add: norm_mult)
-    also have "\<dots> \<le> ?k" by simp
-    finally have "cmod (poly (c # cs) z) \<le> ?k" .}
-  with kp show ?case by blast
-qed
-
-
-text{* Offsetting the variable in a polynomial gives another of same degree *}
-  (* FIXME : Lemma holds also in locale --- fix it later *)
-lemma  poly_offset_lemma:
-  shows "\<exists>b q. (length q = length p) \<and> (\<forall>x. poly (b#q) (x::complex) = (a + x) * poly p x)"
-proof(induct p)
-  case Nil thus ?case by simp
-next
-  case (Cons c cs)
-  from Cons.hyps obtain b q where 
-    bq: "length q = length cs" "\<forall>x. poly (b # q) x = (a + x) * poly cs x"
-    by blast
-  let ?b = "a*c"
-  let ?q = "(b+c)#q"
-  have lg: "length ?q = length (c#cs)" using bq(1) by simp
-  {fix x
-    from bq(2)[rule_format, of x]
-    have "x*poly (b # q) x = x*((a + x) * poly cs x)" by simp
-    hence "poly (?b# ?q) x = (a + x) * poly (c # cs) x"
-      by (simp add: ring_simps)}
-  with lg  show ?case by blast 
-qed
-
-    (* FIXME : This one too*)
-lemma poly_offset: "\<exists> q. length q = length p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
-proof (induct p)
-  case Nil thus ?case by simp
-next
-  case (Cons c cs)
-  from Cons.hyps obtain q where q: "length q = length cs" "\<forall>x. poly q x = poly cs (a + x)" by blast
-  from poly_offset_lemma[of q a] obtain b p where 
-    bp: "length p = length q" "\<forall>x. poly (b # p) x = (a + x) * poly q x"
-    by blast
-  thus ?case using q bp by - (rule exI[where x="(c + b)#p"], simp)
-qed
-
-text{* An alternative useful formulation of completeness of the reals *}
-lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
-  shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
-proof-
-  from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
-  from ex have thx:"\<exists>x. x \<in> Collect P" by blast
-  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y" 
-    by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
-  from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
-    by blast
-  from Y[OF x] have xY: "x < Y" .
-  from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)  
-  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y" 
-    apply (clarsimp, atomize (full)) by auto 
-  from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
-  {fix y
-    {fix z assume z: "P z" "y < z"
-      from L' z have "y < L" by auto }
-    moreover
-    {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
-      hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
-      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def) 
-      with yL(1) have False  by arith}
-    ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
-  thus ?thesis by blast
-qed
-
-
-subsection{* Some theorems about Sequences*}
-text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
-
-lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
-  unfolding Ex1_def
-  apply (rule_tac x="nat_rec e f" in exI)
-  apply (rule conjI)+
-apply (rule def_nat_rec_0, simp)
-apply (rule allI, rule def_nat_rec_Suc, simp)
-apply (rule allI, rule impI, rule ext)
-apply (erule conjE)
-apply (induct_tac x)
-apply (simp add: nat_rec_0)
-apply (erule_tac x="n" in allE)
-apply (simp)
-done
-
- text{* An equivalent formulation of monotony -- Not used here, but might be useful *}
-lemma mono_Suc: "mono f = (\<forall>n. (f n :: 'a :: order) \<le> f (Suc n))"
-unfolding mono_def
-proof auto
-  fix A B :: nat
-  assume H: "\<forall>n. f n \<le> f (Suc n)" "A \<le> B"
-  hence "\<exists>k. B = A + k" apply -  apply (thin_tac "\<forall>n. f n \<le> f (Suc n)") 
-    by presburger
-  then obtain k where k: "B = A + k" by blast
-  {fix a k
-    have "f a \<le> f (a + k)"
-    proof (induct k)
-      case 0 thus ?case by simp
-    next
-      case (Suc k)
-      from Suc.hyps H(1)[rule_format, of "a + k"] show ?case by simp
-    qed}
-  with k show "f A \<le> f B" by blast
-qed
-
-text{* for any sequence, there is a mootonic subsequence *}
-lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
-proof-
-  {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
-    let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
-    from num_Axiom[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
-    obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
-    have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
-      using H apply - 
-      apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
-      unfolding order_le_less by blast 
-    hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
-    {fix n
-      have "?P (f (Suc n)) (f n)" 
-	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
-	using H apply - 
-      apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
-      unfolding order_le_less by blast 
-    hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
-  note fSuc = this
-    {fix p q assume pq: "p \<ge> f q"
-      have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
-	by (cases q, simp_all) }
-    note pqth = this
-    {fix q
-      have "f (Suc q) > f q" apply (induct q) 
-	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
-    note fss = this
-    from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
-    {fix a b 
-      have "f a \<le> f (a + b)"
-      proof(induct b)
-	case 0 thus ?case by simp
-      next
-	case (Suc b)
-	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
-      qed}
-    note fmon0 = this
-    have "monoseq (\<lambda>n. s (f n))" 
-    proof-
-      {fix n
-	have "s (f n) \<ge> s (f (Suc n))" 
-	proof(cases n)
-	  case 0
-	  assume n0: "n = 0"
-	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
-	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
-	next
-	  case (Suc m)
-	  assume m: "n = Suc m"
-	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
-	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
-	qed}
-      thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
-    qed
-    with th1 have ?thesis by blast}
-  moreover
-  {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
-    {fix p assume p: "p \<ge> Suc N" 
-      hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
-      have "m \<noteq> p" using m(2) by auto 
-      with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
-    note th0 = this
-    let ?P = "\<lambda>m x. m > x \<and> s x < s m"
-    from num_Axiom[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
-    obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
-      "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
-    have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
-      using N apply - 
-      apply (erule allE[where x="Suc N"], clarsimp)
-      apply (rule_tac x="m" in exI)
-      apply auto
-      apply (subgoal_tac "Suc N \<noteq> m")
-      apply simp
-      apply (rule ccontr, simp)
-      done
-    hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
-    {fix n
-      have "f n > N \<and> ?P (f (Suc n)) (f n)"
-	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
-      proof (induct n)
-	case 0 thus ?case
-	  using f0 N apply auto 
-	  apply (erule allE[where x="f 0"], clarsimp) 
-	  apply (rule_tac x="m" in exI, simp)
-	  by (subgoal_tac "f 0 \<noteq> m", auto)
-      next
-	case (Suc n)
-	from Suc.hyps have Nfn: "N < f n" by blast
-	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
-	with Nfn have mN: "m > N" by arith
-	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
-	
-	from key have th0: "f (Suc n) > N" by simp
-	from N[rule_format, OF th0]
-	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
-	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
-	hence "m' > f (Suc n)" using m'(1) by simp
-	with key m'(2) show ?case by auto
-      qed}
-    note fSuc = this
-    {fix n
-      have "f n \<ge> Suc N \<and> f(Suc n) > f n \<and> s(f n) < s(f(Suc n))" using fSuc[of n] by auto 
-      hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
-    note thf = this
-    have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
-    have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
-      apply -
-      apply (rule disjI1)
-      apply auto
-      apply (rule order_less_imp_le)
-      apply blast
-      done
-    then have ?thesis  using sqf by blast}
-  ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
-qed
-
-lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
-proof(induct n)
-  case 0 thus ?case by simp
-next
-  case (Suc n)
-  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
-  have "n < f (Suc n)" by arith 
-  thus ?case by arith
-qed
-
-subsection {* Fundamental theorem of algebra *}
-lemma  unimodular_reduce_norm:
-  assumes md: "cmod z = 1"
-  shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
-proof-
-  obtain x y where z: "z = Complex x y " by (cases z, auto)
-  from md z have xy: "x^2 + y^2 = 1" by (simp add: cmod_def)
-  {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
-    from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
-      by (simp_all add: cmod_def power2_eq_square ring_simps)
-    hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
-    hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
-      by - (rule power_mono, simp, simp)+
-    hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1" 
-      by (simp_all  add: power2_abs power_mult_distrib)
-    from add_mono[OF th0] xy have False by simp }
-  thus ?thesis unfolding linorder_not_le[symmetric] by blast
-qed
-
-text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
-lemma reduce_poly_simple:
- assumes b: "b \<noteq> 0" and n: "n\<noteq>0"
-  shows "\<exists>z. cmod (1 + b * z^n) < 1"
-using n
-proof(induct n rule: nat_less_induct)
-  fix n
-  assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" and n: "n \<noteq> 0"
-  let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
-  {assume e: "even n"
-    hence "\<exists>m. n = 2*m" by presburger
-    then obtain m where m: "n = 2*m" by blast
-    from n m have "m\<noteq>0" "m < n" by presburger+
-    with IH[rule_format, of m] obtain z where z: "?P z m" by blast
-    from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
-    hence "\<exists>z. ?P z n" ..}
-  moreover
-  {assume o: "odd n"
-    from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
-    have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
-    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) = 
-    ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
-    also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2" 
-      apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
-      by (simp add: power2_eq_square)
-    finally 
-    have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
-    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
-    1" 
-      apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
-      using right_inverse[OF b']
-      by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] ring_simps)
-    have th0: "cmod (complex_of_real (cmod b) / b) = 1"
-      apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse ring_simps )
-      by (simp add: real_sqrt_mult[symmetric] th0)        
-    from o have "\<exists>m. n = Suc (2*m)" by presburger+
-    then obtain m where m: "n = Suc (2*m)" by blast
-    from unimodular_reduce_norm[OF th0] o
-    have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
-      apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
-      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_def)
-      apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
-      apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
-      apply (rule_tac x="- ii" in exI, simp add: m power_mult)
-      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_def)
-      apply (rule_tac x="ii" in exI, simp add: m power_mult diff_def)
-      done
-    then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
-    let ?w = "v / complex_of_real (root n (cmod b))"
-    from odd_real_root_pow[OF o, of "cmod b"]
-    have th1: "?w ^ n = v^n / complex_of_real (cmod b)" 
-      by (simp add: power_divide complex_of_real_power)
-    have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: norm_divide)
-    hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
-    have th4: "cmod (complex_of_real (cmod b) / b) *
-   cmod (1 + b * (v ^ n / complex_of_real (cmod b)))
-   < cmod (complex_of_real (cmod b) / b) * 1"
-      apply (simp only: norm_mult[symmetric] right_distrib)
-      using b v by (simp add: th2)
-
-    from mult_less_imp_less_left[OF th4 th3]
-    have "?P ?w n" unfolding th1 . 
-    hence "\<exists>z. ?P z n" .. }
-  ultimately show "\<exists>z. ?P z n" by blast
-qed
-
-
-text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
-
-lemma metric_bound_lemma: "cmod (x - y) <= \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
-  using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
-  unfolding cmod_def by simp
-
-lemma bolzano_weierstrass_complex_disc:
-  assumes r: "\<forall>n. cmod (s n) \<le> r"
-  shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
-proof-
-  from seq_monosub[of "Re o s"] 
-  obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))" 
-    unfolding o_def by blast
-  from seq_monosub[of "Im o s o f"] 
-  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast  
-  let ?h = "f o g"
-  from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith 
-  have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>" 
-  proof
-    fix n
-    from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
-  qed
-  have conv1: "convergent (\<lambda>n. Re (s ( f n)))"
-    apply (rule Bseq_monoseq_convergent)
-    apply (simp add: Bseq_def)
-    apply (rule exI[where x= "r + 1"])
-    using th rp apply simp
-    using f(2) .
-  have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>" 
-  proof
-    fix n
-    from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
-  qed
-
-  have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
-    apply (rule Bseq_monoseq_convergent)
-    apply (simp add: Bseq_def)
-    apply (rule exI[where x= "r + 1"])
-    using th rp apply simp
-    using g(2) .
-
-  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x" 
-    by blast 
-  hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r" 
-    unfolding LIMSEQ_def real_norm_def .
-
-  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y" 
-    by blast 
-  hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r" 
-    unfolding LIMSEQ_def real_norm_def .
-  let ?w = "Complex x y"
-  from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto 
-  {fix e assume ep: "e > (0::real)"
-    hence e2: "e/2 > 0" by simp
-    from x[rule_format, OF e2] y[rule_format, OF e2]
-    obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2" and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2" by blast
-    {fix n assume nN12: "n \<ge> N1 + N2"
-      hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
-      from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
-      have "cmod (s (?h n) - ?w) < e" 
-	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
-    hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
-  with hs show ?thesis  by blast  
-qed
-
-text{* Polynomial is continuous. *}
-
-lemma poly_cont:
-  assumes ep: "e > 0" 
-  shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
-proof-
-  from poly_offset[of p z] obtain q where q: "length q = length p" "\<And>x. poly q x = poly p (z + x)" by blast
-  {fix w
-    note q(2)[of "w - z", simplified]}
-  note th = this
-  show ?thesis unfolding th[symmetric]
-  proof(induct q)
-    case Nil thus ?case  using ep by auto
-  next
-    case (Cons c cs)
-    from poly_bound_exists[of 1 "cs"] 
-    obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
-    from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
-    have one0: "1 > (0::real)"  by arith
-    from real_lbound_gt_zero[OF one0 em0] 
-    obtain d where d: "d >0" "d < 1" "d < e / m" by blast
-    from d(1,3) m(1) have dm: "d*m > 0" "d*m < e" 
-      by (simp_all add: field_simps real_mult_order)
-    show ?case 
-      proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
-	fix d w
-	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
-	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
-	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
-	from H have th: "cmod (w-z) \<le> d" by simp 
-	from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
-	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
-      qed  
-    qed
-qed
-
-text{* Hence a polynomial attains minimum on a closed disc 
-  in the complex plane. *}
-lemma  poly_minimum_modulus_disc:
-  "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
-proof-
-  {assume "\<not> r \<ge> 0" hence ?thesis unfolding linorder_not_le
-      apply -
-      apply (rule exI[where x=0]) 
-      apply auto
-      apply (subgoal_tac "cmod w < 0")
-      apply simp
-      apply arith
-      done }
-  moreover
-  {assume rp: "r \<ge> 0"
-    from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp 
-    hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
-    {fix x z
-      assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
-      hence "- x < 0 " by arith
-      with H(2) norm_ge_zero[of "poly p z"]  have False by simp }
-    then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
-    from real_sup_exists[OF mth1 mth2] obtain s where 
-      s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow>(y < s)" by blast
-    let ?m = "-s"
-    {fix y
-      from s[rule_format, of "-y"] have 
-    "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" 
-	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
-    note s1 = this[unfolded minus_minus]
-    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" 
-      by auto
-    {fix n::nat
-      from s1[rule_format, of "?m + 1/real (Suc n)"] 
-      have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
-	by simp}
-    hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
-    from choice[OF th] obtain g where 
-      g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)" 
-      by blast
-    from bolzano_weierstrass_complex_disc[OF g(1)] 
-    obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
-      by blast    
-    {fix w 
-      assume wr: "cmod w \<le> r"
-      let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
-      {assume e: "?e > 0"
-	hence e2: "?e/2 > 0" by simp
-	from poly_cont[OF e2, of z p] obtain d where
-	  d: "d>0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2" by blast
-	{fix w assume w: "cmod (w - z) < d"
-	  have "cmod(poly p w - poly p z) < ?e / 2"
-	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
-	note th1 = this
-	
-	from fz(2)[rule_format, OF d(1)] obtain N1 where 
-	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
-	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
-	  N2: "2/?e < real N2" by blast
-	have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
-	  using N1[rule_format, of "N1 + N2"] th1 by simp
-	{fix a b e2 m :: real
-	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
-          ==> False" by arith}
-      note th0 = this
-      have ath: 
-	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
-      from s1m[OF g(1)[rule_format]]
-      have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
-      from seq_suble[OF fz(1), of "N1+N2"]
-      have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
-      have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"  
-	using N2 by auto
-      from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
-      from g(2)[rule_format, of "f (N1 + N2)"]
-      have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
-      from order_less_le_trans[OF th01 th00]
-      have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
-      from N2 have "2/?e < real (Suc (N1 + N2))" by arith
-      with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
-      have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
-      with ath[OF th31 th32]
-      have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith  
-      have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c" 
-	by arith
-      have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
-\<le> cmod (poly p (g (f (N1 + N2))) - poly p z)" 
-	by (simp add: norm_triangle_ineq3)
-      from ath2[OF th22, of ?m]
-      have thc2: "2*(?e/2) \<le> \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)" by simp
-      from th0[OF th2 thc1 thc2] have False .}
-      hence "?e = 0" by auto
-      then have "cmod (poly p z) = ?m" by simp  
-      with s1m[OF wr]
-      have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
-    hence ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma "(rcis (sqrt (abs r)) (a/2)) ^ 2 = rcis (abs r) a"
-  unfolding power2_eq_square
-  apply (simp add: rcis_mult)
-  apply (simp add: power2_eq_square[symmetric])
-  done
-
-lemma cispi: "cis pi = -1" 
-  unfolding cis_def
-  by simp
-
-lemma "(rcis (sqrt (abs r)) ((pi + a)/2)) ^ 2 = rcis (- abs r) a"
-  unfolding power2_eq_square
-  apply (simp add: rcis_mult add_divide_distrib)
-  apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
-  done
-
-text {* Nonzero polynomial in z goes to infinity as z does. *}
-
-instance complex::idom_char_0 by (intro_classes)
-instance complex :: recpower_idom_char_0 by intro_classes
-
-lemma poly_infinity:
-  assumes ex: "list_ex (\<lambda>c. c \<noteq> 0) p"
-  shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (a#p) z)"
-using ex
-proof(induct p arbitrary: a d)
-  case (Cons c cs a d) 
-  {assume H: "list_ex (\<lambda>c. c\<noteq>0) cs"
-    with Cons.hyps obtain r where r: "\<forall>z. r \<le> cmod z \<longrightarrow> d + cmod a \<le> cmod (poly (c # cs) z)" by blast
-    let ?r = "1 + \<bar>r\<bar>"
-    {fix z assume h: "1 + \<bar>r\<bar> \<le> cmod z"
-      have r0: "r \<le> cmod z" using h by arith
-      from r[rule_format, OF r0]
-      have th0: "d + cmod a \<le> 1 * cmod(poly (c#cs) z)" by arith
-      from h have z1: "cmod z \<ge> 1" by arith
-      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (c#cs) z"]]]
-      have th1: "d \<le> cmod(z * poly (c#cs) z) - cmod a"
-	unfolding norm_mult by (simp add: ring_simps)
-      from complex_mod_triangle_sub[of "z * poly (c#cs) z" a]
-      have th2: "cmod(z * poly (c#cs) z) - cmod a \<le> cmod (poly (a#c#cs) z)" 
-	by (simp add: diff_le_eq ring_simps) 
-      from th1 th2 have "d \<le> cmod (poly (a#c#cs) z)"  by arith}
-    hence ?case by blast}
-  moreover
-  {assume cs0: "\<not> (list_ex (\<lambda>c. c \<noteq> 0) cs)"
-    with Cons.prems have c0: "c \<noteq> 0" by simp
-    from cs0 have cs0': "list_all (\<lambda>c. c = 0) cs" 
-      by (auto simp add: list_all_iff list_ex_iff)
-    {fix z
-      assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
-      from c0 have "cmod c > 0" by simp
-      from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)" 
-	by (simp add: field_simps norm_mult)
-      have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
-      from complex_mod_triangle_sub[of "z*c" a ]
-      have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
-	by (simp add: ring_simps)
-      from ath[OF th1 th0] have "d \<le> cmod (poly (a # c # cs) z)" 
-	using poly_0[OF cs0'] by simp}
-    then have ?case  by blast}
-  ultimately show ?case by blast
-qed simp
-
-text {* Hence polynomial's modulus attains its minimum somewhere. *}
-lemma poly_minimum_modulus:
-  "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
-proof(induct p)
-  case (Cons c cs) 
-  {assume cs0: "list_ex (\<lambda>c. c \<noteq> 0) cs"
-    from poly_infinity[OF cs0, of "cmod (poly (c#cs) 0)" c]
-    obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (c # cs) 0) \<le> cmod (poly (c # cs) z)" by blast
-    have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
-    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "c#cs"] 
-    obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) w)" by blast
-    {fix z assume z: "r \<le> cmod z"
-      from v[of 0] r[OF z] 
-      have "cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) z)"
-	by simp }
-    note v0 = this
-    from v0 v ath[of r] have ?case by blast}
-  moreover
-  {assume cs0: "\<not> (list_ex (\<lambda>c. c\<noteq>0) cs)"
-    hence th:"list_all (\<lambda>c. c = 0) cs" by (simp add: list_all_iff list_ex_iff)
-    from poly_0[OF th] Cons.hyps have ?case by simp}
-  ultimately show ?case by blast
-qed simp
-
-text{* Constant function (non-syntactic characterization). *}
-definition "constant f = (\<forall>x y. f x = f y)"
-
-lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> length p \<ge> 2"
-  unfolding constant_def
-  apply (induct p, auto)
-  apply (unfold not_less[symmetric])
-  apply simp
-  apply (rule ccontr)
-  apply auto
-  done
- 
-lemma poly_replicate_append:
-  "poly ((replicate n 0)@p) (x::'a::{recpower, comm_ring}) = x^n * poly p x"
-  by(induct n, auto simp add: power_Suc ring_simps)
-
-text {* Decomposition of polynomial, skipping zero coefficients 
-  after the first.  *}
-
-lemma poly_decompose_lemma:
- assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
-  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (length q + k) = length p \<and> 
-                 (\<forall>z. poly p z = z^k * poly (a#q) z)"
-using nz
-proof(induct p)
-  case Nil thus ?case by simp
-next
-  case (Cons c cs)
-  {assume c0: "c = 0"
-    
-    from Cons.hyps Cons.prems c0 have ?case apply auto
-      apply (rule_tac x="k+1" in exI)
-      apply (rule_tac x="a" in exI, clarsimp)
-      apply (rule_tac x="q" in exI)
-      by (auto simp add: power_Suc)}
-  moreover
-  {assume c0: "c\<noteq>0"
-    hence ?case apply-
-      apply (rule exI[where x=0])
-      apply (rule exI[where x=c], clarsimp)
-      apply (rule exI[where x=cs])
-      apply auto
-      done}
-  ultimately show ?case by blast
-qed
-
-lemma poly_decompose:
-  assumes nc: "~constant(poly p)"
-  shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
-               length q + k + 1 = length p \<and> 
-              (\<forall>z. poly p z = poly p 0 + z^k * poly (a#q) z)"
-using nc 
-proof(induct p)
-  case Nil thus ?case by (simp add: constant_def)
-next
-  case (Cons c cs)
-  {assume C:"\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
-    {fix x y
-      from C have "poly (c#cs) x = poly (c#cs) y" by (cases "x=0", auto)}
-    with Cons.prems have False by (auto simp add: constant_def)}
-  hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
-  from poly_decompose_lemma[OF th] 
-  show ?case 
-    apply clarsimp    
-    apply (rule_tac x="k+1" in exI)
-    apply (rule_tac x="a" in exI)
-    apply simp
-    apply (rule_tac x="q" in exI)
-    apply (auto simp add: power_Suc)
-    done
-qed
-
-text{* Fundamental theorem of algebral *}
-
-lemma fundamental_theorem_of_algebra:
-  assumes nc: "~constant(poly p)"
-  shows "\<exists>z::complex. poly p z = 0"
-using nc
-proof(induct n\<equiv> "length p" arbitrary: p rule: nat_less_induct)
-  fix n fix p :: "complex list"
-  let ?p = "poly p"
-  assume H: "\<forall>m<n. \<forall>p. \<not> constant (poly p) \<longrightarrow> m = length p \<longrightarrow> (\<exists>(z::complex). poly p z = 0)" and nc: "\<not> constant ?p" and n: "n = length p"
-  let ?ths = "\<exists>z. ?p z = 0"
-
-  from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
-  from poly_minimum_modulus obtain c where 
-    c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" by blast
-  {assume pc: "?p c = 0" hence ?ths by blast}
-  moreover
-  {assume pc0: "?p c \<noteq> 0"
-    from poly_offset[of p c] obtain q where
-      q: "length q = length p" "\<forall>x. poly q x = ?p (c+x)" by blast
-    {assume h: "constant (poly q)"
-      from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
-      {fix x y
-	from th have "?p x = poly q (x - c)" by auto 
-	also have "\<dots> = poly q (y - c)" 
-	  using h unfolding constant_def by blast
-	also have "\<dots> = ?p y" using th by auto
-	finally have "?p x = ?p y" .}
-      with nc have False unfolding constant_def by blast }
-    hence qnc: "\<not> constant (poly q)" by blast
-    from q(2) have pqc0: "?p c = poly q 0" by simp
-    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp 
-    let ?a0 = "poly q 0"
-    from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
-    from a00 
-    have qr: "\<forall>z. poly q z = poly (map (op * (inverse ?a0)) q) z * ?a0"
-      by (simp add: poly_cmult_map)
-    let ?r = "map (op * (inverse ?a0)) q"
-    have lgqr: "length q = length ?r" by simp 
-    {assume h: "\<And>x y. poly ?r x = poly ?r y"
-      {fix x y
-	from qr[rule_format, of x] 
-	have "poly q x = poly ?r x * ?a0" by auto
-	also have "\<dots> = poly ?r y * ?a0" using h by simp
-	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
-	finally have "poly q x = poly q y" .} 
-      with qnc have False unfolding constant_def by blast}
-    hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
-    from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
-    {fix w 
-      have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
-	using qr[rule_format, of w] a00 by simp
-      also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
-	using a00 unfolding norm_divide by (simp add: field_simps)
-      finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
-    note mrmq_eq = this
-    from poly_decompose[OF rnc] obtain k a s where 
-      kas: "a\<noteq>0" "k\<noteq>0" "length s + k + 1 = length ?r" 
-      "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (a#s) z" by blast
-    {assume "k + 1 = n"
-      with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=[]" by auto
-      {fix w
-	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
-	  using kas(4)[rule_format, of w] s0 r01 by (simp add: ring_simps)}
-      note hth = this [symmetric]
-	from reduce_poly_simple[OF kas(1,2)] 
-      have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
-    moreover
-    {assume kn: "k+1 \<noteq> n"
-      from kn kas(3) q(1) n[symmetric] have k1n: "k + 1 < n" by simp
-      have th01: "\<not> constant (poly (1#((replicate (k - 1) 0)@[a])))" 
-	unfolding constant_def poly_Nil poly_Cons poly_replicate_append
-	using kas(1) apply simp 
-	by (rule exI[where x=0], rule exI[where x=1], simp)
-      from kas(2) have th02: "k+1 = length (1#((replicate (k - 1) 0)@[a]))" 
-	by simp
-      from H[rule_format, OF k1n th01 th02]
-      obtain w where w: "1 + w^k * a = 0"
-	unfolding poly_Nil poly_Cons poly_replicate_append
-	using kas(2) by (auto simp add: power_Suc[symmetric, of _ "k - Suc 0"] 
-	  mult_assoc[of _ _ a, symmetric])
-      from poly_bound_exists[of "cmod w" s] obtain m where 
-	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
-      have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
-      from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
-      then have wm1: "w^k * a = - 1" by simp
-      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" 
-	using norm_ge_zero[of w] w0 m(1)
-	  by (simp add: inverse_eq_divide zero_less_mult_iff)
-      with real_down2[OF zero_less_one] obtain t where
-	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
-      let ?ct = "complex_of_real t"
-      let ?w = "?ct * w"
-      have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: ring_simps power_mult_distrib)
-      also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
-	unfolding wm1 by (simp)
-      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
-	apply -
-	apply (rule cong[OF refl[of cmod]])
-	apply assumption
-	done
-      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] 
-      have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" unfolding norm_of_real by simp 
-      have ath: "\<And>x (t::real). 0\<le> x \<Longrightarrow> x < t \<Longrightarrow> t\<le>1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1" by arith
-      have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
-      then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult) 
-      from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
-	by (simp add: inverse_eq_divide field_simps)
-      with zero_less_power[OF t(1), of k] 
-      have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
-	apply - apply (rule mult_strict_left_mono) by simp_all
-      have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
-	by (simp add: ring_simps power_mult_distrib norm_of_real norm_power norm_mult)
-      then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
-	using t(1,2) m(2)[rule_format, OF tw] w0
-	apply (simp only: )
-	apply auto
-	apply (rule mult_mono, simp_all add: norm_ge_zero)+
-	apply (simp add: zero_le_mult_iff zero_le_power)
-	done
-      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp 
-      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1" 
-	by auto
-      from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
-      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" . 
-      from th11 th12
-      have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith 
-      then have "cmod (poly ?r ?w) < 1" 
-	unfolding kas(4)[rule_format, of ?w] r01 by simp 
-      then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
-    ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
-    from cr0_contr cq0 q(2)
-    have ?ths unfolding mrmq_eq not_less[symmetric] by auto}
-  ultimately show ?ths by blast
-qed
-
-text {* Alternative version with a syntactic notion of constant polynomial. *}
-
-lemma fundamental_theorem_of_algebra_alt:
-  assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> list_all(\<lambda>b. b = 0) l \<and> p = a#l)"
-  shows "\<exists>z. poly p z = (0::complex)"
-using nc
-proof(induct p)
-  case (Cons c cs)
-  {assume "c=0" hence ?case by auto}
-  moreover
-  {assume c0: "c\<noteq>0"
-    {assume nc: "constant (poly (c#cs))"
-      from nc[unfolded constant_def, rule_format, of 0] 
-      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
-      hence "list_all (\<lambda>c. c=0) cs"
-	proof(induct cs)
-	  case (Cons d ds)
-	  {assume "d=0" hence ?case using Cons.prems Cons.hyps by simp}
-	  moreover
-	  {assume d0: "d\<noteq>0"
-	    from poly_bound_exists[of 1 ds] obtain m where 
-	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
-	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
-	    from real_down2[OF dm zero_less_one] obtain x where 
-	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
-	    let ?x = "complex_of_real x"
-	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
-	    from Cons.prems[rule_format, OF cx(1)]
-	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
-	    from m(2)[rule_format, OF cx(2)] x(1)
-	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
-	      by (simp add: norm_mult)
-	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
-	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
-	    with cth  have ?case by blast}
-	  ultimately show ?case by blast 
-	qed simp}
-      then have nc: "\<not> constant (poly (c#cs))" using Cons.prems c0 
-	by blast
-      from fundamental_theorem_of_algebra[OF nc] have ?case .}
-  ultimately show ?case by blast  
-qed simp
-
-subsection{* Nullstellenstatz, degrees and divisibility of polynomials *}
-
-lemma nullstellensatz_lemma:
-  fixes p :: "complex list"
-  assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
-  and "degree p = n" and "n \<noteq> 0"
-  shows "p divides (pexp q n)"
-using prems
-proof(induct n arbitrary: p q rule: nat_less_induct)
-  fix n::nat fix p q :: "complex list"
-  assume IH: "\<forall>m<n. \<forall>p q.
-                 (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
-                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p divides (q %^ m)"
-    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
-    and dpn: "degree p = n" and n0: "n \<noteq> 0"
-  let ?ths = "p divides (q %^ n)"
-  {fix a assume a: "poly p a = 0"
-    {assume p0: "poly p = poly []" 
-      hence ?ths unfolding divides_def  using pq0 n0
-	apply - apply (rule exI[where x="[]"], rule ext)
-	by (auto simp add: poly_mult poly_exp)}
-    moreover
-    {assume p0: "poly p \<noteq> poly []" 
-      and oa: "order  a p \<noteq> 0"
-      from p0 have pne: "p \<noteq> []" by auto
-      let ?op = "order a p"
-      from p0 have ap: "([- a, 1] %^ ?op) divides p" 
-	"\<not> pexp [- a, 1] (Suc ?op) divides p" using order by blast+ 
-      note oop = order_degree[OF p0, unfolded dpn]
-      {assume q0: "q = []"
-	hence ?ths using n0 unfolding divides_def 
-	  apply simp
-	  apply (rule exI[where x="[]"], rule ext)
-	  by (simp add: divides_def poly_exp poly_mult)}
-      moreover
-      {assume q0: "q\<noteq>[]"
-	from pq0[rule_format, OF a, unfolded poly_linear_divides] q0
-	obtain r where r: "q = pmult [- a, 1] r" by blast
-	from ap[unfolded divides_def] obtain s where
-	  s: "poly p = poly (pmult (pexp [- a, 1] ?op) s)" by blast
-	have s0: "poly s \<noteq> poly []"
-	  using s p0 by (simp add: poly_entire)
-	hence pns0: "poly (pnormalize s) \<noteq> poly []" and sne: "s\<noteq>[]" by auto
-	{assume ds0: "degree s = 0"
-	  from ds0 pns0 have "\<exists>k. pnormalize s = [k]" unfolding degree_def 
-	    by (cases "pnormalize s", auto)
-	  then obtain k where kpn: "pnormalize s = [k]" by blast
-	  from pns0[unfolded poly_zero] kpn have k: "k \<noteq>0" "poly s = poly [k]"
-	    using poly_normalize[of s] by simp_all
-	  let ?w = "pmult (pmult [1/k] (pexp [-a,1] (n - ?op))) (pexp r n)"
-	  from k r s oop have "poly (pexp q n) = poly (pmult p ?w)"
-	    by - (rule ext, simp add: poly_mult poly_exp poly_cmult poly_add power_add[symmetric] ring_simps power_mult_distrib[symmetric])
-	  hence ?ths unfolding divides_def by blast}
-	moreover
-	{assume ds0: "degree s \<noteq> 0"
-	  from ds0 s0 dpn degree_unique[OF s, unfolded linear_pow_mul_degree] oa
-	    have dsn: "degree s < n" by auto 
-	    {fix x assume h: "poly s x = 0"
-	      {assume xa: "x = a"
-		from h[unfolded xa poly_linear_divides] sne obtain u where
-		  u: "s = pmult [- a, 1] u" by blast
-		have "poly p = poly (pmult (pexp [- a, 1] (Suc ?op)) u)"
-		  unfolding s u
-		  apply (rule ext)
-		  by (simp add: ring_simps power_mult_distrib[symmetric] poly_mult poly_cmult poly_add poly_exp)
-		with ap(2)[unfolded divides_def] have False by blast}
-	      note xa = this
-	      from h s have "poly p x = 0" by (simp add: poly_mult)
-	      with pq0 have "poly q x = 0" by blast
-	      with r xa have "poly r x = 0"
-		by (auto simp add: poly_mult poly_add poly_cmult eq_diff_eq[symmetric])}
-	    note impth = this
-	    from IH[rule_format, OF dsn, of s r] impth ds0
-	    have "s divides (pexp r (degree s))" by blast
-	    then obtain u where u: "poly (pexp r (degree s)) = poly (pmult s u)"
-	      unfolding divides_def by blast
-	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
-	      by (simp add: poly_mult[symmetric] poly_exp[symmetric])
-	    let ?w = "pmult (pmult u (pexp [-a,1] (n - ?op))) (pexp r (n - degree s))"
-	    from u' s r oop[of a] dsn have "poly (pexp q n) = poly (pmult p ?w)"
-	      apply - apply (rule ext)
-	      apply (simp only:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult ring_simps)
-	      
-	      apply (simp add:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult mult_assoc[symmetric])
-	      done
-	    hence ?ths unfolding divides_def by blast}
-      ultimately have ?ths by blast }
-      ultimately have ?ths by blast}
-    ultimately have ?ths using a order_root by blast}
-  moreover
-  {assume exa: "\<not> (\<exists>a. poly p a = 0)"
-    from fundamental_theorem_of_algebra_alt[of p] exa obtain c cs where
-      ccs: "c\<noteq>0" "list_all (\<lambda>c. c = 0) cs" "p = c#cs" by blast
-    
-    from poly_0[OF ccs(2)] ccs(3) 
-    have pp: "\<And>x. poly p x =  c" by simp
-    let ?w = "pmult [1/c] (pexp q n)"
-    from pp ccs(1) 
-    have "poly (pexp q n) = poly (pmult p ?w) "
-      apply - apply (rule ext)
-      unfolding poly_mult_assoc[symmetric] by (simp add: poly_mult)
-    hence ?ths unfolding divides_def by blast}
-  ultimately show ?ths by blast
-qed
-
-lemma nullstellensatz_univariate:
-  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
-    p divides (q %^ (degree p)) \<or> (poly p = poly [] \<and> poly q = poly [])"
-proof-
-  {assume pe: "poly p = poly []"
-    hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> poly q = poly []"
-      apply auto
-      by (rule ext, simp)
-    {assume "p divides (pexp q (degree p))"
-      then obtain r where r: "poly (pexp q (degree p)) = poly (pmult p r)" 
-	unfolding divides_def by blast
-      from cong[OF r refl] pe degree_unique[OF pe]
-      have False by (simp add: poly_mult degree_def)}
-    with eq pe have ?thesis by blast}
-  moreover
-  {assume pe: "poly p \<noteq> poly []"
-    have p0: "poly [0] = poly []" by (rule ext, simp)
-    {assume dp: "degree p = 0"
-      then obtain k where "pnormalize p = [k]" using pe poly_normalize[of p]
-	unfolding degree_def by (cases "pnormalize p", auto)
-      hence k: "pnormalize p = [k]" "poly p = poly [k]" "k\<noteq>0"
-	using pe poly_normalize[of p] by (auto simp add: p0)
-      hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
-      from k(2,3) dp have "poly (pexp q (degree p)) = poly (pmult p [1/k]) "
-	by - (rule ext, simp add: poly_mult poly_exp)
-      hence th2: "p divides (pexp q (degree p))" unfolding divides_def by blast
-      from th1 th2 pe have ?thesis by blast}
-    moreover
-    {assume dp: "degree p \<noteq> 0"
-      then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
-      {assume "p divides (pexp q (Suc n))"
-	then obtain u where u: "poly (pexp q (Suc n)) = poly (pmult p u)"
-	  unfolding divides_def by blast
-	hence u' :"\<And>x. poly (pexp q (Suc n)) x = poly (pmult p u) x" by simp_all
-	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
-	  hence "poly (pexp q (Suc n)) x \<noteq> 0" by (simp only: poly_exp) simp	  
-	  hence False using u' h(1) by (simp only: poly_mult poly_exp) simp}}
-	with n nullstellensatz_lemma[of p q "degree p"] dp 
-	have ?thesis by auto}
-    ultimately have ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-text{* Useful lemma *}
-
-lemma (in idom_char_0) constant_degree: "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
-proof
-  assume l: ?lhs
-  from l[unfolded constant_def, rule_format, of _ "zero"]
-  have th: "poly p = poly [poly p 0]" apply - by (rule ext, simp)
-  from degree_unique[OF th] show ?rhs by (simp add: degree_def)
-next
-  assume r: ?rhs
-  from r have "pnormalize p = [] \<or> (\<exists>k. pnormalize p = [k])"
-    unfolding degree_def by (cases "pnormalize p", auto)
-  then show ?lhs unfolding constant_def poly_normalize[of p, symmetric]
-    by (auto simp del: poly_normalize)
-qed
-
-(* It would be nicer to prove this without using algebraic closure...        *)
-
-lemma divides_degree_lemma: assumes dpn: "degree (p::complex list) = n"
-  shows "n \<le> degree (p *** q) \<or> poly (p *** q) = poly []"
-  using dpn
-proof(induct n arbitrary: p q)
-  case 0 thus ?case by simp
-next
-  case (Suc n p q)
-  from Suc.prems fundamental_theorem_of_algebra[of p] constant_degree[of p]
-  obtain a where a: "poly p a = 0" by auto
-  then obtain r where r: "p = pmult [-a, 1] r" unfolding poly_linear_divides
-    using Suc.prems by (auto simp add: degree_def)
-  {assume h: "poly (pmult r q) = poly []"
-    hence "poly (pmult p q) = poly []" using r
-      apply - apply (rule ext)  by (auto simp add: poly_entire poly_mult poly_add poly_cmult) hence ?case by blast}
-  moreover
-  {assume h: "poly (pmult r q) \<noteq> poly []" 
-    hence r0: "poly r \<noteq> poly []" and q0: "poly q \<noteq> poly []"
-      by (auto simp add: poly_entire)
-    have eq: "poly (pmult p q) = poly (pmult [-a, 1] (pmult r q))"
-      apply - apply (rule ext)
-      by (simp add: r poly_mult poly_add poly_cmult ring_simps)
-    from linear_mul_degree[OF h, of "- a"]
-    have dqe: "degree (pmult p q) = degree (pmult r q) + 1"
-      unfolding degree_unique[OF eq] .
-    from linear_mul_degree[OF r0, of "- a", unfolded r[symmetric]] r Suc.prems 
-    have dr: "degree r = n" by auto
-    from  Suc.hyps[OF dr, of q] have "Suc n \<le> degree (pmult p q)"
-      unfolding dqe using h by (auto simp del: poly.simps) 
-    hence ?case by blast}
-  ultimately show ?case by blast
-qed
-
-lemma divides_degree: assumes pq: "p divides (q:: complex list)"
-  shows "degree p \<le> degree q \<or> poly q = poly []"
-using pq  divides_degree_lemma[OF refl, of p]
-apply (auto simp add: divides_def poly_entire)
-apply atomize
-apply (erule_tac x="qa" in allE, auto)
-apply (subgoal_tac "degree q = degree (p *** qa)", simp)
-apply (rule degree_unique, simp)
-done
-
-(* Arithmetic operations on multivariate polynomials.                        *)
-
-lemma mpoly_base_conv: 
-  "(0::complex) \<equiv> poly [] x" "c \<equiv> poly [c] x" "x \<equiv> poly [0,1] x" by simp_all
-
-lemma mpoly_norm_conv: 
-  "poly [0] (x::complex) \<equiv> poly [] x" "poly [poly [] y] x \<equiv> poly [] x" by simp_all
-
-lemma mpoly_sub_conv: 
-  "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
-  by (simp add: diff_def)
-
-lemma poly_pad_rule: "poly p x = 0 ==> poly (0#p) x = (0::complex)" by simp
-
-lemma poly_cancel_eq_conv: "p = (0::complex) \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (q = 0) \<equiv> (a * q - b * p = 0)" apply (atomize (full)) by auto
-
-lemma resolve_eq_raw:  "poly [] x \<equiv> 0" "poly [c] x \<equiv> (c::complex)" by auto
-lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
-  \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
-lemma expand_ex_beta_conv: "list_ex P [c] \<equiv> P c" by simp
-
-lemma poly_divides_pad_rule: 
-  fixes p q :: "complex list"
-  assumes pq: "p divides q"
-  shows "p divides ((0::complex)#q)"
-proof-
-  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
-  hence "poly (0#q) = poly (p *** ([0,1] *** r))" 
-    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
-  thus ?thesis unfolding divides_def by blast
-qed
-
-lemma poly_divides_pad_const_rule: 
-  fixes p q :: "complex list"
-  assumes pq: "p divides q"
-  shows "p divides (a %* q)"
-proof-
-  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
-  hence "poly (a %* q) = poly (p *** (a %* r))" 
-    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
-  thus ?thesis unfolding divides_def by blast
-qed
-
-
-lemma poly_divides_conv0:  
-  fixes p :: "complex list"
-  assumes lgpq: "length q < length p" and lq:"last p \<noteq> 0"
-  shows "p divides q \<equiv> (\<not> (list_ex (\<lambda>c. c \<noteq> 0) q))" (is "?lhs \<equiv> ?rhs")
-proof-
-  {assume r: ?rhs 
-    hence eq: "poly q = poly []" unfolding poly_zero 
-      by (simp add: list_all_iff list_ex_iff)
-    hence "poly q = poly (p *** [])" by - (rule ext, simp add: poly_mult)
-    hence ?lhs unfolding divides_def  by blast}
-  moreover
-  {assume l: ?lhs
-    have ath: "\<And>lq lp dq::nat. lq < lp ==> lq \<noteq> 0 \<Longrightarrow> dq <= lq - 1 ==> dq < lp - 1"
-      by arith
-    {assume q0: "length q = 0"
-      hence "q = []" by simp
-      hence ?rhs by simp}
-    moreover
-    {assume lgq0: "length q \<noteq> 0"
-      from pnormalize_length[of q] have dql: "degree q \<le> length q - 1" 
-	unfolding degree_def by simp
-      from ath[OF lgpq lgq0 dql, unfolded pnormal_degree[OF lq, symmetric]] divides_degree[OF l] have "poly q = poly []" by auto
-      hence ?rhs unfolding poly_zero by (simp add: list_all_iff list_ex_iff)}
-    ultimately have ?rhs by blast }
-  ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
-qed
-
-lemma poly_divides_conv1: 
-  assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex list) divides p'"
-  and qrp': "\<And>x. a * poly q x - poly p' x \<equiv> poly r x"
-  shows "p divides q \<equiv> p divides (r::complex list)" (is "?lhs \<equiv> ?rhs")
-proof-
-  {
-  from pp' obtain t where t: "poly p' = poly (p *** t)" 
-    unfolding divides_def by blast
-  {assume l: ?lhs
-    then obtain u where u: "poly q = poly (p *** u)" unfolding divides_def by blast
-     have "poly r = poly (p *** ((a %* u) +++ (-- t)))"
-       using u qrp' t
-       by - (rule ext, 
-	 simp add: poly_add poly_mult poly_cmult poly_minus ring_simps)
-     then have ?rhs unfolding divides_def by blast}
-  moreover
-  {assume r: ?rhs
-    then obtain u where u: "poly r = poly (p *** u)" unfolding divides_def by blast
-    from u t qrp' a0 have "poly q = poly (p *** ((1/a) %* (u +++ t)))"
-      by - (rule ext, atomize (full), simp add: poly_mult poly_add poly_cmult field_simps)
-    hence ?lhs  unfolding divides_def by blast}
-  ultimately have "?lhs = ?rhs" by blast }
-thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
-qed
-
-lemma basic_cqe_conv1:
-  "(\<exists>x. poly p x = 0 \<and> poly [] x \<noteq> 0) \<equiv> False"
-  "(\<exists>x. poly [] x \<noteq> 0) \<equiv> False"
-  "(\<exists>x. poly [c] x \<noteq> 0) \<equiv> c\<noteq>0"
-  "(\<exists>x. poly [] x = 0) \<equiv> True"
-  "(\<exists>x. poly [c] x = 0) \<equiv> c = 0" by simp_all
-
-lemma basic_cqe_conv2: 
-  assumes l:"last (a#b#p) \<noteq> 0" 
-  shows "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True"
-proof-
-  {fix h t
-    assume h: "h\<noteq>0" "list_all (\<lambda>c. c=(0::complex)) t"  "a#b#p = h#t"
-    hence "list_all (\<lambda>c. c= 0) (b#p)" by simp
-    moreover have "last (b#p) \<in> set (b#p)" by simp
-    ultimately have "last (b#p) = 0" by (simp add: list_all_iff)
-    with l have False by simp}
-  hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> list_all (\<lambda>c. c=0) t \<and> a#b#p = h#t)"
-    by blast
-  from fundamental_theorem_of_algebra_alt[OF th] 
-  show "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True" by auto
-qed
-
-lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
-proof-
-  have "\<not> (list_ex (\<lambda>c. c \<noteq> 0) p) \<longleftrightarrow> poly p = poly []" 
-    by (simp add: poly_zero list_all_iff list_ex_iff)
-  also have "\<dots> \<longleftrightarrow> (\<not> (\<exists>x. poly p x \<noteq> 0))" by (auto intro: ext)
-  finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
-    by - (atomize (full), blast)
-qed
-
-lemma basic_cqe_conv3:
-  fixes p q :: "complex list"
-  assumes l: "last (a#p) \<noteq> 0" 
-  shows "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
-proof-
-  note np = pnormalize_eq[OF l]
-  {assume "poly (a#p) = poly []" hence False using l
-      unfolding poly_zero apply (auto simp add: list_all_iff del: last.simps)
-      apply (cases p, simp_all) done}
-  then have p0: "poly (a#p) \<noteq> poly []"  by blast
-  from np have dp:"degree (a#p) = length p" by (simp add: degree_def)
-  from nullstellensatz_univariate[of "a#p" q] p0 dp
-  show "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
-    by - (atomize (full), auto)
-qed
-
-lemma basic_cqe_conv4:
-  fixes p q :: "complex list"
-  assumes h: "\<And>x. poly (q %^ n) x \<equiv> poly r x"
-  shows "p divides (q %^ n) \<equiv> p divides r"
-proof-
-  from h have "poly (q %^ n) = poly r" by (auto intro: ext)  
-  thus "p divides (q %^ n) \<equiv> p divides r" unfolding divides_def by simp
-qed
-
-lemma pmult_Cons_Cons: "((a::complex)#b#p) *** q = (a %*q) +++ (0#((b#p) *** q))"
-  by simp
-
-lemma elim_neg_conv: "- z \<equiv> (-1) * (z::complex)" by simp
-lemma eqT_intr: "PROP P \<Longrightarrow> (True \<Longrightarrow> PROP P )" "PROP P \<Longrightarrow> True" by blast+
-lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
-lemma last_simps: "last [x] = x" "last (x#y#ys) = last (y#ys)" by simp_all
-lemma length_simps: "length [] = 0" "length (x#y#xs) = length xs + 2" "length [x] = 1" by simp_all
-
-lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
-lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
-  by (atomize (full)) simp_all
-lemma cqe_conv1: "poly [] x = 0 \<longleftrightarrow> True"  by simp
-lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
-proof
-  assume "p \<Longrightarrow> q \<equiv> r" thus "p \<and> q \<equiv> p \<and> r" apply - apply (atomize (full)) by blast
-next
-  assume "p \<and> q \<equiv> p \<and> r" "p"
-  thus "q \<equiv> r" apply - apply (atomize (full)) apply blast done
-qed
-lemma poly_const_conv: "poly [c] (x::complex) = y \<longleftrightarrow> c = y" by simp
-
-end
\ No newline at end of file
--- a/src/HOL/Complex/README.html	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,67 +0,0 @@
-<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
-
-<!-- $Id$ -->
-
-<HTML>
-
-<HEAD>
-  <meta http-equiv="content-type" content="text/html;charset=iso-8859-1">
-  <TITLE>HOL/Complex/README</TITLE>
-</HEAD>
-
-<BODY>
-
-<H1>Complex: The Complex Numbers</H1>
-		<P>This directory defines the type <KBD>complex</KBD> of the complex numbers,
-with numeric constants and some complex analysis.  The development includes
-nonstandard analysis for the complex numbers.  Note that the image
-<KBD>HOL-Complex</KBD> includes theories from the directories 
-<KBD><a href="#Anchor-Real">HOL/Real</a></KBD>  and <KBD><a href="#Anchor-Hyperreal">HOL/Hyperreal</a></KBD>. They define other types including <kbd>real</kbd> (the real numbers) and <kbd>hypreal</kbd> (the hyperreal or non-standard reals).
-
-<ul>
-<li><a href="CLim.html">CLim</a> Limits, continuous functions, and derivatives for the complex numbers
-<li><a href="CSeries.html">CSeries</a> Finite summation and infinite series for the complex numbers
-<li><a href="CStar.html">CStar</a> Star-transforms for the complex numbers, to form non-standard extensions of sets and functions
-<li><a href="Complex.html">Complex</a> The complex numbers
-<li><a href="NSCA.html">NSCA</a> Nonstandard complex analysis
-<li><a href="NSComplex.html">NSComplex</a> Ultrapower construction of the nonstandard complex numbers
-</ul>
-
-<h2><a name="Anchor-Real" id="Anchor-Real"></a>Real: Dedekind Cut Construction of the Real Line</h2>
-
-<ul>
-<li><a href="Lubs.html">Lubs</a> Definition of upper bounds, lubs and so on, to support completeness proofs.
-<li><a href="PReal.html">PReal</a> The positive reals constructed using Dedekind cuts
-<li><a href="Rational.html">Rational</a> The rational numbers constructed as equivalence classes of integers
-<li><a href="RComplete.html">RComplete</a> The reals are complete: they satisfy the supremum property. They also have the Archimedean property.
-<li><a href="RealDef.html">RealDef</a> The real numbers, their ordering properties, and embedding of the integers and the natural numbers
-<li><a href="RealPow.html">RealPow</a> Real numbers raised to natural number powers
-</ul>
-<h2><a name="Anchor-Hyperreal" id="Anchor-Hyperreal"></a>Hyperreal: Ultrafilter Construction of the Non-Standard Reals</h2>
-See J. D. Fleuriot and L. C. Paulson. Mechanizing Nonstandard Real Analysis. LMS J. Computation and Mathematics 3 (2000), 140-190.
-<ul>
-<li><a href="Filter.html">Filter</a> Theory of Filters and Ultrafilters. Main result is a version of the Ultrafilter Theorem proved using Zorn's Lemma.
-<li><a href="HLog.html">HLog</a> Non-standard logarithms
-<li><a href="HSeries.html">HSeries</a> Non-standard theory of finite summation and infinite series
-<li><a href="HTranscendental.html">HTranscendental</a> Non-standard extensions of transcendental functions
-<li><a href="HyperDef.html">HyperDef</a> Ultrapower construction of the hyperreals
-<li><a href="HyperNat.html">HyperNat</a> Ultrapower construction of the hypernaturals
-<li><a href="HyperPow.html">HyperPow</a> Powers theory for the hyperreals
-<!-- <li><a href="IntFloor.html">IntFloor</a> Floor and Ceiling functions relating the reals and integers -->
-<li><a href="Integration.html">Integration</a> Gage integrals
-<li><a href="Lim.html">Lim</a> Theory of limits, continuous functions, and derivatives
-<li><a href="Log.html">Log</a> Logarithms for the reals
-<li><a href="MacLaurin.html">MacLaurin</a> MacLaurin series
-<li><a href="NatStar.html">NatStar</a> Star-transforms for the hypernaturals, to form non-standard extensions of sets and functions involving the naturals or reals
-<li><a href="NthRoot.html">NthRoot</a> Existence of n-th roots of real numbers
-<li><a href="NSA.html">NSA</a> Theory defining sets of infinite numbers, infinitesimals, the infinitely close relation, and their various algebraic properties.
-<li><a href="Poly.html">Poly</a> Univariate real polynomials
-<li><a href="SEQ.html">SEQ</a> Convergence of sequences and series using standard and nonstandard analysis
-<li><a href="Series.html">Series</a> Finite summation and infinite series for the reals
-<li><a href="Star.html">Star</a> Nonstandard extensions of real sets and real functions
-<li><a href="Transcendental.html">Transcendental</a> Power series and transcendental functions
-</ul>
-<HR>
-<P>Last modified $Date$
-</BODY>
-</HTML>
--- a/src/HOL/Complex/document/root.tex	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,32 +0,0 @@
-
-% $Id$
-
-\documentclass[11pt,a4paper]{article}
-\usepackage{graphicx,isabelle,isabellesym,latexsym}
-\usepackage[latin1]{inputenc}
-\usepackage{pdfsetup}
-
-\urlstyle{rm}
-\isabellestyle{it}
-\pagestyle{myheadings}
-
-\begin{document}
-
-\title{Isabelle/HOL-Complex --- Higher-Order Logic with Complex Numbers}
-\maketitle
-
-\tableofcontents
-
-\begin{center}
-  \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph}
-\end{center}
-
-\newpage
-
-\renewcommand{\isamarkupheader}[1]%
-{\section{\isabellecontext: #1}\markright{THEORY~``\isabellecontext''}}
-
-\parindent 0pt\parskip 0.5ex
-\input{session}
-
-\end{document}
--- a/src/HOL/Complex_Main.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Complex_Main.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,15 +1,10 @@
-(*  Title:      HOL/Complex_Main.thy
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2003  University of Cambridge
-*)
-
-header{*Comprehensive Complex Theory*}
+header {* Comprehensive Complex Theory *}
 
 theory Complex_Main
 imports
   Main
   Real
-  "~~/src/HOL/Complex/Fundamental_Theorem_Algebra"
+  Fundamental_Theorem_Algebra
   Log
   Ln
   Taylor
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Dense_Linear_Order.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,877 @@
+(* Author: Amine Chaieb, TU Muenchen *)
+
+header {* Dense linear order without endpoints
+  and a quantifier elimination procedure in Ferrante and Rackoff style *}
+
+theory Dense_Linear_Order
+imports Plain Groebner_Basis
+uses
+  "~~/src/HOL/Tools/Qelim/langford_data.ML"
+  "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML"
+  ("~~/src/HOL/Tools/Qelim/langford.ML")
+  ("~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML")
+begin
+
+setup {* Langford_Data.setup #> Ferrante_Rackoff_Data.setup *}
+
+context linorder
+begin
+
+lemma less_not_permute: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear)
+
+lemma gather_simps: 
+  shows 
+  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)"
+  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)"
+  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))"
+  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))"  by auto
+
+lemma 
+  gather_start: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
+  by simp
+
+text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
+lemma minf_lt:  "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto
+lemma minf_gt: "\<exists>z . \<forall>x. x < z \<longrightarrow>  (t < x \<longleftrightarrow>  False)"
+  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
+
+lemma minf_le: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le)
+lemma minf_ge: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)"
+  by (auto simp add: less_le not_less not_le)
+lemma minf_eq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
+lemma minf_neq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
+lemma minf_P: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
+
+text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
+lemma pinf_gt:  "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
+lemma pinf_lt: "\<exists>z . \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
+  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
+
+lemma pinf_ge: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le)
+lemma pinf_le: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)"
+  by (auto simp add: less_le not_less not_le)
+lemma pinf_eq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
+lemma pinf_neq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
+lemma pinf_P: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast
+
+lemma nmi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma nmi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)"
+  by (auto simp add: le_less)
+lemma  nmi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x\<le> t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma  nmi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t\<le> x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma  nmi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma  nmi_neq: "t \<in> U \<Longrightarrow>\<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma  nmi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma  nmi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
+  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
+  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma  nmi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
+  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
+  \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+
+lemma  npi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x < t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by (auto simp add: le_less)
+lemma  npi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+lemma  npi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x \<le> t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+lemma  npi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t \<le> x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+lemma  npi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+lemma  npi_neq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u )" by auto
+lemma  npi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+lemma  npi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ;  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
+  \<Longrightarrow>  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+lemma  npi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
+  \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+
+lemma lin_dense_lt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x < t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y < t)"
+proof(clarsimp)
+  fix x l u y  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x"
+    and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  {assume H: "t < y"
+    from less_trans[OF lx px] less_trans[OF H yu]
+    have "l < t \<and> t < u"  by simp
+    with tU noU have "False" by auto}
+  hence "\<not> t < y"  by auto hence "y \<le> t" by (simp add: not_less)
+  thus "y < t" using tny by (simp add: less_le)
+qed
+
+lemma lin_dense_gt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> t < x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t < y)"
+proof(clarsimp)
+  fix x l u y
+  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
+  and px: "t < x" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  {assume H: "y< t"
+    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto}
+  hence "\<not> y<t"  by auto hence "t \<le> y" by (auto simp add: not_less)
+  thus "t < y" using tny by (simp add:less_le)
+qed
+
+lemma lin_dense_le: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<le> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<le> t)"
+proof(clarsimp)
+  fix x l u y
+  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
+  and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  {assume H: "t < y"
+    from less_le_trans[OF lx px] less_trans[OF H yu]
+    have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto}
+  hence "\<not> t < y"  by auto thus "y \<le> t" by (simp add: not_less)
+qed
+
+lemma lin_dense_ge: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> t \<le> x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t \<le> y)"
+proof(clarsimp)
+  fix x l u y
+  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
+  and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  {assume H: "y< t"
+    from less_trans[OF ly H] le_less_trans[OF px xu]
+    have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto}
+  hence "\<not> y<t"  by auto thus "t \<le> y" by (simp add: not_less)
+qed
+lemma lin_dense_eq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x = t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y= t)"  by auto
+lemma lin_dense_neq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<noteq> t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<noteq> t)"  by auto
+lemma lin_dense_P: "\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P)"  by auto
+
+lemma lin_dense_conj:
+  "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
+  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
+  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<and> P2 x)
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))"
+  by blast
+lemma lin_dense_disj:
+  "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
+  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
+  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<or> P2 x)
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))"
+  by blast
+
+lemma npmibnd: "\<lbrakk>\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<le> x); \<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk>
+  \<Longrightarrow> \<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<le> x \<and> x \<le> u')"
+by auto
+
+lemma finite_set_intervals:
+  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
+  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
+  shows "\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x"
+proof-
+  let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
+  let ?xM = "{y. y\<in> S \<and> x \<le> y}"
+  let ?a = "Max ?Mx"
+  let ?b = "Min ?xM"
+  have MxS: "?Mx \<subseteq> S" by blast
+  hence fMx: "finite ?Mx" using fS finite_subset by auto
+  from lx linS have linMx: "l \<in> ?Mx" by blast
+  hence Mxne: "?Mx \<noteq> {}" by blast
+  have xMS: "?xM \<subseteq> S" by blast
+  hence fxM: "finite ?xM" using fS finite_subset by auto
+  from xu uinS have linxM: "u \<in> ?xM" by blast
+  hence xMne: "?xM \<noteq> {}" by blast
+  have ax:"?a \<le> x" using Mxne fMx by auto
+  have xb:"x \<le> ?b" using xMne fxM by auto
+  have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
+  have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
+  have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
+  proof(clarsimp)
+    fix y   assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
+    from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
+    moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
+    moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
+    ultimately show "False" by blast
+  qed
+  from ainS binS noy ax xb px show ?thesis by blast
+qed
+
+lemma finite_set_intervals2:
+  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
+  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
+  shows "(\<exists> s\<in> S. P s) \<or> (\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a < x \<and> x < b \<and> P x)"
+proof-
+  from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
+  obtain a and b where
+    as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
+    and axb: "a \<le> x \<and> x \<le> b \<and> P x"  by auto
+  from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
+  thus ?thesis using px as bs noS by blast
+qed
+
+end
+
+section {* The classical QE after Langford for dense linear orders *}
+
+context dense_linear_order
+begin
+
+lemma interval_empty_iff:
+  "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
+  by (auto dest: dense)
+
+lemma dlo_qe_bnds: 
+  assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite U"
+  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> (\<forall> l \<in> L. \<forall>u \<in> U. l < u)"
+proof (simp only: atomize_eq, rule iffI)
+  assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)"
+  then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
+  {fix l u assume l: "l \<in> L" and u: "u \<in> U"
+    have "l < x" using xL l by blast
+    also have "x < u" using xU u by blast
+    finally (less_trans) have "l < u" .}
+  thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
+next
+  assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u"
+  let ?ML = "Max L"
+  let ?MU = "Min U"  
+  from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" by auto
+  from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" by auto
+  from th1 th2 H have "?ML < ?MU" by auto
+  with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast
+  from th3 th1' have "\<forall>l \<in> L. l < w" by auto
+  moreover from th4 th2' have "\<forall>u \<in> U. w < u" by auto
+  ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" by auto
+qed
+
+lemma dlo_qe_noub: 
+  assumes ne: "L \<noteq> {}" and fL: "finite L"
+  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
+proof(simp add: atomize_eq)
+  from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast
+  from ne fL have "\<forall>x \<in> L. x \<le> Max L" by simp
+  with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans)
+  thus "\<exists>x. \<forall>y\<in>L. y < x" by blast
+qed
+
+lemma dlo_qe_nolb: 
+  assumes ne: "U \<noteq> {}" and fU: "finite U"
+  shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
+proof(simp add: atomize_eq)
+  from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast
+  from ne fU have "\<forall>x \<in> U. Min U \<le> x" by simp
+  with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans)
+  thus "\<exists>x. \<forall>y\<in>U. x < y" by blast
+qed
+
+lemma exists_neq: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x" 
+  using gt_ex[of t] by auto
+
+lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq 
+  le_less neq_iff linear less_not_permute
+
+lemma axiom: "dense_linear_order (op \<le>) (op <)" by (rule dense_linear_order_axioms)
+lemma atoms:
+  shows "TERM (less :: 'a \<Rightarrow> _)"
+    and "TERM (less_eq :: 'a \<Rightarrow> _)"
+    and "TERM (op = :: 'a \<Rightarrow> _)" .
+
+declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms]
+declare dlo_simps[langfordsimp]
+
+end
+
+(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
+lemma dnf:
+  "(P & (Q | R)) = ((P&Q) | (P&R))" 
+  "((Q | R) & P) = ((Q&P) | (R&P))"
+  by blast+
+
+lemmas weak_dnf_simps = simp_thms dnf
+
+lemma nnf_simps:
+    "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
+    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
+  by blast+
+
+lemma ex_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" by blast
+
+lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib
+
+use "~~/src/HOL/Tools/Qelim/langford.ML"
+method_setup dlo = {*
+  Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac)
+*} "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"} *}
+
+text {* Linear order without upper bounds *}
+
+class_locale linorder_stupid_syntax = linorder
+begin
+notation
+  less_eq  ("op \<sqsubseteq>") and
+  less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
+  less  ("op \<sqsubset>") and
+  less  ("(_/ \<sqsubset> _)"  [51, 51] 50)
+
+end
+
+class_locale linorder_no_ub = linorder_stupid_syntax +
+  assumes gt_ex: "\<exists>y. less x y"
+begin
+lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
+
+text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
+lemma pinf_conj:
+  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
+proof-
+  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
+  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
+  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
+  {fix x assume H: "z \<sqsubset> x"
+    from less_trans[OF zz1 H] less_trans[OF zz2 H]
+    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
+  }
+  thus ?thesis by blast
+qed
+
+lemma pinf_disj:
+  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
+proof-
+  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
+  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
+  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
+  {fix x assume H: "z \<sqsubset> x"
+    from less_trans[OF zz1 H] less_trans[OF zz2 H]
+    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
+  }
+  thus ?thesis by blast
+qed
+
+lemma pinf_ex: assumes ex:"\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
+proof-
+  from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
+  from gt_ex obtain x where x: "z \<sqsubset> x" by blast
+  from z x p1 show ?thesis by blast
+qed
+
+end
+
+text {* Linear order without upper bounds *}
+
+class_locale linorder_no_lb = linorder_stupid_syntax +
+  assumes lt_ex: "\<exists>y. less y x"
+begin
+lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
+
+
+text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
+lemma minf_conj:
+  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
+proof-
+  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
+  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
+  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
+  {fix x assume H: "x \<sqsubset> z"
+    from less_trans[OF H zz1] less_trans[OF H zz2]
+    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
+  }
+  thus ?thesis by blast
+qed
+
+lemma minf_disj:
+  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
+proof-
+  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
+  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
+  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
+  {fix x assume H: "x \<sqsubset> z"
+    from less_trans[OF H zz1] less_trans[OF H zz2]
+    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
+  }
+  thus ?thesis by blast
+qed
+
+lemma minf_ex: assumes ex:"\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
+proof-
+  from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
+  from lt_ex obtain x where x: "x \<sqsubset> z" by blast
+  from z x p1 show ?thesis by blast
+qed
+
+end
+
+
+class_locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
+  fixes between
+  assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
+     and  between_same: "between x x = x"
+
+class_interpretation  constr_dense_linear_order < dense_linear_order 
+  apply unfold_locales
+  using gt_ex lt_ex between_less
+    by (auto, rule_tac x="between x y" in exI, simp)
+
+context  constr_dense_linear_order
+begin
+
+lemma rinf_U:
+  assumes fU: "finite U"
+  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
+  \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
+  and nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')"
+  and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  P x"
+  shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
+proof-
+  from ex obtain x where px: "P x" by blast
+  from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto
+  then obtain u and u' where uU:"u\<in> U" and uU': "u' \<in> U" and ux:"u \<sqsubseteq> x" and xu':"x \<sqsubseteq> u'" by auto
+  from uU have Une: "U \<noteq> {}" by auto
+  term "linorder.Min less_eq"
+  let ?l = "linorder.Min less_eq U"
+  let ?u = "linorder.Max less_eq U"
+  have linM: "?l \<in> U" using fU Une by simp
+  have uinM: "?u \<in> U" using fU Une by simp
+  have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU by auto
+  have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU by auto
+  have th:"?l \<sqsubseteq> u" using uU Une lM by auto
+  from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" .
+  have th: "u' \<sqsubseteq> ?u" using uU' Une Mu by simp
+  from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" .
+  from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
+  have "(\<exists> s\<in> U. P s) \<or>
+      (\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x)" .
+  moreover { fix u assume um: "u\<in>U" and pu: "P u"
+    have "between u u = u" by (simp add: between_same)
+    with um pu have "P (between u u)" by simp
+    with um have ?thesis by blast}
+  moreover{
+    assume "\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x"
+      then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
+        and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U" and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x"
+        by blast
+      from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
+      let ?u = "between t1 t2"
+      from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
+      from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
+      with t1M t2M have ?thesis by blast}
+    ultimately show ?thesis by blast
+  qed
+
+theorem fr_eq:
+  assumes fU: "finite U"
+  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
+   \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
+  and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
+  and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
+  and mi: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x = MP)"  and pi: "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x = PP)"
+  shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))"
+  (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
+proof-
+ {
+   assume px: "\<exists> x. P x"
+   have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
+   moreover {assume "MP \<or> PP" hence "?D" by blast}
+   moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
+     from npmibnd[OF nmibnd npibnd]
+     have nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')" .
+     from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
+   ultimately have "?D" by blast}
+ moreover
+ { assume "?D"
+   moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
+   moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
+   moreover {assume f:"?F" hence "?E" by blast}
+   ultimately have "?E" by blast}
+ ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
+qed
+
+lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
+lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P
+
+lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P
+lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P
+lemmas lin_dense_thms = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P
+
+lemma ferrack_axiom: "constr_dense_linear_order less_eq less between"
+  by (rule constr_dense_linear_order_axioms)
+lemma atoms:
+  shows "TERM (less :: 'a \<Rightarrow> _)"
+    and "TERM (less_eq :: 'a \<Rightarrow> _)"
+    and "TERM (op = :: 'a \<Rightarrow> _)" .
+
+declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms
+    nmi: nmi_thms npi: npi_thms lindense:
+    lin_dense_thms qe: fr_eq atoms: atoms]
+
+declaration {*
+let
+fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
+fun generic_whatis phi =
+ let
+  val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
+  fun h x t =
+   case term_of t of
+     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+                            else Ferrante_Rackoff_Data.Nox
+   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+                            else Ferrante_Rackoff_Data.Nox
+   | b$y$z => if Term.could_unify (b, lt) then
+                 if term_of x aconv y then Ferrante_Rackoff_Data.Lt
+                 else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
+                 else Ferrante_Rackoff_Data.Nox
+             else if Term.could_unify (b, le) then
+                 if term_of x aconv y then Ferrante_Rackoff_Data.Le
+                 else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
+                 else Ferrante_Rackoff_Data.Nox
+             else Ferrante_Rackoff_Data.Nox
+   | _ => Ferrante_Rackoff_Data.Nox
+ in h end
+ fun ss phi = HOL_ss addsimps (simps phi)
+in
+ Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
+  {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
+end
+*}
+
+end
+
+use "~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML"
+
+method_setup ferrack = {*
+  Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
+*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
+
+subsection {* Ferrante and Rackoff algorithm over ordered fields *}
+
+lemma neg_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
+proof-
+  assume H: "c < 0"
+  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
+  also have "\<dots> = (0 < x)" by simp
+  finally show  "(c*x < 0) == (x > 0)" by simp
+qed
+
+lemma pos_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
+proof-
+  assume H: "c > 0"
+  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
+  also have "\<dots> = (0 > x)" by simp
+  finally show  "(c*x < 0) == (x < 0)" by simp
+qed
+
+lemma neg_prod_sum_lt: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t< 0) == (x > (- 1/c)*t))"
+proof-
+  assume H: "c < 0"
+  have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
+  also have "\<dots> = ((- 1/c)*t < x)" by simp
+  finally show  "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
+qed
+
+lemma pos_prod_sum_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t < 0) == (x < (- 1/c)*t))"
+proof-
+  assume H: "c > 0"
+  have "c*x + t< 0 = (c*x < -t)"  by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
+  also have "\<dots> = ((- 1/c)*t > x)" by simp
+  finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" by simp
+qed
+
+lemma sum_lt:"((x::'a::pordered_ab_group_add) + t < 0) == (x < - t)"
+  using less_diff_eq[where a= x and b=t and c=0] by simp
+
+lemma neg_prod_le:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
+proof-
+  assume H: "c < 0"
+  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
+  also have "\<dots> = (0 <= x)" by simp
+  finally show  "(c*x <= 0) == (x >= 0)" by simp
+qed
+
+lemma pos_prod_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
+proof-
+  assume H: "c > 0"
+  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
+  also have "\<dots> = (0 >= x)" by simp
+  finally show  "(c*x <= 0) == (x <= 0)" by simp
+qed
+
+lemma neg_prod_sum_le: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t <= 0) == (x >= (- 1/c)*t))"
+proof-
+  assume H: "c < 0"
+  have "c*x + t <= 0 = (c*x <= -t)"  by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
+  also have "\<dots> = ((- 1/c)*t <= x)" by simp
+  finally show  "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
+qed
+
+lemma pos_prod_sum_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t <= 0) == (x <= (- 1/c)*t))"
+proof-
+  assume H: "c > 0"
+  have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
+  also have "\<dots> = ((- 1/c)*t >= x)" by simp
+  finally show  "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
+qed
+
+lemma sum_le:"((x::'a::pordered_ab_group_add) + t <= 0) == (x <= - t)"
+  using le_diff_eq[where a= x and b=t and c=0] by simp
+
+lemma nz_prod_eq:"(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x = 0) == (x = 0))" by simp
+lemma nz_prod_sum_eq: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x + t = 0) == (x = (- 1/c)*t))"
+proof-
+  assume H: "c \<noteq> 0"
+  have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] ring_simps)
+  finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" by simp
+qed
+lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)"
+  using eq_diff_eq[where a= x and b=t and c=0] by simp
+
+
+class_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)
+  fix x y::'a assume lt: "x < y"
+  from  less_half_sum[OF lt] show "x < (x + y) /2" by simp
+next
+  fix x y::'a assume lt: "x < y"
+  from  gt_half_sum[OF lt] show "(x + y) /2 < y" by simp
+qed
+
+declaration{*
+let
+fun earlier [] x y = false
+        | earlier (h::t) x y =
+    if h aconvc y then false else if h aconvc x then true else earlier t x y;
+
+fun dest_frac ct = case term_of ct of
+   Const (@{const_name "HOL.divide"},_) $ a $ b=>
+    Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
+ | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
+
+fun mk_frac phi cT x =
+ let val (a, b) = Rat.quotient_of_rat x
+ in if b = 1 then Numeral.mk_cnumber cT a
+    else Thm.capply
+         (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
+                     (Numeral.mk_cnumber cT a))
+         (Numeral.mk_cnumber cT b)
+ end
+
+fun whatis x ct = case term_of ct of
+  Const(@{const_name "HOL.plus"}, _)$(Const(@{const_name "HOL.times"},_)$_$y)$_ =>
+     if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
+     else ("Nox",[])
+| Const(@{const_name "HOL.plus"}, _)$y$_ =>
+     if y aconv term_of x then ("x+t",[Thm.dest_arg ct])
+     else ("Nox",[])
+| Const(@{const_name "HOL.times"}, _)$_$y =>
+     if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct])
+     else ("Nox",[])
+| t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
+
+fun xnormalize_conv ctxt [] ct = reflexive ct
+| xnormalize_conv ctxt (vs as (x::_)) ct =
+   case term_of ct of
+   Const(@{const_name HOL.less},_)$_$Const(@{const_name "HOL.zero"},_) =>
+    (case whatis x (Thm.dest_arg1 ct) of
+    ("c*x+t",[c,t]) =>
+       let
+        val cr = dest_frac c
+        val clt = Thm.dest_fun2 ct
+        val cz = Thm.dest_arg ct
+        val neg = cr </ Rat.zero
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+               (Thm.capply @{cterm "Trueprop"}
+                  (if neg then Thm.capply (Thm.capply clt c) cz
+                    else Thm.capply (Thm.capply clt cz) c))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
+             (if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+      in rth end
+    | ("x+t",[t]) =>
+       let
+        val T = ctyp_of_term x
+        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+       in  rth end
+    | ("c*x",[c]) =>
+       let
+        val cr = dest_frac c
+        val clt = Thm.dest_fun2 ct
+        val cz = Thm.dest_arg ct
+        val neg = cr </ Rat.zero
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+               (Thm.capply @{cterm "Trueprop"}
+                  (if neg then Thm.capply (Thm.capply clt c) cz
+                    else Thm.capply (Thm.capply clt cz) c))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
+             (if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
+        val rth = th
+      in rth end
+    | _ => reflexive ct)
+
+
+|  Const(@{const_name HOL.less_eq},_)$_$Const(@{const_name "HOL.zero"},_) =>
+   (case whatis x (Thm.dest_arg1 ct) of
+    ("c*x+t",[c,t]) =>
+       let
+        val T = ctyp_of_term x
+        val cr = dest_frac c
+        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
+        val cz = Thm.dest_arg ct
+        val neg = cr </ Rat.zero
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+               (Thm.capply @{cterm "Trueprop"}
+                  (if neg then Thm.capply (Thm.capply clt c) cz
+                    else Thm.capply (Thm.capply clt cz) c))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim (instantiate' [SOME T] (map SOME [c,x,t])
+             (if neg then @{thm neg_prod_sum_le} else @{thm pos_prod_sum_le})) cth
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+      in rth end
+    | ("x+t",[t]) =>
+       let
+        val T = ctyp_of_term x
+        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"}
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+       in  rth end
+    | ("c*x",[c]) =>
+       let
+        val T = ctyp_of_term x
+        val cr = dest_frac c
+        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
+        val cz = Thm.dest_arg ct
+        val neg = cr </ Rat.zero
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+               (Thm.capply @{cterm "Trueprop"}
+                  (if neg then Thm.capply (Thm.capply clt c) cz
+                    else Thm.capply (Thm.capply clt cz) c))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
+             (if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
+        val rth = th
+      in rth end
+    | _ => reflexive ct)
+
+|  Const("op =",_)$_$Const(@{const_name "HOL.zero"},_) =>
+   (case whatis x (Thm.dest_arg1 ct) of
+    ("c*x+t",[c,t]) =>
+       let
+        val T = ctyp_of_term x
+        val cr = dest_frac c
+        val ceq = Thm.dest_fun2 ct
+        val cz = Thm.dest_arg ct
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+            (Thm.capply @{cterm "Trueprop"}
+             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim
+                 (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+      in rth end
+    | ("x+t",[t]) =>
+       let
+        val T = ctyp_of_term x
+        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+       in  rth end
+    | ("c*x",[c]) =>
+       let
+        val T = ctyp_of_term x
+        val cr = dest_frac c
+        val ceq = Thm.dest_fun2 ct
+        val cz = Thm.dest_arg ct
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+            (Thm.capply @{cterm "Trueprop"}
+             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val rth = implies_elim
+                 (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth
+      in rth end
+    | _ => reflexive ct);
+
+local
+  val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"}
+  val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"}
+  val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
+in
+fun field_isolate_conv phi ctxt vs ct = case term_of ct of
+  Const(@{const_name HOL.less},_)$a$b =>
+   let val (ca,cb) = Thm.dest_binop ct
+       val T = ctyp_of_term ca
+       val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0
+       val nth = Conv.fconv_rule
+         (Conv.arg_conv (Conv.arg1_conv
+              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
+       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+   in rth end
+| Const(@{const_name HOL.less_eq},_)$a$b =>
+   let val (ca,cb) = Thm.dest_binop ct
+       val T = ctyp_of_term ca
+       val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0
+       val nth = Conv.fconv_rule
+         (Conv.arg_conv (Conv.arg1_conv
+              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
+       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+   in rth end
+
+| Const("op =",_)$a$b =>
+   let val (ca,cb) = Thm.dest_binop ct
+       val T = ctyp_of_term ca
+       val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
+       val nth = Conv.fconv_rule
+         (Conv.arg_conv (Conv.arg1_conv
+              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
+       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+   in rth end
+| @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
+| _ => reflexive ct
+end;
+
+fun classfield_whatis phi =
+ let
+  fun h x t =
+   case term_of t of
+     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+                            else Ferrante_Rackoff_Data.Nox
+   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+                            else Ferrante_Rackoff_Data.Nox
+   | Const(@{const_name HOL.less},_)$y$z =>
+       if term_of x aconv y then Ferrante_Rackoff_Data.Lt
+        else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
+        else Ferrante_Rackoff_Data.Nox
+   | Const (@{const_name HOL.less_eq},_)$y$z =>
+         if term_of x aconv y then Ferrante_Rackoff_Data.Le
+         else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
+         else Ferrante_Rackoff_Data.Nox
+   | _ => Ferrante_Rackoff_Data.Nox
+ in h end;
+fun class_field_ss phi =
+   HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}])
+   addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}]
+
+in
+Ferrante_Rackoff_Data.funs @{thm "class_ordered_field_dense_linear_order.ferrack_axiom"}
+  {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
+end
+*}
+
+
+end 
--- a/src/HOL/Divides.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Divides.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -639,7 +639,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"]
+class_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 dvd_diff: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)"
--- a/src/HOL/Finite_Set.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Finite_Set.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -750,7 +750,7 @@
 assumes "finite A" and "a \<notin> A"
 shows "fold_image times g z (insert a A) = g a * (fold_image times g z A)"
 proof -
-  interpret I: fun_left_comm ["%x y. (g x) * y"]
+  interpret I: fun_left_comm "%x y. (g x) * y"
     by unfold_locales (simp add: mult_ac)
   show ?thesis using assms by(simp add:fold_image_def I.fold_insert)
 qed
@@ -798,7 +798,7 @@
     and hyp: "\<And>x y. h (g x y) = times x (h y)"
   shows "h (fold g j w A) = fold times j (h w) A"
 proof -
-  interpret ab_semigroup_mult [g] by fact
+  class_interpret ab_semigroup_mult [g] by fact
   show ?thesis using fin hyp by (induct set: finite) simp_all
 qed
 *)
@@ -873,7 +873,7 @@
 
 subsection {* Generalized summation over a set *}
 
-interpretation comm_monoid_add: comm_monoid_mult ["0::'a::comm_monoid_add" "op +"]
+class_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"
@@ -1760,7 +1760,7 @@
 proof (induct rule: finite_induct)
   case empty then show ?case by simp
 next
-  interpret ab_semigroup_mult ["op Un"]
+  class_interpret ab_semigroup_mult ["op Un"]
     proof qed auto
   case insert 
   then show ?case by simp
@@ -1943,7 +1943,7 @@
 assumes fold: "fold_graph times (b::'a) A y" and "b \<notin> A"
 shows "fold_graph times z (insert b A) (z * y)"
 proof -
-  interpret fun_left_comm ["op *::'a \<Rightarrow> 'a \<Rightarrow> 'a"] by (rule fun_left_comm)
+  interpret fun_left_comm "op *::'a \<Rightarrow> 'a \<Rightarrow> 'a" by (rule fun_left_comm)
 from assms show ?thesis
 proof (induct rule: fold_graph.induct)
   case emptyI thus ?case by (force simp add: fold_insert_aux mult_commute)
@@ -1983,7 +1983,7 @@
 lemma fold1_eq_fold:
 assumes "finite A" "a \<notin> A" shows "fold1 times (insert a A) = fold times a A"
 proof -
-  interpret fun_left_comm ["op *::'a \<Rightarrow> 'a \<Rightarrow> 'a"] by (rule fun_left_comm)
+  interpret fun_left_comm "op *::'a \<Rightarrow> 'a \<Rightarrow> 'a" by (rule fun_left_comm)
   from assms show ?thesis
 apply (simp add: fold1_def fold_def)
 apply (rule the_equality)
@@ -2010,7 +2010,7 @@
   assumes nonempty: "A \<noteq> {}" and A: "finite A" "x \<notin> A"
   shows "fold1 times (insert x A) = x * fold1 times A"
 proof -
-  interpret fun_left_comm ["op *::'a \<Rightarrow> 'a \<Rightarrow> 'a"] by (rule fun_left_comm)
+  interpret fun_left_comm "op *::'a \<Rightarrow> 'a \<Rightarrow> 'a" by (rule fun_left_comm)
   from nonempty obtain a A' where "A = insert a A' & a ~: A'"
     by (auto simp add: nonempty_iff)
   with A show ?thesis
@@ -2033,7 +2033,7 @@
   assumes nonempty: "A \<noteq> {}" and A: "finite A" 
   shows "fold1 times (insert x A) = x * fold1 times A"
 proof -
-  interpret fun_left_comm_idem ["op *::'a \<Rightarrow> 'a \<Rightarrow> 'a"]
+  interpret fun_left_comm_idem "op *::'a \<Rightarrow> 'a \<Rightarrow> 'a"
     by (rule fun_left_comm_idem)
   from nonempty obtain a A' where A': "A = insert a A' & a ~: A'"
     by (auto simp add: nonempty_iff)
@@ -2198,7 +2198,7 @@
   assumes "finite A" "A \<noteq> {}"
   shows "x \<le> fold1 inf A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
 proof -
-  interpret ab_semigroup_idem_mult [inf]
+  class_interpret ab_semigroup_idem_mult [inf]
     by (rule ab_semigroup_idem_mult_inf)
   show ?thesis using assms by (induct rule: finite_ne_induct) simp_all
 qed
@@ -2213,7 +2213,7 @@
   proof (induct rule: finite_ne_induct)
     case singleton thus ?case by simp
   next
-    interpret ab_semigroup_idem_mult [inf]
+    class_interpret ab_semigroup_idem_mult [inf]
       by (rule ab_semigroup_idem_mult_inf)
     case (insert x F)
     from insert(5) have "a = x \<or> a \<in> F" by simp
@@ -2288,7 +2288,7 @@
     and "A \<noteq> {}"
   shows "sup x (\<Sqinter>\<^bsub>fin\<^esub>A) = \<Sqinter>\<^bsub>fin\<^esub>{sup x a|a. a \<in> A}"
 proof -
-  interpret ab_semigroup_idem_mult [inf]
+  class_interpret ab_semigroup_idem_mult [inf]
     by (rule ab_semigroup_idem_mult_inf)
   from assms show ?thesis
     by (simp add: Inf_fin_def image_def
@@ -2303,7 +2303,7 @@
   case singleton thus ?case
     by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def])
 next
-  interpret ab_semigroup_idem_mult [inf]
+  class_interpret ab_semigroup_idem_mult [inf]
     by (rule ab_semigroup_idem_mult_inf)
   case (insert x A)
   have finB: "finite {sup x b |b. b \<in> B}"
@@ -2333,7 +2333,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "inf x (\<Squnion>\<^bsub>fin\<^esub>A) = \<Squnion>\<^bsub>fin\<^esub>{inf x a|a. a \<in> A}"
 proof -
-  interpret ab_semigroup_idem_mult [sup]
+  class_interpret ab_semigroup_idem_mult [sup]
     by (rule ab_semigroup_idem_mult_sup)
   from assms show ?thesis
     by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1])
@@ -2357,7 +2357,7 @@
     thus ?thesis by(simp add: insert(1) B(1))
   qed
   have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
-  interpret ab_semigroup_idem_mult [sup]
+  class_interpret ab_semigroup_idem_mult [sup]
     by (rule ab_semigroup_idem_mult_sup)
   have "inf (\<Squnion>\<^bsub>fin\<^esub>(insert x A)) (\<Squnion>\<^bsub>fin\<^esub>B) = inf (sup x (\<Squnion>\<^bsub>fin\<^esub>A)) (\<Squnion>\<^bsub>fin\<^esub>B)"
     using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def])
@@ -2386,7 +2386,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "\<Sqinter>\<^bsub>fin\<^esub>A = Inf A"
 proof -
-  interpret ab_semigroup_idem_mult [inf]
+  class_interpret ab_semigroup_idem_mult [inf]
     by (rule ab_semigroup_idem_mult_inf)
   from assms show ?thesis
   unfolding Inf_fin_def by (induct A set: finite)
@@ -2397,7 +2397,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "\<Squnion>\<^bsub>fin\<^esub>A = Sup A"
 proof -
-  interpret ab_semigroup_idem_mult [sup]
+  class_interpret ab_semigroup_idem_mult [sup]
     by (rule ab_semigroup_idem_mult_sup)
   from assms show ?thesis
   unfolding Sup_fin_def by (induct A set: finite)
@@ -2446,7 +2446,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "x < fold1 min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
 proof -
-  interpret ab_semigroup_idem_mult [min]
+  class_interpret ab_semigroup_idem_mult [min]
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
   by (induct rule: finite_ne_induct)
@@ -2457,7 +2457,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "fold1 min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
 proof -
-  interpret ab_semigroup_idem_mult [min]
+  class_interpret ab_semigroup_idem_mult [min]
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
   by (induct rule: finite_ne_induct)
@@ -2468,7 +2468,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "fold1 min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
 proof -
-  interpret ab_semigroup_idem_mult [min]
+  class_interpret ab_semigroup_idem_mult [min]
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
   by (induct rule: finite_ne_induct)
@@ -2481,7 +2481,7 @@
 proof cases
   assume "A = B" thus ?thesis by simp
 next
-  interpret ab_semigroup_idem_mult [min]
+  class_interpret ab_semigroup_idem_mult [min]
     by (rule ab_semigroup_idem_mult_min)
   assume "A \<noteq> B"
   have B: "B = A \<union> (B-A)" using `A \<subseteq> B` by blast
@@ -2515,7 +2515,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Min (insert x A) = min x (Min A)"
 proof -
-  interpret ab_semigroup_idem_mult [min]
+  class_interpret ab_semigroup_idem_mult [min]
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis by (rule fold1_insert_idem_def [OF Min_def])
 qed
@@ -2524,7 +2524,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Max (insert x A) = max x (Max A)"
 proof -
-  interpret ab_semigroup_idem_mult [max]
+  class_interpret ab_semigroup_idem_mult [max]
     by (rule ab_semigroup_idem_mult_max)
   from assms show ?thesis by (rule fold1_insert_idem_def [OF Max_def])
 qed
@@ -2533,7 +2533,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Min A \<in> A"
 proof -
-  interpret ab_semigroup_idem_mult [min]
+  class_interpret ab_semigroup_idem_mult [min]
     by (rule ab_semigroup_idem_mult_min)
   from assms fold1_in show ?thesis by (fastsimp simp: Min_def min_def)
 qed
@@ -2542,7 +2542,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Max A \<in> A"
 proof -
-  interpret ab_semigroup_idem_mult [max]
+  class_interpret ab_semigroup_idem_mult [max]
     by (rule ab_semigroup_idem_mult_max)
   from assms fold1_in [of A] show ?thesis by (fastsimp simp: Max_def max_def)
 qed
@@ -2551,7 +2551,7 @@
   assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
   shows "Min (A \<union> B) = min (Min A) (Min B)"
 proof -
-  interpret ab_semigroup_idem_mult [min]
+  class_interpret ab_semigroup_idem_mult [min]
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
     by (simp add: Min_def fold1_Un2)
@@ -2561,7 +2561,7 @@
   assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
   shows "Max (A \<union> B) = max (Max A) (Max B)"
 proof -
-  interpret ab_semigroup_idem_mult [max]
+  class_interpret ab_semigroup_idem_mult [max]
     by (rule ab_semigroup_idem_mult_max)
   from assms show ?thesis
     by (simp add: Max_def fold1_Un2)
@@ -2572,7 +2572,7 @@
     and "finite N" and "N \<noteq> {}"
   shows "h (Min N) = Min (h ` N)"
 proof -
-  interpret ab_semigroup_idem_mult [min]
+  class_interpret ab_semigroup_idem_mult [min]
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
     by (simp add: Min_def hom_fold1_commute)
@@ -2583,7 +2583,7 @@
     and "finite N" and "N \<noteq> {}"
   shows "h (Max N) = Max (h ` N)"
 proof -
-  interpret ab_semigroup_idem_mult [max]
+  class_interpret ab_semigroup_idem_mult [max]
     by (rule ab_semigroup_idem_mult_max)
   from assms show ?thesis
     by (simp add: Max_def hom_fold1_commute [of h])
@@ -2593,7 +2593,7 @@
   assumes "finite A" and "x \<in> A"
   shows "Min A \<le> x"
 proof -
-  interpret lower_semilattice ["op \<le>" "op <" min]
+  class_interpret lower_semilattice ["op \<le>" "op <" min]
     by (rule min_lattice)
   from assms show ?thesis by (simp add: Min_def fold1_belowI)
 qed
@@ -2611,7 +2611,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
 proof -
-  interpret lower_semilattice ["op \<le>" "op <" min]
+  class_interpret lower_semilattice ["op \<le>" "op <" min]
     by (rule min_lattice)
   from assms show ?thesis by (simp add: Min_def below_fold1_iff)
 qed
@@ -2629,7 +2629,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
 proof -
-  interpret lower_semilattice ["op \<le>" "op <" min]
+  class_interpret lower_semilattice ["op \<le>" "op <" min]
     by (rule min_lattice)
   from assms show ?thesis by (simp add: Min_def strict_below_fold1_iff)
 qed
@@ -2639,7 +2639,7 @@
   shows "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
 proof -
   note Max = Max_def
-  interpret linorder ["op \<ge>" "op >"]
+  class_interpret linorder ["op \<ge>" "op >"]
     by (rule dual_linorder)
   from assms show ?thesis
     by (simp add: Max strict_below_fold1_iff [folded dual_max])
@@ -2649,7 +2649,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
 proof -
-  interpret lower_semilattice ["op \<le>" "op <" min]
+  class_interpret lower_semilattice ["op \<le>" "op <" min]
     by (rule min_lattice)
   from assms show ?thesis
     by (simp add: Min_def fold1_below_iff)
@@ -2660,7 +2660,7 @@
   shows "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
 proof -
   note Max = Max_def
-  interpret linorder ["op \<ge>" "op >"]
+  class_interpret linorder ["op \<ge>" "op >"]
     by (rule dual_linorder)
   from assms show ?thesis
     by (simp add: Max fold1_below_iff [folded dual_max])
@@ -2670,7 +2670,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
 proof -
-  interpret lower_semilattice ["op \<le>" "op <" min]
+  class_interpret lower_semilattice ["op \<le>" "op <" min]
     by (rule min_lattice)
   from assms show ?thesis
     by (simp add: Min_def fold1_strict_below_iff)
@@ -2681,7 +2681,7 @@
   shows "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
 proof -
   note Max = Max_def
-  interpret linorder ["op \<ge>" "op >"]
+  class_interpret linorder ["op \<ge>" "op >"]
     by (rule dual_linorder)
   from assms show ?thesis
     by (simp add: Max fold1_strict_below_iff [folded dual_max])
@@ -2691,7 +2691,7 @@
   assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
   shows "Min N \<le> Min M"
 proof -
-  interpret distrib_lattice ["op \<le>" "op <" min max]
+  class_interpret distrib_lattice ["op \<le>" "op <" min max]
     by (rule distrib_lattice_min_max)
   from assms show ?thesis by (simp add: Min_def fold1_antimono)
 qed
@@ -2701,7 +2701,7 @@
   shows "Max M \<le> Max N"
 proof -
   note Max = Max_def
-  interpret linorder ["op \<ge>" "op >"]
+  class_interpret linorder ["op \<ge>" "op >"]
     by (rule dual_linorder)
   from assms show ?thesis
     by (simp add: Max fold1_antimono [folded dual_max])
--- a/src/HOL/FrechetDeriv.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/FrechetDeriv.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -65,8 +65,8 @@
   assumes "bounded_linear g"
   shows "bounded_linear (\<lambda>x. f x + g x)"
 proof -
-  interpret f: bounded_linear [f] by fact
-  interpret g: bounded_linear [g] by fact
+  interpret f: bounded_linear f by fact
+  interpret g: bounded_linear g by fact
   show ?thesis apply (unfold_locales)
     apply (simp only: f.add g.add add_ac)
     apply (simp only: f.scaleR g.scaleR scaleR_right_distrib)
@@ -124,7 +124,7 @@
   assumes "bounded_linear f"
   shows "bounded_linear (\<lambda>x. - f x)"
 proof -
-  interpret f: bounded_linear [f] by fact
+  interpret f: bounded_linear f by fact
   show ?thesis apply (unfold_locales)
     apply (simp add: f.add)
     apply (simp add: f.scaleR)
@@ -151,7 +151,7 @@
   assumes f: "FDERIV f x :> F"
   shows "isCont f x"
 proof -
-  from f interpret F: bounded_linear ["F"] by (rule FDERIV_bounded_linear)
+  from f interpret F: bounded_linear "F" by (rule FDERIV_bounded_linear)
   have "(\<lambda>h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0"
     by (rule FDERIV_D [OF f])
   hence "(\<lambda>h. norm (f (x + h) - f x - F h) / norm h * norm h) -- 0 --> 0"
@@ -180,8 +180,8 @@
   assumes "bounded_linear g"
   shows "bounded_linear (\<lambda>x. f (g x))"
 proof -
-  interpret f: bounded_linear [f] by fact
-  interpret g: bounded_linear [g] by fact
+  interpret f: bounded_linear f by fact
+  interpret g: bounded_linear g by fact
   show ?thesis proof (unfold_locales)
     fix x y show "f (g (x + y)) = f (g x) + f (g y)"
       by (simp only: f.add g.add)
@@ -223,8 +223,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
 
@@ -375,9 +375,9 @@
     by (simp only: FDERIV_lemma)
 qed
 
-lemmas FDERIV_mult = bounded_bilinear_locale.mult.prod.FDERIV
+lemmas FDERIV_mult = mult.FDERIV
 
-lemmas FDERIV_scaleR = bounded_bilinear_locale.scaleR.prod.FDERIV
+lemmas FDERIV_scaleR = scaleR.FDERIV
 
 
 subsection {* Powers *}
@@ -409,10 +409,10 @@
 by (simp add: right_diff_distrib left_diff_distrib mult_assoc)
 
 lemmas bounded_linear_mult_const =
-  bounded_bilinear_locale.mult.prod.bounded_linear_left [THEN bounded_linear_compose]
+  mult.bounded_linear_left [THEN bounded_linear_compose]
 
 lemmas bounded_linear_const_mult =
-  bounded_bilinear_locale.mult.prod.bounded_linear_right [THEN bounded_linear_compose]
+  mult.bounded_linear_right [THEN bounded_linear_compose]
 
 lemma FDERIV_inverse:
   fixes x :: "'a::real_normed_div_algebra"
@@ -492,7 +492,7 @@
   fixes x :: "'a::real_normed_field" shows
   "FDERIV f x :> (\<lambda>h. h * D) = (\<lambda>h. (f (x + h) - f x) / h) -- 0 --> D"
  apply (unfold fderiv_def)
- apply (simp add: bounded_bilinear_locale.mult.prod.bounded_linear_left)
+ apply (simp add: mult.bounded_linear_left)
  apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric])
  apply (subst diff_divide_distrib)
  apply (subst times_divide_eq_left [symmetric])
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Fundamental_Theorem_Algebra.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,1327 @@
+(* Author: Amine Chaieb, TU Muenchen *)
+
+header{*Fundamental Theorem of Algebra*}
+
+theory Fundamental_Theorem_Algebra
+imports Univ_Poly Dense_Linear_Order Complex
+begin
+
+subsection {* Square root of complex numbers *}
+definition csqrt :: "complex \<Rightarrow> complex" where
+"csqrt z = (if Im z = 0 then
+            if 0 \<le> Re z then Complex (sqrt(Re z)) 0
+            else Complex 0 (sqrt(- Re z))
+           else Complex (sqrt((cmod z + Re z) /2))
+                        ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
+
+lemma csqrt[algebra]: "csqrt z ^ 2 = z"
+proof-
+  obtain x y where xy: "z = Complex x y" by (cases z)
+  {assume y0: "y = 0"
+    {assume x0: "x \<ge> 0" 
+      then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
+	by (simp add: csqrt_def power2_eq_square)}
+    moreover
+    {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
+      then have ?thesis using y0 xy real_sqrt_pow2[OF x0] 
+	by (simp add: csqrt_def power2_eq_square) }
+    ultimately have ?thesis by blast}
+  moreover
+  {assume y0: "y\<noteq>0"
+    {fix x y
+      let ?z = "Complex x y"
+      from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
+      hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+ 
+      hence "(sqrt (x * x + y * y) + x) / 2 \<ge> 0" "(sqrt (x * x + y * y) - x) / 2 \<ge> 0" by (simp_all add: power2_eq_square) }
+    note th = this
+    have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2" 
+      by (simp add: power2_eq_square) 
+    from th[of x y]
+    have sq4': "sqrt (((sqrt (x * x + y * y) + x)^2 / 4)) = (sqrt (x * x + y * y) + x) / 2" "sqrt (((sqrt (x * x + y * y) - x)^2 / 4)) = (sqrt (x * x + y * y) - x) / 2" unfolding sq4 by simp_all
+    then have th1: "sqrt ((sqrt (x * x + y * y) + x) * (sqrt (x * x + y * y) + x) / 4) - sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) - x) / 4) = x"
+      unfolding power2_eq_square by simp 
+    have "sqrt 4 = sqrt (2^2)" by simp 
+    hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
+    have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
+      using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
+      unfolding power2_eq_square 
+      by (simp add: ring_simps real_sqrt_divide sqrt4)
+     from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
+       apply (simp add: real_sqrt_sum_squares_mult_ge_zero[of x y] real_sqrt_pow2[OF th(1)[of x y], unfolded power2_eq_square] real_sqrt_pow2[OF th(2)[of x y], unfolded power2_eq_square] real_sqrt_mult[symmetric])
+      using th1 th2  ..}
+  ultimately show ?thesis by blast
+qed
+
+
+subsection{* More lemmas about module of complex numbers *}
+
+lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
+  by (rule of_real_power [symmetric])
+
+lemma real_down2: "(0::real) < d1 \<Longrightarrow> 0 < d2 ==> EX e. 0 < e & e < d1 & e < d2"
+  apply ferrack apply arith done
+
+text{* The triangle inequality for cmod *}
+lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
+  using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
+
+subsection{* Basic lemmas about complex polynomials *}
+
+lemma poly_bound_exists:
+  shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
+proof(induct p)
+  case Nil thus ?case by (rule exI[where x=1], simp) 
+next
+  case (Cons c cs)
+  from Cons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
+    by blast
+  let ?k = " 1 + cmod c + \<bar>r * m\<bar>"
+  have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
+  {fix z
+    assume H: "cmod z \<le> r"
+    from m H have th: "cmod (poly cs z) \<le> m" by blast
+    from H have rp: "r \<ge> 0" using norm_ge_zero[of z] by arith
+    have "cmod (poly (c # cs) z) \<le> cmod c + cmod (z* poly cs z)"
+      using norm_triangle_ineq[of c "z* poly cs z"] by simp
+    also have "\<dots> \<le> cmod c + r*m" using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]] by (simp add: norm_mult)
+    also have "\<dots> \<le> ?k" by simp
+    finally have "cmod (poly (c # cs) z) \<le> ?k" .}
+  with kp show ?case by blast
+qed
+
+
+text{* Offsetting the variable in a polynomial gives another of same degree *}
+  (* FIXME : Lemma holds also in locale --- fix it later *)
+lemma  poly_offset_lemma:
+  shows "\<exists>b q. (length q = length p) \<and> (\<forall>x. poly (b#q) (x::complex) = (a + x) * poly p x)"
+proof(induct p)
+  case Nil thus ?case by simp
+next
+  case (Cons c cs)
+  from Cons.hyps obtain b q where 
+    bq: "length q = length cs" "\<forall>x. poly (b # q) x = (a + x) * poly cs x"
+    by blast
+  let ?b = "a*c"
+  let ?q = "(b+c)#q"
+  have lg: "length ?q = length (c#cs)" using bq(1) by simp
+  {fix x
+    from bq(2)[rule_format, of x]
+    have "x*poly (b # q) x = x*((a + x) * poly cs x)" by simp
+    hence "poly (?b# ?q) x = (a + x) * poly (c # cs) x"
+      by (simp add: ring_simps)}
+  with lg  show ?case by blast 
+qed
+
+    (* FIXME : This one too*)
+lemma poly_offset: "\<exists> q. length q = length p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
+proof (induct p)
+  case Nil thus ?case by simp
+next
+  case (Cons c cs)
+  from Cons.hyps obtain q where q: "length q = length cs" "\<forall>x. poly q x = poly cs (a + x)" by blast
+  from poly_offset_lemma[of q a] obtain b p where 
+    bp: "length p = length q" "\<forall>x. poly (b # p) x = (a + x) * poly q x"
+    by blast
+  thus ?case using q bp by - (rule exI[where x="(c + b)#p"], simp)
+qed
+
+text{* An alternative useful formulation of completeness of the reals *}
+lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
+  shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
+proof-
+  from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
+  from ex have thx:"\<exists>x. x \<in> Collect P" by blast
+  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y" 
+    by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
+  from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
+    by blast
+  from Y[OF x] have xY: "x < Y" .
+  from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)  
+  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y" 
+    apply (clarsimp, atomize (full)) by auto 
+  from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
+  {fix y
+    {fix z assume z: "P z" "y < z"
+      from L' z have "y < L" by auto }
+    moreover
+    {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
+      hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
+      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def) 
+      with yL(1) have False  by arith}
+    ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
+  thus ?thesis by blast
+qed
+
+
+subsection{* Some theorems about Sequences*}
+text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
+
+lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
+  unfolding Ex1_def
+  apply (rule_tac x="nat_rec e f" in exI)
+  apply (rule conjI)+
+apply (rule def_nat_rec_0, simp)
+apply (rule allI, rule def_nat_rec_Suc, simp)
+apply (rule allI, rule impI, rule ext)
+apply (erule conjE)
+apply (induct_tac x)
+apply (simp add: nat_rec_0)
+apply (erule_tac x="n" in allE)
+apply (simp)
+done
+
+ text{* An equivalent formulation of monotony -- Not used here, but might be useful *}
+lemma mono_Suc: "mono f = (\<forall>n. (f n :: 'a :: order) \<le> f (Suc n))"
+unfolding mono_def
+proof auto
+  fix A B :: nat
+  assume H: "\<forall>n. f n \<le> f (Suc n)" "A \<le> B"
+  hence "\<exists>k. B = A + k" apply -  apply (thin_tac "\<forall>n. f n \<le> f (Suc n)") 
+    by presburger
+  then obtain k where k: "B = A + k" by blast
+  {fix a k
+    have "f a \<le> f (a + k)"
+    proof (induct k)
+      case 0 thus ?case by simp
+    next
+      case (Suc k)
+      from Suc.hyps H(1)[rule_format, of "a + k"] show ?case by simp
+    qed}
+  with k show "f A \<le> f B" by blast
+qed
+
+text{* for any sequence, there is a mootonic subsequence *}
+lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
+proof-
+  {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
+    let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
+    from num_Axiom[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
+    obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
+    have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
+      using H apply - 
+      apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
+      unfolding order_le_less by blast 
+    hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
+    {fix n
+      have "?P (f (Suc n)) (f n)" 
+	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
+	using H apply - 
+      apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
+      unfolding order_le_less by blast 
+    hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
+  note fSuc = this
+    {fix p q assume pq: "p \<ge> f q"
+      have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
+	by (cases q, simp_all) }
+    note pqth = this
+    {fix q
+      have "f (Suc q) > f q" apply (induct q) 
+	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
+    note fss = this
+    from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
+    {fix a b 
+      have "f a \<le> f (a + b)"
+      proof(induct b)
+	case 0 thus ?case by simp
+      next
+	case (Suc b)
+	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
+      qed}
+    note fmon0 = this
+    have "monoseq (\<lambda>n. s (f n))" 
+    proof-
+      {fix n
+	have "s (f n) \<ge> s (f (Suc n))" 
+	proof(cases n)
+	  case 0
+	  assume n0: "n = 0"
+	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
+	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
+	next
+	  case (Suc m)
+	  assume m: "n = Suc m"
+	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
+	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
+	qed}
+      thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
+    qed
+    with th1 have ?thesis by blast}
+  moreover
+  {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
+    {fix p assume p: "p \<ge> Suc N" 
+      hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
+      have "m \<noteq> p" using m(2) by auto 
+      with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
+    note th0 = this
+    let ?P = "\<lambda>m x. m > x \<and> s x < s m"
+    from num_Axiom[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
+    obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
+      "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
+    have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
+      using N apply - 
+      apply (erule allE[where x="Suc N"], clarsimp)
+      apply (rule_tac x="m" in exI)
+      apply auto
+      apply (subgoal_tac "Suc N \<noteq> m")
+      apply simp
+      apply (rule ccontr, simp)
+      done
+    hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
+    {fix n
+      have "f n > N \<and> ?P (f (Suc n)) (f n)"
+	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
+      proof (induct n)
+	case 0 thus ?case
+	  using f0 N apply auto 
+	  apply (erule allE[where x="f 0"], clarsimp) 
+	  apply (rule_tac x="m" in exI, simp)
+	  by (subgoal_tac "f 0 \<noteq> m", auto)
+      next
+	case (Suc n)
+	from Suc.hyps have Nfn: "N < f n" by blast
+	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
+	with Nfn have mN: "m > N" by arith
+	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
+	
+	from key have th0: "f (Suc n) > N" by simp
+	from N[rule_format, OF th0]
+	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
+	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
+	hence "m' > f (Suc n)" using m'(1) by simp
+	with key m'(2) show ?case by auto
+      qed}
+    note fSuc = this
+    {fix n
+      have "f n \<ge> Suc N \<and> f(Suc n) > f n \<and> s(f n) < s(f(Suc n))" using fSuc[of n] by auto 
+      hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
+    note thf = this
+    have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
+    have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
+      apply -
+      apply (rule disjI1)
+      apply auto
+      apply (rule order_less_imp_le)
+      apply blast
+      done
+    then have ?thesis  using sqf by blast}
+  ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
+qed
+
+lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
+proof(induct n)
+  case 0 thus ?case by simp
+next
+  case (Suc n)
+  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
+  have "n < f (Suc n)" by arith 
+  thus ?case by arith
+qed
+
+subsection {* Fundamental theorem of algebra *}
+lemma  unimodular_reduce_norm:
+  assumes md: "cmod z = 1"
+  shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
+proof-
+  obtain x y where z: "z = Complex x y " by (cases z, auto)
+  from md z have xy: "x^2 + y^2 = 1" by (simp add: cmod_def)
+  {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
+    from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
+      by (simp_all add: cmod_def power2_eq_square ring_simps)
+    hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
+    hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
+      by - (rule power_mono, simp, simp)+
+    hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1" 
+      by (simp_all  add: power2_abs power_mult_distrib)
+    from add_mono[OF th0] xy have False by simp }
+  thus ?thesis unfolding linorder_not_le[symmetric] by blast
+qed
+
+text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
+lemma reduce_poly_simple:
+ assumes b: "b \<noteq> 0" and n: "n\<noteq>0"
+  shows "\<exists>z. cmod (1 + b * z^n) < 1"
+using n
+proof(induct n rule: nat_less_induct)
+  fix n
+  assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" and n: "n \<noteq> 0"
+  let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
+  {assume e: "even n"
+    hence "\<exists>m. n = 2*m" by presburger
+    then obtain m where m: "n = 2*m" by blast
+    from n m have "m\<noteq>0" "m < n" by presburger+
+    with IH[rule_format, of m] obtain z where z: "?P z m" by blast
+    from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
+    hence "\<exists>z. ?P z n" ..}
+  moreover
+  {assume o: "odd n"
+    from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
+    have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
+    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) = 
+    ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
+    also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2" 
+      apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
+      by (simp add: power2_eq_square)
+    finally 
+    have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
+    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
+    1" 
+      apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
+      using right_inverse[OF b']
+      by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] ring_simps)
+    have th0: "cmod (complex_of_real (cmod b) / b) = 1"
+      apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse ring_simps )
+      by (simp add: real_sqrt_mult[symmetric] th0)        
+    from o have "\<exists>m. n = Suc (2*m)" by presburger+
+    then obtain m where m: "n = Suc (2*m)" by blast
+    from unimodular_reduce_norm[OF th0] o
+    have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
+      apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
+      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_def)
+      apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
+      apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
+      apply (rule_tac x="- ii" in exI, simp add: m power_mult)
+      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_def)
+      apply (rule_tac x="ii" in exI, simp add: m power_mult diff_def)
+      done
+    then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
+    let ?w = "v / complex_of_real (root n (cmod b))"
+    from odd_real_root_pow[OF o, of "cmod b"]
+    have th1: "?w ^ n = v^n / complex_of_real (cmod b)" 
+      by (simp add: power_divide complex_of_real_power)
+    have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: norm_divide)
+    hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
+    have th4: "cmod (complex_of_real (cmod b) / b) *
+   cmod (1 + b * (v ^ n / complex_of_real (cmod b)))
+   < cmod (complex_of_real (cmod b) / b) * 1"
+      apply (simp only: norm_mult[symmetric] right_distrib)
+      using b v by (simp add: th2)
+
+    from mult_less_imp_less_left[OF th4 th3]
+    have "?P ?w n" unfolding th1 . 
+    hence "\<exists>z. ?P z n" .. }
+  ultimately show "\<exists>z. ?P z n" by blast
+qed
+
+
+text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
+
+lemma metric_bound_lemma: "cmod (x - y) <= \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
+  using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
+  unfolding cmod_def by simp
+
+lemma bolzano_weierstrass_complex_disc:
+  assumes r: "\<forall>n. cmod (s n) \<le> r"
+  shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
+proof-
+  from seq_monosub[of "Re o s"] 
+  obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))" 
+    unfolding o_def by blast
+  from seq_monosub[of "Im o s o f"] 
+  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast  
+  let ?h = "f o g"
+  from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith 
+  have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>" 
+  proof
+    fix n
+    from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
+  qed
+  have conv1: "convergent (\<lambda>n. Re (s ( f n)))"
+    apply (rule Bseq_monoseq_convergent)
+    apply (simp add: Bseq_def)
+    apply (rule exI[where x= "r + 1"])
+    using th rp apply simp
+    using f(2) .
+  have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>" 
+  proof
+    fix n
+    from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
+  qed
+
+  have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
+    apply (rule Bseq_monoseq_convergent)
+    apply (simp add: Bseq_def)
+    apply (rule exI[where x= "r + 1"])
+    using th rp apply simp
+    using g(2) .
+
+  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x" 
+    by blast 
+  hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r" 
+    unfolding LIMSEQ_def real_norm_def .
+
+  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y" 
+    by blast 
+  hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r" 
+    unfolding LIMSEQ_def real_norm_def .
+  let ?w = "Complex x y"
+  from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto 
+  {fix e assume ep: "e > (0::real)"
+    hence e2: "e/2 > 0" by simp
+    from x[rule_format, OF e2] y[rule_format, OF e2]
+    obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2" and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2" by blast
+    {fix n assume nN12: "n \<ge> N1 + N2"
+      hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
+      from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
+      have "cmod (s (?h n) - ?w) < e" 
+	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
+    hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
+  with hs show ?thesis  by blast  
+qed
+
+text{* Polynomial is continuous. *}
+
+lemma poly_cont:
+  assumes ep: "e > 0" 
+  shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
+proof-
+  from poly_offset[of p z] obtain q where q: "length q = length p" "\<And>x. poly q x = poly p (z + x)" by blast
+  {fix w
+    note q(2)[of "w - z", simplified]}
+  note th = this
+  show ?thesis unfolding th[symmetric]
+  proof(induct q)
+    case Nil thus ?case  using ep by auto
+  next
+    case (Cons c cs)
+    from poly_bound_exists[of 1 "cs"] 
+    obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
+    from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
+    have one0: "1 > (0::real)"  by arith
+    from real_lbound_gt_zero[OF one0 em0] 
+    obtain d where d: "d >0" "d < 1" "d < e / m" by blast
+    from d(1,3) m(1) have dm: "d*m > 0" "d*m < e" 
+      by (simp_all add: field_simps real_mult_order)
+    show ?case 
+      proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
+	fix d w
+	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
+	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
+	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
+	from H have th: "cmod (w-z) \<le> d" by simp 
+	from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
+	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
+      qed  
+    qed
+qed
+
+text{* Hence a polynomial attains minimum on a closed disc 
+  in the complex plane. *}
+lemma  poly_minimum_modulus_disc:
+  "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
+proof-
+  {assume "\<not> r \<ge> 0" hence ?thesis unfolding linorder_not_le
+      apply -
+      apply (rule exI[where x=0]) 
+      apply auto
+      apply (subgoal_tac "cmod w < 0")
+      apply simp
+      apply arith
+      done }
+  moreover
+  {assume rp: "r \<ge> 0"
+    from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp 
+    hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
+    {fix x z
+      assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
+      hence "- x < 0 " by arith
+      with H(2) norm_ge_zero[of "poly p z"]  have False by simp }
+    then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
+    from real_sup_exists[OF mth1 mth2] obtain s where 
+      s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow>(y < s)" by blast
+    let ?m = "-s"
+    {fix y
+      from s[rule_format, of "-y"] have 
+    "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" 
+	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
+    note s1 = this[unfolded minus_minus]
+    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" 
+      by auto
+    {fix n::nat
+      from s1[rule_format, of "?m + 1/real (Suc n)"] 
+      have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
+	by simp}
+    hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
+    from choice[OF th] obtain g where 
+      g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)" 
+      by blast
+    from bolzano_weierstrass_complex_disc[OF g(1)] 
+    obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
+      by blast    
+    {fix w 
+      assume wr: "cmod w \<le> r"
+      let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
+      {assume e: "?e > 0"
+	hence e2: "?e/2 > 0" by simp
+	from poly_cont[OF e2, of z p] obtain d where
+	  d: "d>0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2" by blast
+	{fix w assume w: "cmod (w - z) < d"
+	  have "cmod(poly p w - poly p z) < ?e / 2"
+	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
+	note th1 = this
+	
+	from fz(2)[rule_format, OF d(1)] obtain N1 where 
+	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
+	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
+	  N2: "2/?e < real N2" by blast
+	have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
+	  using N1[rule_format, of "N1 + N2"] th1 by simp
+	{fix a b e2 m :: real
+	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
+          ==> False" by arith}
+      note th0 = this
+      have ath: 
+	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
+      from s1m[OF g(1)[rule_format]]
+      have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
+      from seq_suble[OF fz(1), of "N1+N2"]
+      have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
+      have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"  
+	using N2 by auto
+      from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
+      from g(2)[rule_format, of "f (N1 + N2)"]
+      have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
+      from order_less_le_trans[OF th01 th00]
+      have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
+      from N2 have "2/?e < real (Suc (N1 + N2))" by arith
+      with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
+      have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
+      with ath[OF th31 th32]
+      have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith  
+      have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c" 
+	by arith
+      have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
+\<le> cmod (poly p (g (f (N1 + N2))) - poly p z)" 
+	by (simp add: norm_triangle_ineq3)
+      from ath2[OF th22, of ?m]
+      have thc2: "2*(?e/2) \<le> \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)" by simp
+      from th0[OF th2 thc1 thc2] have False .}
+      hence "?e = 0" by auto
+      then have "cmod (poly p z) = ?m" by simp  
+      with s1m[OF wr]
+      have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
+    hence ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma "(rcis (sqrt (abs r)) (a/2)) ^ 2 = rcis (abs r) a"
+  unfolding power2_eq_square
+  apply (simp add: rcis_mult)
+  apply (simp add: power2_eq_square[symmetric])
+  done
+
+lemma cispi: "cis pi = -1" 
+  unfolding cis_def
+  by simp
+
+lemma "(rcis (sqrt (abs r)) ((pi + a)/2)) ^ 2 = rcis (- abs r) a"
+  unfolding power2_eq_square
+  apply (simp add: rcis_mult add_divide_distrib)
+  apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
+  done
+
+text {* Nonzero polynomial in z goes to infinity as z does. *}
+
+instance complex::idom_char_0 by (intro_classes)
+instance complex :: recpower_idom_char_0 by intro_classes
+
+lemma poly_infinity:
+  assumes ex: "list_ex (\<lambda>c. c \<noteq> 0) p"
+  shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (a#p) z)"
+using ex
+proof(induct p arbitrary: a d)
+  case (Cons c cs a d) 
+  {assume H: "list_ex (\<lambda>c. c\<noteq>0) cs"
+    with Cons.hyps obtain r where r: "\<forall>z. r \<le> cmod z \<longrightarrow> d + cmod a \<le> cmod (poly (c # cs) z)" by blast
+    let ?r = "1 + \<bar>r\<bar>"
+    {fix z assume h: "1 + \<bar>r\<bar> \<le> cmod z"
+      have r0: "r \<le> cmod z" using h by arith
+      from r[rule_format, OF r0]
+      have th0: "d + cmod a \<le> 1 * cmod(poly (c#cs) z)" by arith
+      from h have z1: "cmod z \<ge> 1" by arith
+      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (c#cs) z"]]]
+      have th1: "d \<le> cmod(z * poly (c#cs) z) - cmod a"
+	unfolding norm_mult by (simp add: ring_simps)
+      from complex_mod_triangle_sub[of "z * poly (c#cs) z" a]
+      have th2: "cmod(z * poly (c#cs) z) - cmod a \<le> cmod (poly (a#c#cs) z)" 
+	by (simp add: diff_le_eq ring_simps) 
+      from th1 th2 have "d \<le> cmod (poly (a#c#cs) z)"  by arith}
+    hence ?case by blast}
+  moreover
+  {assume cs0: "\<not> (list_ex (\<lambda>c. c \<noteq> 0) cs)"
+    with Cons.prems have c0: "c \<noteq> 0" by simp
+    from cs0 have cs0': "list_all (\<lambda>c. c = 0) cs" 
+      by (auto simp add: list_all_iff list_ex_iff)
+    {fix z
+      assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
+      from c0 have "cmod c > 0" by simp
+      from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)" 
+	by (simp add: field_simps norm_mult)
+      have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
+      from complex_mod_triangle_sub[of "z*c" a ]
+      have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
+	by (simp add: ring_simps)
+      from ath[OF th1 th0] have "d \<le> cmod (poly (a # c # cs) z)" 
+	using poly_0[OF cs0'] by simp}
+    then have ?case  by blast}
+  ultimately show ?case by blast
+qed simp
+
+text {* Hence polynomial's modulus attains its minimum somewhere. *}
+lemma poly_minimum_modulus:
+  "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
+proof(induct p)
+  case (Cons c cs) 
+  {assume cs0: "list_ex (\<lambda>c. c \<noteq> 0) cs"
+    from poly_infinity[OF cs0, of "cmod (poly (c#cs) 0)" c]
+    obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (c # cs) 0) \<le> cmod (poly (c # cs) z)" by blast
+    have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
+    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "c#cs"] 
+    obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) w)" by blast
+    {fix z assume z: "r \<le> cmod z"
+      from v[of 0] r[OF z] 
+      have "cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) z)"
+	by simp }
+    note v0 = this
+    from v0 v ath[of r] have ?case by blast}
+  moreover
+  {assume cs0: "\<not> (list_ex (\<lambda>c. c\<noteq>0) cs)"
+    hence th:"list_all (\<lambda>c. c = 0) cs" by (simp add: list_all_iff list_ex_iff)
+    from poly_0[OF th] Cons.hyps have ?case by simp}
+  ultimately show ?case by blast
+qed simp
+
+text{* Constant function (non-syntactic characterization). *}
+definition "constant f = (\<forall>x y. f x = f y)"
+
+lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> length p \<ge> 2"
+  unfolding constant_def
+  apply (induct p, auto)
+  apply (unfold not_less[symmetric])
+  apply simp
+  apply (rule ccontr)
+  apply auto
+  done
+ 
+lemma poly_replicate_append:
+  "poly ((replicate n 0)@p) (x::'a::{recpower, comm_ring}) = x^n * poly p x"
+  by(induct n, auto simp add: power_Suc ring_simps)
+
+text {* Decomposition of polynomial, skipping zero coefficients 
+  after the first.  *}
+
+lemma poly_decompose_lemma:
+ assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
+  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (length q + k) = length p \<and> 
+                 (\<forall>z. poly p z = z^k * poly (a#q) z)"
+using nz
+proof(induct p)
+  case Nil thus ?case by simp
+next
+  case (Cons c cs)
+  {assume c0: "c = 0"
+    
+    from Cons.hyps Cons.prems c0 have ?case apply auto
+      apply (rule_tac x="k+1" in exI)
+      apply (rule_tac x="a" in exI, clarsimp)
+      apply (rule_tac x="q" in exI)
+      by (auto simp add: power_Suc)}
+  moreover
+  {assume c0: "c\<noteq>0"
+    hence ?case apply-
+      apply (rule exI[where x=0])
+      apply (rule exI[where x=c], clarsimp)
+      apply (rule exI[where x=cs])
+      apply auto
+      done}
+  ultimately show ?case by blast
+qed
+
+lemma poly_decompose:
+  assumes nc: "~constant(poly p)"
+  shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
+               length q + k + 1 = length p \<and> 
+              (\<forall>z. poly p z = poly p 0 + z^k * poly (a#q) z)"
+using nc 
+proof(induct p)
+  case Nil thus ?case by (simp add: constant_def)
+next
+  case (Cons c cs)
+  {assume C:"\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
+    {fix x y
+      from C have "poly (c#cs) x = poly (c#cs) y" by (cases "x=0", auto)}
+    with Cons.prems have False by (auto simp add: constant_def)}
+  hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
+  from poly_decompose_lemma[OF th] 
+  show ?case 
+    apply clarsimp    
+    apply (rule_tac x="k+1" in exI)
+    apply (rule_tac x="a" in exI)
+    apply simp
+    apply (rule_tac x="q" in exI)
+    apply (auto simp add: power_Suc)
+    done
+qed
+
+text{* Fundamental theorem of algebral *}
+
+lemma fundamental_theorem_of_algebra:
+  assumes nc: "~constant(poly p)"
+  shows "\<exists>z::complex. poly p z = 0"
+using nc
+proof(induct n\<equiv> "length p" arbitrary: p rule: nat_less_induct)
+  fix n fix p :: "complex list"
+  let ?p = "poly p"
+  assume H: "\<forall>m<n. \<forall>p. \<not> constant (poly p) \<longrightarrow> m = length p \<longrightarrow> (\<exists>(z::complex). poly p z = 0)" and nc: "\<not> constant ?p" and n: "n = length p"
+  let ?ths = "\<exists>z. ?p z = 0"
+
+  from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
+  from poly_minimum_modulus obtain c where 
+    c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" by blast
+  {assume pc: "?p c = 0" hence ?ths by blast}
+  moreover
+  {assume pc0: "?p c \<noteq> 0"
+    from poly_offset[of p c] obtain q where
+      q: "length q = length p" "\<forall>x. poly q x = ?p (c+x)" by blast
+    {assume h: "constant (poly q)"
+      from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
+      {fix x y
+	from th have "?p x = poly q (x - c)" by auto 
+	also have "\<dots> = poly q (y - c)" 
+	  using h unfolding constant_def by blast
+	also have "\<dots> = ?p y" using th by auto
+	finally have "?p x = ?p y" .}
+      with nc have False unfolding constant_def by blast }
+    hence qnc: "\<not> constant (poly q)" by blast
+    from q(2) have pqc0: "?p c = poly q 0" by simp
+    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp 
+    let ?a0 = "poly q 0"
+    from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
+    from a00 
+    have qr: "\<forall>z. poly q z = poly (map (op * (inverse ?a0)) q) z * ?a0"
+      by (simp add: poly_cmult_map)
+    let ?r = "map (op * (inverse ?a0)) q"
+    have lgqr: "length q = length ?r" by simp 
+    {assume h: "\<And>x y. poly ?r x = poly ?r y"
+      {fix x y
+	from qr[rule_format, of x] 
+	have "poly q x = poly ?r x * ?a0" by auto
+	also have "\<dots> = poly ?r y * ?a0" using h by simp
+	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
+	finally have "poly q x = poly q y" .} 
+      with qnc have False unfolding constant_def by blast}
+    hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
+    from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
+    {fix w 
+      have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
+	using qr[rule_format, of w] a00 by simp
+      also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
+	using a00 unfolding norm_divide by (simp add: field_simps)
+      finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
+    note mrmq_eq = this
+    from poly_decompose[OF rnc] obtain k a s where 
+      kas: "a\<noteq>0" "k\<noteq>0" "length s + k + 1 = length ?r" 
+      "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (a#s) z" by blast
+    {assume "k + 1 = n"
+      with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=[]" by auto
+      {fix w
+	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
+	  using kas(4)[rule_format, of w] s0 r01 by (simp add: ring_simps)}
+      note hth = this [symmetric]
+	from reduce_poly_simple[OF kas(1,2)] 
+      have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
+    moreover
+    {assume kn: "k+1 \<noteq> n"
+      from kn kas(3) q(1) n[symmetric] have k1n: "k + 1 < n" by simp
+      have th01: "\<not> constant (poly (1#((replicate (k - 1) 0)@[a])))" 
+	unfolding constant_def poly_Nil poly_Cons poly_replicate_append
+	using kas(1) apply simp 
+	by (rule exI[where x=0], rule exI[where x=1], simp)
+      from kas(2) have th02: "k+1 = length (1#((replicate (k - 1) 0)@[a]))" 
+	by simp
+      from H[rule_format, OF k1n th01 th02]
+      obtain w where w: "1 + w^k * a = 0"
+	unfolding poly_Nil poly_Cons poly_replicate_append
+	using kas(2) by (auto simp add: power_Suc[symmetric, of _ "k - Suc 0"] 
+	  mult_assoc[of _ _ a, symmetric])
+      from poly_bound_exists[of "cmod w" s] obtain m where 
+	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
+      have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
+      from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
+      then have wm1: "w^k * a = - 1" by simp
+      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" 
+	using norm_ge_zero[of w] w0 m(1)
+	  by (simp add: inverse_eq_divide zero_less_mult_iff)
+      with real_down2[OF zero_less_one] obtain t where
+	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
+      let ?ct = "complex_of_real t"
+      let ?w = "?ct * w"
+      have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: ring_simps power_mult_distrib)
+      also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
+	unfolding wm1 by (simp)
+      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
+	apply -
+	apply (rule cong[OF refl[of cmod]])
+	apply assumption
+	done
+      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] 
+      have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" unfolding norm_of_real by simp 
+      have ath: "\<And>x (t::real). 0\<le> x \<Longrightarrow> x < t \<Longrightarrow> t\<le>1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1" by arith
+      have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
+      then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult) 
+      from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
+	by (simp add: inverse_eq_divide field_simps)
+      with zero_less_power[OF t(1), of k] 
+      have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
+	apply - apply (rule mult_strict_left_mono) by simp_all
+      have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
+	by (simp add: ring_simps power_mult_distrib norm_of_real norm_power norm_mult)
+      then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
+	using t(1,2) m(2)[rule_format, OF tw] w0
+	apply (simp only: )
+	apply auto
+	apply (rule mult_mono, simp_all add: norm_ge_zero)+
+	apply (simp add: zero_le_mult_iff zero_le_power)
+	done
+      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp 
+      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1" 
+	by auto
+      from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
+      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" . 
+      from th11 th12
+      have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith 
+      then have "cmod (poly ?r ?w) < 1" 
+	unfolding kas(4)[rule_format, of ?w] r01 by simp 
+      then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
+    ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
+    from cr0_contr cq0 q(2)
+    have ?ths unfolding mrmq_eq not_less[symmetric] by auto}
+  ultimately show ?ths by blast
+qed
+
+text {* Alternative version with a syntactic notion of constant polynomial. *}
+
+lemma fundamental_theorem_of_algebra_alt:
+  assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> list_all(\<lambda>b. b = 0) l \<and> p = a#l)"
+  shows "\<exists>z. poly p z = (0::complex)"
+using nc
+proof(induct p)
+  case (Cons c cs)
+  {assume "c=0" hence ?case by auto}
+  moreover
+  {assume c0: "c\<noteq>0"
+    {assume nc: "constant (poly (c#cs))"
+      from nc[unfolded constant_def, rule_format, of 0] 
+      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
+      hence "list_all (\<lambda>c. c=0) cs"
+	proof(induct cs)
+	  case (Cons d ds)
+	  {assume "d=0" hence ?case using Cons.prems Cons.hyps by simp}
+	  moreover
+	  {assume d0: "d\<noteq>0"
+	    from poly_bound_exists[of 1 ds] obtain m where 
+	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
+	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
+	    from real_down2[OF dm zero_less_one] obtain x where 
+	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
+	    let ?x = "complex_of_real x"
+	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
+	    from Cons.prems[rule_format, OF cx(1)]
+	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
+	    from m(2)[rule_format, OF cx(2)] x(1)
+	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
+	      by (simp add: norm_mult)
+	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
+	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
+	    with cth  have ?case by blast}
+	  ultimately show ?case by blast 
+	qed simp}
+      then have nc: "\<not> constant (poly (c#cs))" using Cons.prems c0 
+	by blast
+      from fundamental_theorem_of_algebra[OF nc] have ?case .}
+  ultimately show ?case by blast  
+qed simp
+
+subsection{* Nullstellenstatz, degrees and divisibility of polynomials *}
+
+lemma nullstellensatz_lemma:
+  fixes p :: "complex list"
+  assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
+  and "degree p = n" and "n \<noteq> 0"
+  shows "p divides (pexp q n)"
+using prems
+proof(induct n arbitrary: p q rule: nat_less_induct)
+  fix n::nat fix p q :: "complex list"
+  assume IH: "\<forall>m<n. \<forall>p q.
+                 (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
+                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p divides (q %^ m)"
+    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
+    and dpn: "degree p = n" and n0: "n \<noteq> 0"
+  let ?ths = "p divides (q %^ n)"
+  {fix a assume a: "poly p a = 0"
+    {assume p0: "poly p = poly []" 
+      hence ?ths unfolding divides_def  using pq0 n0
+	apply - apply (rule exI[where x="[]"], rule ext)
+	by (auto simp add: poly_mult poly_exp)}
+    moreover
+    {assume p0: "poly p \<noteq> poly []" 
+      and oa: "order  a p \<noteq> 0"
+      from p0 have pne: "p \<noteq> []" by auto
+      let ?op = "order a p"
+      from p0 have ap: "([- a, 1] %^ ?op) divides p" 
+	"\<not> pexp [- a, 1] (Suc ?op) divides p" using order by blast+ 
+      note oop = order_degree[OF p0, unfolded dpn]
+      {assume q0: "q = []"
+	hence ?ths using n0 unfolding divides_def 
+	  apply simp
+	  apply (rule exI[where x="[]"], rule ext)
+	  by (simp add: divides_def poly_exp poly_mult)}
+      moreover
+      {assume q0: "q\<noteq>[]"
+	from pq0[rule_format, OF a, unfolded poly_linear_divides] q0
+	obtain r where r: "q = pmult [- a, 1] r" by blast
+	from ap[unfolded divides_def] obtain s where
+	  s: "poly p = poly (pmult (pexp [- a, 1] ?op) s)" by blast
+	have s0: "poly s \<noteq> poly []"
+	  using s p0 by (simp add: poly_entire)
+	hence pns0: "poly (pnormalize s) \<noteq> poly []" and sne: "s\<noteq>[]" by auto
+	{assume ds0: "degree s = 0"
+	  from ds0 pns0 have "\<exists>k. pnormalize s = [k]" unfolding degree_def 
+	    by (cases "pnormalize s", auto)
+	  then obtain k where kpn: "pnormalize s = [k]" by blast
+	  from pns0[unfolded poly_zero] kpn have k: "k \<noteq>0" "poly s = poly [k]"
+	    using poly_normalize[of s] by simp_all
+	  let ?w = "pmult (pmult [1/k] (pexp [-a,1] (n - ?op))) (pexp r n)"
+	  from k r s oop have "poly (pexp q n) = poly (pmult p ?w)"
+	    by - (rule ext, simp add: poly_mult poly_exp poly_cmult poly_add power_add[symmetric] ring_simps power_mult_distrib[symmetric])
+	  hence ?ths unfolding divides_def by blast}
+	moreover
+	{assume ds0: "degree s \<noteq> 0"
+	  from ds0 s0 dpn degree_unique[OF s, unfolded linear_pow_mul_degree] oa
+	    have dsn: "degree s < n" by auto 
+	    {fix x assume h: "poly s x = 0"
+	      {assume xa: "x = a"
+		from h[unfolded xa poly_linear_divides] sne obtain u where
+		  u: "s = pmult [- a, 1] u" by blast
+		have "poly p = poly (pmult (pexp [- a, 1] (Suc ?op)) u)"
+		  unfolding s u
+		  apply (rule ext)
+		  by (simp add: ring_simps power_mult_distrib[symmetric] poly_mult poly_cmult poly_add poly_exp)
+		with ap(2)[unfolded divides_def] have False by blast}
+	      note xa = this
+	      from h s have "poly p x = 0" by (simp add: poly_mult)
+	      with pq0 have "poly q x = 0" by blast
+	      with r xa have "poly r x = 0"
+		by (auto simp add: poly_mult poly_add poly_cmult eq_diff_eq[symmetric])}
+	    note impth = this
+	    from IH[rule_format, OF dsn, of s r] impth ds0
+	    have "s divides (pexp r (degree s))" by blast
+	    then obtain u where u: "poly (pexp r (degree s)) = poly (pmult s u)"
+	      unfolding divides_def by blast
+	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
+	      by (simp add: poly_mult[symmetric] poly_exp[symmetric])
+	    let ?w = "pmult (pmult u (pexp [-a,1] (n - ?op))) (pexp r (n - degree s))"
+	    from u' s r oop[of a] dsn have "poly (pexp q n) = poly (pmult p ?w)"
+	      apply - apply (rule ext)
+	      apply (simp only:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult ring_simps)
+	      
+	      apply (simp add:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult mult_assoc[symmetric])
+	      done
+	    hence ?ths unfolding divides_def by blast}
+      ultimately have ?ths by blast }
+      ultimately have ?ths by blast}
+    ultimately have ?ths using a order_root by blast}
+  moreover
+  {assume exa: "\<not> (\<exists>a. poly p a = 0)"
+    from fundamental_theorem_of_algebra_alt[of p] exa obtain c cs where
+      ccs: "c\<noteq>0" "list_all (\<lambda>c. c = 0) cs" "p = c#cs" by blast
+    
+    from poly_0[OF ccs(2)] ccs(3) 
+    have pp: "\<And>x. poly p x =  c" by simp
+    let ?w = "pmult [1/c] (pexp q n)"
+    from pp ccs(1) 
+    have "poly (pexp q n) = poly (pmult p ?w) "
+      apply - apply (rule ext)
+      unfolding poly_mult_assoc[symmetric] by (simp add: poly_mult)
+    hence ?ths unfolding divides_def by blast}
+  ultimately show ?ths by blast
+qed
+
+lemma nullstellensatz_univariate:
+  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
+    p divides (q %^ (degree p)) \<or> (poly p = poly [] \<and> poly q = poly [])"
+proof-
+  {assume pe: "poly p = poly []"
+    hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> poly q = poly []"
+      apply auto
+      by (rule ext, simp)
+    {assume "p divides (pexp q (degree p))"
+      then obtain r where r: "poly (pexp q (degree p)) = poly (pmult p r)" 
+	unfolding divides_def by blast
+      from cong[OF r refl] pe degree_unique[OF pe]
+      have False by (simp add: poly_mult degree_def)}
+    with eq pe have ?thesis by blast}
+  moreover
+  {assume pe: "poly p \<noteq> poly []"
+    have p0: "poly [0] = poly []" by (rule ext, simp)
+    {assume dp: "degree p = 0"
+      then obtain k where "pnormalize p = [k]" using pe poly_normalize[of p]
+	unfolding degree_def by (cases "pnormalize p", auto)
+      hence k: "pnormalize p = [k]" "poly p = poly [k]" "k\<noteq>0"
+	using pe poly_normalize[of p] by (auto simp add: p0)
+      hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
+      from k(2,3) dp have "poly (pexp q (degree p)) = poly (pmult p [1/k]) "
+	by - (rule ext, simp add: poly_mult poly_exp)
+      hence th2: "p divides (pexp q (degree p))" unfolding divides_def by blast
+      from th1 th2 pe have ?thesis by blast}
+    moreover
+    {assume dp: "degree p \<noteq> 0"
+      then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
+      {assume "p divides (pexp q (Suc n))"
+	then obtain u where u: "poly (pexp q (Suc n)) = poly (pmult p u)"
+	  unfolding divides_def by blast
+	hence u' :"\<And>x. poly (pexp q (Suc n)) x = poly (pmult p u) x" by simp_all
+	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
+	  hence "poly (pexp q (Suc n)) x \<noteq> 0" by (simp only: poly_exp) simp	  
+	  hence False using u' h(1) by (simp only: poly_mult poly_exp) simp}}
+	with n nullstellensatz_lemma[of p q "degree p"] dp 
+	have ?thesis by auto}
+    ultimately have ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+text{* Useful lemma *}
+
+lemma (in idom_char_0) constant_degree: "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
+proof
+  assume l: ?lhs
+  from l[unfolded constant_def, rule_format, of _ "zero"]
+  have th: "poly p = poly [poly p 0]" apply - by (rule ext, simp)
+  from degree_unique[OF th] show ?rhs by (simp add: degree_def)
+next
+  assume r: ?rhs
+  from r have "pnormalize p = [] \<or> (\<exists>k. pnormalize p = [k])"
+    unfolding degree_def by (cases "pnormalize p", auto)
+  then show ?lhs unfolding constant_def poly_normalize[of p, symmetric]
+    by (auto simp del: poly_normalize)
+qed
+
+(* It would be nicer to prove this without using algebraic closure...        *)
+
+lemma divides_degree_lemma: assumes dpn: "degree (p::complex list) = n"
+  shows "n \<le> degree (p *** q) \<or> poly (p *** q) = poly []"
+  using dpn
+proof(induct n arbitrary: p q)
+  case 0 thus ?case by simp
+next
+  case (Suc n p q)
+  from Suc.prems fundamental_theorem_of_algebra[of p] constant_degree[of p]
+  obtain a where a: "poly p a = 0" by auto
+  then obtain r where r: "p = pmult [-a, 1] r" unfolding poly_linear_divides
+    using Suc.prems by (auto simp add: degree_def)
+  {assume h: "poly (pmult r q) = poly []"
+    hence "poly (pmult p q) = poly []" using r
+      apply - apply (rule ext)  by (auto simp add: poly_entire poly_mult poly_add poly_cmult) hence ?case by blast}
+  moreover
+  {assume h: "poly (pmult r q) \<noteq> poly []" 
+    hence r0: "poly r \<noteq> poly []" and q0: "poly q \<noteq> poly []"
+      by (auto simp add: poly_entire)
+    have eq: "poly (pmult p q) = poly (pmult [-a, 1] (pmult r q))"
+      apply - apply (rule ext)
+      by (simp add: r poly_mult poly_add poly_cmult ring_simps)
+    from linear_mul_degree[OF h, of "- a"]
+    have dqe: "degree (pmult p q) = degree (pmult r q) + 1"
+      unfolding degree_unique[OF eq] .
+    from linear_mul_degree[OF r0, of "- a", unfolded r[symmetric]] r Suc.prems 
+    have dr: "degree r = n" by auto
+    from  Suc.hyps[OF dr, of q] have "Suc n \<le> degree (pmult p q)"
+      unfolding dqe using h by (auto simp del: poly.simps) 
+    hence ?case by blast}
+  ultimately show ?case by blast
+qed
+
+lemma divides_degree: assumes pq: "p divides (q:: complex list)"
+  shows "degree p \<le> degree q \<or> poly q = poly []"
+using pq  divides_degree_lemma[OF refl, of p]
+apply (auto simp add: divides_def poly_entire)
+apply atomize
+apply (erule_tac x="qa" in allE, auto)
+apply (subgoal_tac "degree q = degree (p *** qa)", simp)
+apply (rule degree_unique, simp)
+done
+
+(* Arithmetic operations on multivariate polynomials.                        *)
+
+lemma mpoly_base_conv: 
+  "(0::complex) \<equiv> poly [] x" "c \<equiv> poly [c] x" "x \<equiv> poly [0,1] x" by simp_all
+
+lemma mpoly_norm_conv: 
+  "poly [0] (x::complex) \<equiv> poly [] x" "poly [poly [] y] x \<equiv> poly [] x" by simp_all
+
+lemma mpoly_sub_conv: 
+  "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
+  by (simp add: diff_def)
+
+lemma poly_pad_rule: "poly p x = 0 ==> poly (0#p) x = (0::complex)" by simp
+
+lemma poly_cancel_eq_conv: "p = (0::complex) \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (q = 0) \<equiv> (a * q - b * p = 0)" apply (atomize (full)) by auto
+
+lemma resolve_eq_raw:  "poly [] x \<equiv> 0" "poly [c] x \<equiv> (c::complex)" by auto
+lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
+  \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
+lemma expand_ex_beta_conv: "list_ex P [c] \<equiv> P c" by simp
+
+lemma poly_divides_pad_rule: 
+  fixes p q :: "complex list"
+  assumes pq: "p divides q"
+  shows "p divides ((0::complex)#q)"
+proof-
+  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
+  hence "poly (0#q) = poly (p *** ([0,1] *** r))" 
+    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
+  thus ?thesis unfolding divides_def by blast
+qed
+
+lemma poly_divides_pad_const_rule: 
+  fixes p q :: "complex list"
+  assumes pq: "p divides q"
+  shows "p divides (a %* q)"
+proof-
+  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
+  hence "poly (a %* q) = poly (p *** (a %* r))" 
+    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
+  thus ?thesis unfolding divides_def by blast
+qed
+
+
+lemma poly_divides_conv0:  
+  fixes p :: "complex list"
+  assumes lgpq: "length q < length p" and lq:"last p \<noteq> 0"
+  shows "p divides q \<equiv> (\<not> (list_ex (\<lambda>c. c \<noteq> 0) q))" (is "?lhs \<equiv> ?rhs")
+proof-
+  {assume r: ?rhs 
+    hence eq: "poly q = poly []" unfolding poly_zero 
+      by (simp add: list_all_iff list_ex_iff)
+    hence "poly q = poly (p *** [])" by - (rule ext, simp add: poly_mult)
+    hence ?lhs unfolding divides_def  by blast}
+  moreover
+  {assume l: ?lhs
+    have ath: "\<And>lq lp dq::nat. lq < lp ==> lq \<noteq> 0 \<Longrightarrow> dq <= lq - 1 ==> dq < lp - 1"
+      by arith
+    {assume q0: "length q = 0"
+      hence "q = []" by simp
+      hence ?rhs by simp}
+    moreover
+    {assume lgq0: "length q \<noteq> 0"
+      from pnormalize_length[of q] have dql: "degree q \<le> length q - 1" 
+	unfolding degree_def by simp
+      from ath[OF lgpq lgq0 dql, unfolded pnormal_degree[OF lq, symmetric]] divides_degree[OF l] have "poly q = poly []" by auto
+      hence ?rhs unfolding poly_zero by (simp add: list_all_iff list_ex_iff)}
+    ultimately have ?rhs by blast }
+  ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
+qed
+
+lemma poly_divides_conv1: 
+  assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex list) divides p'"
+  and qrp': "\<And>x. a * poly q x - poly p' x \<equiv> poly r x"
+  shows "p divides q \<equiv> p divides (r::complex list)" (is "?lhs \<equiv> ?rhs")
+proof-
+  {
+  from pp' obtain t where t: "poly p' = poly (p *** t)" 
+    unfolding divides_def by blast
+  {assume l: ?lhs
+    then obtain u where u: "poly q = poly (p *** u)" unfolding divides_def by blast
+     have "poly r = poly (p *** ((a %* u) +++ (-- t)))"
+       using u qrp' t
+       by - (rule ext, 
+	 simp add: poly_add poly_mult poly_cmult poly_minus ring_simps)
+     then have ?rhs unfolding divides_def by blast}
+  moreover
+  {assume r: ?rhs
+    then obtain u where u: "poly r = poly (p *** u)" unfolding divides_def by blast
+    from u t qrp' a0 have "poly q = poly (p *** ((1/a) %* (u +++ t)))"
+      by - (rule ext, atomize (full), simp add: poly_mult poly_add poly_cmult field_simps)
+    hence ?lhs  unfolding divides_def by blast}
+  ultimately have "?lhs = ?rhs" by blast }
+thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
+qed
+
+lemma basic_cqe_conv1:
+  "(\<exists>x. poly p x = 0 \<and> poly [] x \<noteq> 0) \<equiv> False"
+  "(\<exists>x. poly [] x \<noteq> 0) \<equiv> False"
+  "(\<exists>x. poly [c] x \<noteq> 0) \<equiv> c\<noteq>0"
+  "(\<exists>x. poly [] x = 0) \<equiv> True"
+  "(\<exists>x. poly [c] x = 0) \<equiv> c = 0" by simp_all
+
+lemma basic_cqe_conv2: 
+  assumes l:"last (a#b#p) \<noteq> 0" 
+  shows "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True"
+proof-
+  {fix h t
+    assume h: "h\<noteq>0" "list_all (\<lambda>c. c=(0::complex)) t"  "a#b#p = h#t"
+    hence "list_all (\<lambda>c. c= 0) (b#p)" by simp
+    moreover have "last (b#p) \<in> set (b#p)" by simp
+    ultimately have "last (b#p) = 0" by (simp add: list_all_iff)
+    with l have False by simp}
+  hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> list_all (\<lambda>c. c=0) t \<and> a#b#p = h#t)"
+    by blast
+  from fundamental_theorem_of_algebra_alt[OF th] 
+  show "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True" by auto
+qed
+
+lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
+proof-
+  have "\<not> (list_ex (\<lambda>c. c \<noteq> 0) p) \<longleftrightarrow> poly p = poly []" 
+    by (simp add: poly_zero list_all_iff list_ex_iff)
+  also have "\<dots> \<longleftrightarrow> (\<not> (\<exists>x. poly p x \<noteq> 0))" by (auto intro: ext)
+  finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
+    by - (atomize (full), blast)
+qed
+
+lemma basic_cqe_conv3:
+  fixes p q :: "complex list"
+  assumes l: "last (a#p) \<noteq> 0" 
+  shows "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
+proof-
+  note np = pnormalize_eq[OF l]
+  {assume "poly (a#p) = poly []" hence False using l
+      unfolding poly_zero apply (auto simp add: list_all_iff del: last.simps)
+      apply (cases p, simp_all) done}
+  then have p0: "poly (a#p) \<noteq> poly []"  by blast
+  from np have dp:"degree (a#p) = length p" by (simp add: degree_def)
+  from nullstellensatz_univariate[of "a#p" q] p0 dp
+  show "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
+    by - (atomize (full), auto)
+qed
+
+lemma basic_cqe_conv4:
+  fixes p q :: "complex list"
+  assumes h: "\<And>x. poly (q %^ n) x \<equiv> poly r x"
+  shows "p divides (q %^ n) \<equiv> p divides r"
+proof-
+  from h have "poly (q %^ n) = poly r" by (auto intro: ext)  
+  thus "p divides (q %^ n) \<equiv> p divides r" unfolding divides_def by simp
+qed
+
+lemma pmult_Cons_Cons: "((a::complex)#b#p) *** q = (a %*q) +++ (0#((b#p) *** q))"
+  by simp
+
+lemma elim_neg_conv: "- z \<equiv> (-1) * (z::complex)" by simp
+lemma eqT_intr: "PROP P \<Longrightarrow> (True \<Longrightarrow> PROP P )" "PROP P \<Longrightarrow> True" by blast+
+lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
+lemma last_simps: "last [x] = x" "last (x#y#ys) = last (y#ys)" by simp_all
+lemma length_simps: "length [] = 0" "length (x#y#xs) = length xs + 2" "length [x] = 1" by simp_all
+
+lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
+lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
+  by (atomize (full)) simp_all
+lemma cqe_conv1: "poly [] x = 0 \<longleftrightarrow> True"  by simp
+lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
+proof
+  assume "p \<Longrightarrow> q \<equiv> r" thus "p \<and> q \<equiv> p \<and> r" apply - apply (atomize (full)) by blast
+next
+  assume "p \<and> q \<equiv> p \<and> r" "p"
+  thus "q \<equiv> r" apply - apply (atomize (full)) apply blast done
+qed
+lemma poly_const_conv: "poly [c] (x::complex) = y \<longleftrightarrow> c = y" by simp
+
+end
\ No newline at end of file
--- a/src/HOL/Groebner_Basis.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Groebner_Basis.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Groebner_Basis.thy
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -164,8 +163,8 @@
 
 end
 
-interpretation class_semiring: gb_semiring
-    ["op +" "op *" "op ^" "0::'a::{comm_semiring_1, recpower}" "1"]
+interpretation class_semiring!: gb_semiring
+    "op +" "op *" "op ^" "0::'a::{comm_semiring_1, recpower}" "1"
   proof qed (auto simp add: ring_simps power_Suc)
 
 lemmas nat_arith =
@@ -243,8 +242,8 @@
 end
 
 
-interpretation class_ring: gb_ring ["op +" "op *" "op ^"
-    "0::'a::{comm_semiring_1,recpower,number_ring}" 1 "op -" "uminus"]
+interpretation class_ring!: gb_ring "op +" "op *" "op ^"
+    "0::'a::{comm_semiring_1,recpower,number_ring}" 1 "op -" "uminus"
   proof qed simp_all
 
 
@@ -344,8 +343,8 @@
   thus "b = 0" by blast
 qed
 
-interpretation class_ringb: ringb
-  ["op +" "op *" "op ^" "0::'a::{idom,recpower,number_ring}" "1" "op -" "uminus"]
+interpretation class_ringb!: ringb
+  "op +" "op *" "op ^" "0::'a::{idom,recpower,number_ring}" "1" "op -" "uminus"
 proof(unfold_locales, simp add: ring_simps power_Suc, auto)
   fix w x y z ::"'a::{idom,recpower,number_ring}"
   assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z"
@@ -360,8 +359,8 @@
 
 declaration {* normalizer_funs @{thm class_ringb.ringb_axioms'} *}
 
-interpretation natgb: semiringb
-  ["op +" "op *" "op ^" "0::nat" "1"]
+interpretation natgb!: semiringb
+  "op +" "op *" "op ^" "0::nat" "1"
 proof (unfold_locales, simp add: ring_simps power_Suc)
   fix w x y z ::"nat"
   { assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z"
@@ -465,8 +464,8 @@
 
 subsection{* Groebner Bases for fields *}
 
-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)
+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
 lemma divide_Numeral0: "(x::'a::{field,number_ring, division_by_zero}) / Numeral0 = 0"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/Bounds.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,82 @@
+(*  Title:      HOL/Real/HahnBanach/Bounds.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Bounds *}
+
+theory Bounds
+imports Main ContNotDenum
+begin
+
+locale lub =
+  fixes A and x
+  assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
+    and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
+
+lemmas [elim?] = lub.least lub.upper
+
+definition
+  the_lub :: "'a::order set \<Rightarrow> 'a" where
+  "the_lub A = The (lub A)"
+
+notation (xsymbols)
+  the_lub  ("\<Squnion>_" [90] 90)
+
+lemma the_lub_equality [elim?]:
+  assumes "lub A x"
+  shows "\<Squnion>A = (x::'a::order)"
+proof -
+  interpret lub A x by fact
+  show ?thesis
+  proof (unfold the_lub_def)
+    from `lub A x` show "The (lub A) = x"
+    proof
+      fix x' assume lub': "lub A x'"
+      show "x' = x"
+      proof (rule order_antisym)
+	from lub' show "x' \<le> x"
+	proof
+          fix a assume "a \<in> A"
+          then show "a \<le> x" ..
+	qed
+	show "x \<le> x'"
+	proof
+          fix a assume "a \<in> A"
+          with lub' show "a \<le> x'" ..
+	qed
+      qed
+    qed
+  qed
+qed
+
+lemma the_lubI_ex:
+  assumes ex: "\<exists>x. lub A x"
+  shows "lub A (\<Squnion>A)"
+proof -
+  from ex obtain x where x: "lub A x" ..
+  also from x have [symmetric]: "\<Squnion>A = x" ..
+  finally show ?thesis .
+qed
+
+lemma lub_compat: "lub A x = isLub UNIV A x"
+proof -
+  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
+    by (rule ext) (simp only: isUb_def)
+  then show ?thesis
+    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
+qed
+
+lemma real_complete:
+  fixes A :: "real set"
+  assumes nonempty: "\<exists>a. a \<in> A"
+    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
+  shows "\<exists>x. lub A x"
+proof -
+  from ex_upper have "\<exists>y. isUb UNIV A y"
+    unfolding isUb_def setle_def by blast
+  with nonempty have "\<exists>x. isLub UNIV A x"
+    by (rule reals_complete)
+  then show ?thesis by (simp only: lub_compat)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/FunctionNorm.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,278 @@
+(*  Title:      HOL/Real/HahnBanach/FunctionNorm.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The norm of a function *}
+
+theory FunctionNorm
+imports NormedSpace FunctionOrder
+begin
+
+subsection {* Continuous linear forms*}
+
+text {*
+  A linear form @{text f} on a normed vector space @{text "(V, \<parallel>\<cdot>\<parallel>)"}
+  is \emph{continuous}, iff it is bounded, i.e.
+  \begin{center}
+  @{text "\<exists>c \<in> R. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
+  \end{center}
+  In our application no other functions than linear forms are
+  considered, so we can define continuous linear forms as bounded
+  linear forms:
+*}
+
+locale continuous = var_V + norm_syntax + linearform +
+  assumes bounded: "\<exists>c. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
+
+declare continuous.intro [intro?] continuous_axioms.intro [intro?]
+
+lemma continuousI [intro]:
+  fixes norm :: "_ \<Rightarrow> real"  ("\<parallel>_\<parallel>")
+  assumes "linearform V f"
+  assumes r: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
+  shows "continuous V norm f"
+proof
+  show "linearform V f" by fact
+  from r have "\<exists>c. \<forall>x\<in>V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" by blast
+  then show "continuous_axioms V norm f" ..
+qed
+
+
+subsection {* The norm of a linear form *}
+
+text {*
+  The least real number @{text c} for which holds
+  \begin{center}
+  @{text "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
+  \end{center}
+  is called the \emph{norm} of @{text f}.
+
+  For non-trivial vector spaces @{text "V \<noteq> {0}"} the norm can be
+  defined as
+  \begin{center}
+  @{text "\<parallel>f\<parallel> = \<sup>x \<noteq> 0. \<bar>f x\<bar> / \<parallel>x\<parallel>"}
+  \end{center}
+
+  For the case @{text "V = {0}"} the supremum would be taken from an
+  empty set. Since @{text \<real>} is unbounded, there would be no supremum.
+  To avoid this situation it must be guaranteed that there is an
+  element in this set. This element must be @{text "{} \<ge> 0"} so that
+  @{text fn_norm} has the norm properties. Furthermore it does not
+  have to change the norm in all other cases, so it must be @{text 0},
+  as all other elements are @{text "{} \<ge> 0"}.
+
+  Thus we define the set @{text B} where the supremum is taken from as
+  follows:
+  \begin{center}
+  @{text "{0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel>. x \<noteq> 0 \<and> x \<in> F}"}
+  \end{center}
+
+  @{text fn_norm} is equal to the supremum of @{text B}, if the
+  supremum exists (otherwise it is undefined).
+*}
+
+locale fn_norm = norm_syntax +
+  fixes B defines "B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
+  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
+  defines "\<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
+
+locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm
+
+lemma (in fn_norm) B_not_empty [intro]: "0 \<in> B V f"
+  by (simp add: B_def)
+
+text {*
+  The following lemma states that every continuous linear form on a
+  normed space @{text "(V, \<parallel>\<cdot>\<parallel>)"} has a function norm.
+*}
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_works:
+  assumes "continuous V norm f"
+  shows "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
+proof -
+  interpret continuous V norm f by fact
+  txt {* The existence of the supremum is shown using the
+    completeness of the reals. Completeness means, that every
+    non-empty bounded set of reals has a supremum. *}
+  have "\<exists>a. lub (B V f) a"
+  proof (rule real_complete)
+    txt {* First we have to show that @{text B} is non-empty: *}
+    have "0 \<in> B V f" ..
+    then show "\<exists>x. x \<in> B V f" ..
+
+    txt {* Then we have to show that @{text B} is bounded: *}
+    show "\<exists>c. \<forall>y \<in> B V f. y \<le> c"
+    proof -
+      txt {* We know that @{text f} is bounded by some value @{text c}. *}
+      from bounded obtain c where c: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
+
+      txt {* To prove the thesis, we have to show that there is some
+        @{text b}, such that @{text "y \<le> b"} for all @{text "y \<in>
+        B"}. Due to the definition of @{text B} there are two cases. *}
+
+      def b \<equiv> "max c 0"
+      have "\<forall>y \<in> B V f. y \<le> b"
+      proof
+        fix y assume y: "y \<in> B V f"
+        show "y \<le> b"
+        proof cases
+          assume "y = 0"
+          then show ?thesis unfolding b_def by arith
+        next
+          txt {* The second case is @{text "y = \<bar>f x\<bar> / \<parallel>x\<parallel>"} for some
+            @{text "x \<in> V"} with @{text "x \<noteq> 0"}. *}
+          assume "y \<noteq> 0"
+          with y obtain x where y_rep: "y = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
+              and x: "x \<in> V" and neq: "x \<noteq> 0"
+            by (auto simp add: B_def real_divide_def)
+          from x neq have gt: "0 < \<parallel>x\<parallel>" ..
+
+          txt {* The thesis follows by a short calculation using the
+            fact that @{text f} is bounded. *}
+
+          note y_rep
+          also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
+          proof (rule mult_right_mono)
+            from c x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
+            from gt have "0 < inverse \<parallel>x\<parallel>" 
+              by (rule positive_imp_inverse_positive)
+            then show "0 \<le> inverse \<parallel>x\<parallel>" by (rule order_less_imp_le)
+          qed
+          also have "\<dots> = c * (\<parallel>x\<parallel> * inverse \<parallel>x\<parallel>)"
+            by (rule real_mult_assoc)
+          also
+          from gt have "\<parallel>x\<parallel> \<noteq> 0" by simp
+          then have "\<parallel>x\<parallel> * inverse \<parallel>x\<parallel> = 1" by simp 
+          also have "c * 1 \<le> b" by (simp add: b_def le_maxI1)
+          finally show "y \<le> b" .
+        qed
+      qed
+      then show ?thesis ..
+    qed
+  qed
+  then show ?thesis unfolding fn_norm_def by (rule the_lubI_ex)
+qed
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_ub [iff?]:
+  assumes "continuous V norm f"
+  assumes b: "b \<in> B V f"
+  shows "b \<le> \<parallel>f\<parallel>\<hyphen>V"
+proof -
+  interpret continuous V norm f by fact
+  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
+    using `continuous V norm f` by (rule fn_norm_works)
+  from this and b show ?thesis ..
+qed
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_leastB:
+  assumes "continuous V norm f"
+  assumes b: "\<And>b. b \<in> B V f \<Longrightarrow> b \<le> y"
+  shows "\<parallel>f\<parallel>\<hyphen>V \<le> y"
+proof -
+  interpret continuous V norm f by fact
+  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
+    using `continuous V norm f` by (rule fn_norm_works)
+  from this and b show ?thesis ..
+qed
+
+text {* The norm of a continuous function is always @{text "\<ge> 0"}. *}
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_ge_zero [iff]:
+  assumes "continuous V norm f"
+  shows "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
+proof -
+  interpret continuous V norm f by fact
+  txt {* The function norm is defined as the supremum of @{text B}.
+    So it is @{text "\<ge> 0"} if all elements in @{text B} are @{text "\<ge>
+    0"}, provided the supremum exists and @{text B} is not empty. *}
+  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
+    using `continuous V norm f` by (rule fn_norm_works)
+  moreover have "0 \<in> B V f" ..
+  ultimately show ?thesis ..
+qed
+
+text {*
+  \medskip The fundamental property of function norms is:
+  \begin{center}
+  @{text "\<bar>f x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
+  \end{center}
+*}
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_le_cong:
+  assumes "continuous V norm f" "linearform V f"
+  assumes x: "x \<in> V"
+  shows "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
+proof -
+  interpret continuous V norm f by fact
+  interpret linearform V f by fact
+  show ?thesis
+  proof cases
+    assume "x = 0"
+    then have "\<bar>f x\<bar> = \<bar>f 0\<bar>" by simp
+    also have "f 0 = 0" by rule unfold_locales
+    also have "\<bar>\<dots>\<bar> = 0" by simp
+    also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
+      using `continuous V norm f` by (rule fn_norm_ge_zero)
+    from x have "0 \<le> norm x" ..
+    with a have "0 \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" by (simp add: zero_le_mult_iff)
+    finally show "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" .
+  next
+    assume "x \<noteq> 0"
+    with x have neq: "\<parallel>x\<parallel> \<noteq> 0" by simp
+    then have "\<bar>f x\<bar> = (\<bar>f x\<bar> * inverse \<parallel>x\<parallel>) * \<parallel>x\<parallel>" by simp
+    also have "\<dots> \<le>  \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
+    proof (rule mult_right_mono)
+      from x show "0 \<le> \<parallel>x\<parallel>" ..
+      from x and neq have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<in> B V f"
+	by (auto simp add: B_def real_divide_def)
+      with `continuous V norm f` show "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>V"
+	by (rule fn_norm_ub)
+    qed
+    finally show ?thesis .
+  qed
+qed
+
+text {*
+  \medskip The function norm is the least positive real number for
+  which the following inequation holds:
+  \begin{center}
+    @{text "\<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
+  \end{center}
+*}
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_least [intro?]:
+  assumes "continuous V norm f"
+  assumes ineq: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" and ge: "0 \<le> c"
+  shows "\<parallel>f\<parallel>\<hyphen>V \<le> c"
+proof -
+  interpret continuous V norm f by fact
+  show ?thesis
+  proof (rule fn_norm_leastB [folded B_def fn_norm_def])
+    fix b assume b: "b \<in> B V f"
+    show "b \<le> c"
+    proof cases
+      assume "b = 0"
+      with ge show ?thesis by simp
+    next
+      assume "b \<noteq> 0"
+      with b obtain x where b_rep: "b = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
+        and x_neq: "x \<noteq> 0" and x: "x \<in> V"
+	by (auto simp add: B_def real_divide_def)
+      note b_rep
+      also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
+      proof (rule mult_right_mono)
+	have "0 < \<parallel>x\<parallel>" using x x_neq ..
+	then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
+	from ineq and x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
+      qed
+      also have "\<dots> = c"
+      proof -
+	from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
+	then show ?thesis by simp
+      qed
+      finally show ?thesis .
+    qed
+  qed (insert `continuous V norm f`, simp_all add: continuous_def)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/FunctionOrder.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,142 @@
+(*  Title:      HOL/Real/HahnBanach/FunctionOrder.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* An order on functions *}
+
+theory FunctionOrder
+imports Subspace Linearform
+begin
+
+subsection {* The graph of a function *}
+
+text {*
+  We define the \emph{graph} of a (real) function @{text f} with
+  domain @{text F} as the set
+  \begin{center}
+  @{text "{(x, f x). x \<in> F}"}
+  \end{center}
+  So we are modeling partial functions by specifying the domain and
+  the mapping function. We use the term ``function'' also for its
+  graph.
+*}
+
+types 'a graph = "('a \<times> real) set"
+
+definition
+  graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph" where
+  "graph F f = {(x, f x) | x. x \<in> F}"
+
+lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
+  unfolding graph_def by blast
+
+lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"
+  unfolding graph_def by blast
+
+lemma graphE [elim?]:
+    "(x, y) \<in> graph F f \<Longrightarrow> (x \<in> F \<Longrightarrow> y = f x \<Longrightarrow> C) \<Longrightarrow> C"
+  unfolding graph_def by blast
+
+
+subsection {* Functions ordered by domain extension *}
+
+text {*
+  A function @{text h'} is an extension of @{text h}, iff the graph of
+  @{text h} is a subset of the graph of @{text h'}.
+*}
+
+lemma graph_extI:
+  "(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'
+    \<Longrightarrow> graph H h \<subseteq> graph H' h'"
+  unfolding graph_def by blast
+
+lemma graph_extD1 [dest?]:
+  "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"
+  unfolding graph_def by blast
+
+lemma graph_extD2 [dest?]:
+  "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"
+  unfolding graph_def by blast
+
+
+subsection {* Domain and function of a graph *}
+
+text {*
+  The inverse functions to @{text graph} are @{text domain} and @{text
+  funct}.
+*}
+
+definition
+  "domain" :: "'a graph \<Rightarrow> 'a set" where
+  "domain g = {x. \<exists>y. (x, y) \<in> g}"
+
+definition
+  funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)" where
+  "funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"
+
+text {*
+  The following lemma states that @{text g} is the graph of a function
+  if the relation induced by @{text g} is unique.
+*}
+
+lemma graph_domain_funct:
+  assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"
+  shows "graph (domain g) (funct g) = g"
+  unfolding domain_def funct_def graph_def
+proof auto  (* FIXME !? *)
+  fix a b assume g: "(a, b) \<in> g"
+  from g show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)
+  from g show "\<exists>y. (a, y) \<in> g" ..
+  from g show "b = (SOME y. (a, y) \<in> g)"
+  proof (rule some_equality [symmetric])
+    fix y assume "(a, y) \<in> g"
+    with g show "y = b" by (rule uniq)
+  qed
+qed
+
+
+subsection {* Norm-preserving extensions of a function *}
+
+text {*
+  Given a linear form @{text f} on the space @{text F} and a seminorm
+  @{text p} on @{text E}. The set of all linear extensions of @{text
+  f}, to superspaces @{text H} of @{text F}, which are bounded by
+  @{text p}, is defined as follows.
+*}
+
+definition
+  norm_pres_extensions ::
+    "'a::{plus, minus, uminus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
+      \<Rightarrow> 'a graph set" where
+    "norm_pres_extensions E p F f
+      = {g. \<exists>H h. g = graph H h
+          \<and> linearform H h
+          \<and> H \<unlhd> E
+          \<and> F \<unlhd> H
+          \<and> graph F f \<subseteq> graph H h
+          \<and> (\<forall>x \<in> H. h x \<le> p x)}"
+
+lemma norm_pres_extensionE [elim]:
+  "g \<in> norm_pres_extensions E p F f
+  \<Longrightarrow> (\<And>H h. g = graph H h \<Longrightarrow> linearform H h
+        \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H \<Longrightarrow> graph F f \<subseteq> graph H h
+        \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x \<Longrightarrow> C) \<Longrightarrow> C"
+  unfolding norm_pres_extensions_def by blast
+
+lemma norm_pres_extensionI2 [intro]:
+  "linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H
+    \<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x
+    \<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"
+  unfolding norm_pres_extensions_def by blast
+
+lemma norm_pres_extensionI:  (* FIXME ? *)
+  "\<exists>H h. g = graph H h
+    \<and> linearform H h
+    \<and> H \<unlhd> E
+    \<and> F \<unlhd> H
+    \<and> graph F f \<subseteq> graph H h
+    \<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"
+  unfolding norm_pres_extensions_def by blast
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/HahnBanach.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,509 @@
+(*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The Hahn-Banach Theorem *}
+
+theory HahnBanach
+imports HahnBanachLemmas
+begin
+
+text {*
+  We present the proof of two different versions of the Hahn-Banach
+  Theorem, closely following \cite[\S36]{Heuser:1986}.
+*}
+
+subsection {* The Hahn-Banach Theorem for vector spaces *}
+
+text {*
+  \textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
+  vector space @{text E}, let @{text p} be a semi-norm on @{text E},
+  and @{text f} be a linear form defined on @{text F} such that @{text
+  f} is bounded by @{text p}, i.e.  @{text "\<forall>x \<in> F. f x \<le> p x"}.  Then
+  @{text f} can be extended to a linear form @{text h} on @{text E}
+  such that @{text h} is norm-preserving, i.e. @{text h} is also
+  bounded by @{text p}.
+
+  \bigskip
+  \textbf{Proof Sketch.}
+  \begin{enumerate}
+
+  \item Define @{text M} as the set of norm-preserving extensions of
+  @{text f} to subspaces of @{text E}. The linear forms in @{text M}
+  are ordered by domain extension.
+
+  \item We show that every non-empty chain in @{text M} has an upper
+  bound in @{text M}.
+
+  \item With Zorn's Lemma we conclude that there is a maximal function
+  @{text g} in @{text M}.
+
+  \item The domain @{text H} of @{text g} is the whole space @{text
+  E}, as shown by classical contradiction:
+
+  \begin{itemize}
+
+  \item Assuming @{text g} is not defined on whole @{text E}, it can
+  still be extended in a norm-preserving way to a super-space @{text
+  H'} of @{text H}.
+
+  \item Thus @{text g} can not be maximal. Contradiction!
+
+  \end{itemize}
+  \end{enumerate}
+*}
+
+theorem HahnBanach:
+  assumes E: "vectorspace E" and "subspace F E"
+    and "seminorm E p" and "linearform F f"
+  assumes fp: "\<forall>x \<in> F. f x \<le> p x"
+  shows "\<exists>h. linearform E h \<and> (\<forall>x \<in> F. h x = f x) \<and> (\<forall>x \<in> E. h x \<le> p x)"
+    -- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
+    -- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
+    -- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
+proof -
+  interpret vectorspace E by fact
+  interpret subspace F E by fact
+  interpret seminorm E p by fact
+  interpret linearform F f by fact
+  def M \<equiv> "norm_pres_extensions E p F f"
+  then have M: "M = \<dots>" by (simp only:)
+  from E have F: "vectorspace F" ..
+  note FE = `F \<unlhd> E`
+  {
+    fix c assume cM: "c \<in> chain M" and ex: "\<exists>x. x \<in> c"
+    have "\<Union>c \<in> M"
+      -- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
+      -- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
+      unfolding M_def
+    proof (rule norm_pres_extensionI)
+      let ?H = "domain (\<Union>c)"
+      let ?h = "funct (\<Union>c)"
+
+      have a: "graph ?H ?h = \<Union>c"
+      proof (rule graph_domain_funct)
+        fix x y z assume "(x, y) \<in> \<Union>c" and "(x, z) \<in> \<Union>c"
+        with M_def cM show "z = y" by (rule sup_definite)
+      qed
+      moreover from M cM a have "linearform ?H ?h"
+        by (rule sup_lf)
+      moreover from a M cM ex FE E have "?H \<unlhd> E"
+        by (rule sup_subE)
+      moreover from a M cM ex FE have "F \<unlhd> ?H"
+        by (rule sup_supF)
+      moreover from a M cM ex have "graph F f \<subseteq> graph ?H ?h"
+        by (rule sup_ext)
+      moreover from a M cM have "\<forall>x \<in> ?H. ?h x \<le> p x"
+        by (rule sup_norm_pres)
+      ultimately show "\<exists>H h. \<Union>c = graph H h
+          \<and> linearform H h
+          \<and> H \<unlhd> E
+          \<and> F \<unlhd> H
+          \<and> graph F f \<subseteq> graph H h
+          \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
+    qed
+  }
+  then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
+  -- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
+  proof (rule Zorn's_Lemma)
+      -- {* We show that @{text M} is non-empty: *}
+    show "graph F f \<in> M"
+      unfolding M_def
+    proof (rule norm_pres_extensionI2)
+      show "linearform F f" by fact
+      show "F \<unlhd> E" by fact
+      from F show "F \<unlhd> F" by (rule vectorspace.subspace_refl)
+      show "graph F f \<subseteq> graph F f" ..
+      show "\<forall>x\<in>F. f x \<le> p x" by fact
+    qed
+  qed
+  then obtain g where gM: "g \<in> M" and gx: "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
+    by blast
+  from gM obtain H h where
+      g_rep: "g = graph H h"
+    and linearform: "linearform H h"
+    and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
+    and graphs: "graph F f \<subseteq> graph H h"
+    and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
+      -- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
+      -- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
+      -- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
+  from HE E have H: "vectorspace H"
+    by (rule subspace.vectorspace)
+
+  have HE_eq: "H = E"
+    -- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
+  proof (rule classical)
+    assume neq: "H \<noteq> E"
+      -- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
+      -- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
+    have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
+    proof -
+      from HE have "H \<subseteq> E" ..
+      with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
+      obtain x': "x' \<noteq> 0"
+      proof
+        show "x' \<noteq> 0"
+        proof
+          assume "x' = 0"
+          with H have "x' \<in> H" by (simp only: vectorspace.zero)
+          with `x' \<notin> H` show False by contradiction
+        qed
+      qed
+
+      def H' \<equiv> "H + lin x'"
+        -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
+      have HH': "H \<unlhd> H'"
+      proof (unfold H'_def)
+        from x'E have "vectorspace (lin x')" ..
+        with H show "H \<unlhd> H + lin x'" ..
+      qed
+
+      obtain xi where
+        xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
+          \<and> xi \<le> p (y + x') - h y"
+        -- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
+        -- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
+           \label{ex-xi-use}\skp *}
+      proof -
+        from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
+            \<and> xi \<le> p (y + x') - h y"
+        proof (rule ex_xi)
+          fix u v assume u: "u \<in> H" and v: "v \<in> H"
+          with HE have uE: "u \<in> E" and vE: "v \<in> E" by auto
+          from H u v linearform have "h v - h u = h (v - u)"
+            by (simp add: linearform.diff)
+          also from hp and H u v have "\<dots> \<le> p (v - u)"
+            by (simp only: vectorspace.diff_closed)
+          also from x'E uE vE have "v - u = x' + - x' + v + - u"
+            by (simp add: diff_eq1)
+          also from x'E uE vE have "\<dots> = v + x' + - (u + x')"
+            by (simp add: add_ac)
+          also from x'E uE vE have "\<dots> = (v + x') - (u + x')"
+            by (simp add: diff_eq1)
+          also from x'E uE vE E have "p \<dots> \<le> p (v + x') + p (u + x')"
+            by (simp add: diff_subadditive)
+          finally have "h v - h u \<le> p (v + x') + p (u + x')" .
+          then show "- p (u + x') - h u \<le> p (v + x') - h v" by simp
+        qed
+        then show thesis by (blast intro: that)
+      qed
+
+      def h' \<equiv> "\<lambda>x. let (y, a) =
+          SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
+        -- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
+
+      have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
+        -- {* @{text h'} is an extension of @{text h} \dots \skp *}
+      proof
+        show "g \<subseteq> graph H' h'"
+        proof -
+          have  "graph H h \<subseteq> graph H' h'"
+          proof (rule graph_extI)
+            fix t assume t: "t \<in> H"
+            from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
+	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
+            with h'_def show "h t = h' t" by (simp add: Let_def)
+          next
+            from HH' show "H \<subseteq> H'" ..
+          qed
+          with g_rep show ?thesis by (simp only:)
+        qed
+
+        show "g \<noteq> graph H' h'"
+        proof -
+          have "graph H h \<noteq> graph H' h'"
+          proof
+            assume eq: "graph H h = graph H' h'"
+            have "x' \<in> H'"
+	      unfolding H'_def
+            proof
+              from H show "0 \<in> H" by (rule vectorspace.zero)
+              from x'E show "x' \<in> lin x'" by (rule x_lin_x)
+              from x'E show "x' = 0 + x'" by simp
+            qed
+            then have "(x', h' x') \<in> graph H' h'" ..
+            with eq have "(x', h' x') \<in> graph H h" by (simp only:)
+            then have "x' \<in> H" ..
+            with `x' \<notin> H` show False by contradiction
+          qed
+          with g_rep show ?thesis by simp
+        qed
+      qed
+      moreover have "graph H' h' \<in> M"
+        -- {* and @{text h'} is norm-preserving. \skp *}
+      proof (unfold M_def)
+        show "graph H' h' \<in> norm_pres_extensions E p F f"
+        proof (rule norm_pres_extensionI2)
+          show "linearform H' h'"
+	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
+	    by (rule h'_lf)
+          show "H' \<unlhd> E"
+	  unfolding H'_def
+          proof
+            show "H \<unlhd> E" by fact
+            show "vectorspace E" by fact
+            from x'E show "lin x' \<unlhd> E" ..
+          qed
+          from H `F \<unlhd> H` HH' show FH': "F \<unlhd> H'"
+            by (rule vectorspace.subspace_trans)
+          show "graph F f \<subseteq> graph H' h'"
+          proof (rule graph_extI)
+            fix x assume x: "x \<in> F"
+            with graphs have "f x = h x" ..
+            also have "\<dots> = h x + 0 * xi" by simp
+            also have "\<dots> = (let (y, a) = (x, 0) in h y + a * xi)"
+              by (simp add: Let_def)
+            also have "(x, 0) =
+                (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
+	      using E HE
+            proof (rule decomp_H'_H [symmetric])
+              from FH x show "x \<in> H" ..
+              from x' show "x' \<noteq> 0" .
+	      show "x' \<notin> H" by fact
+	      show "x' \<in> E" by fact
+            qed
+            also have
+              "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
+              in h y + a * xi) = h' x" by (simp only: h'_def)
+            finally show "f x = h' x" .
+          next
+            from FH' show "F \<subseteq> H'" ..
+          qed
+          show "\<forall>x \<in> H'. h' x \<le> p x"
+	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
+	      `seminorm E p` linearform and hp xi
+	    by (rule h'_norm_pres)
+        qed
+      qed
+      ultimately show ?thesis ..
+    qed
+    then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
+      -- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
+    with gx show "H = E" by contradiction
+  qed
+
+  from HE_eq and linearform have "linearform E h"
+    by (simp only:)
+  moreover have "\<forall>x \<in> F. h x = f x"
+  proof
+    fix x assume "x \<in> F"
+    with graphs have "f x = h x" ..
+    then show "h x = f x" ..
+  qed
+  moreover from HE_eq and hp have "\<forall>x \<in> E. h x \<le> p x"
+    by (simp only:)
+  ultimately show ?thesis by blast
+qed
+
+
+subsection  {* Alternative formulation *}
+
+text {*
+  The following alternative formulation of the Hahn-Banach
+  Theorem\label{abs-HahnBanach} uses the fact that for a real linear
+  form @{text f} and a seminorm @{text p} the following inequations
+  are equivalent:\footnote{This was shown in lemma @{thm [source]
+  abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
+  \begin{center}
+  \begin{tabular}{lll}
+  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
+  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
+  \end{tabular}
+  \end{center}
+*}
+
+theorem abs_HahnBanach:
+  assumes E: "vectorspace E" and FE: "subspace F E"
+    and lf: "linearform F f" and sn: "seminorm E p"
+  assumes fp: "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
+  shows "\<exists>g. linearform E g
+    \<and> (\<forall>x \<in> F. g x = f x)
+    \<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
+proof -
+  interpret vectorspace E by fact
+  interpret subspace F E by fact
+  interpret linearform F f by fact
+  interpret seminorm E p by fact
+  have "\<exists>g. linearform E g \<and> (\<forall>x \<in> F. g x = f x) \<and> (\<forall>x \<in> E. g x \<le> p x)"
+    using E FE sn lf
+  proof (rule HahnBanach)
+    show "\<forall>x \<in> F. f x \<le> p x"
+      using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
+  qed
+  then obtain g where lg: "linearform E g" and *: "\<forall>x \<in> F. g x = f x"
+      and **: "\<forall>x \<in> E. g x \<le> p x" by blast
+  have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
+    using _ E sn lg **
+  proof (rule abs_ineq_iff [THEN iffD2])
+    show "E \<unlhd> E" ..
+  qed
+  with lg * show ?thesis by blast
+qed
+
+
+subsection {* The Hahn-Banach Theorem for normed spaces *}
+
+text {*
+  Every continuous linear form @{text f} on a subspace @{text F} of a
+  norm space @{text E}, can be extended to a continuous linear form
+  @{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
+*}
+
+theorem norm_HahnBanach:
+  fixes V and norm ("\<parallel>_\<parallel>")
+  fixes B defines "\<And>V f. B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
+  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
+  defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
+  assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
+    and linearform: "linearform F f" and "continuous F norm f"
+  shows "\<exists>g. linearform E g
+     \<and> continuous E norm g
+     \<and> (\<forall>x \<in> F. g x = f x)
+     \<and> \<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
+proof -
+  interpret normed_vectorspace E norm by fact
+  interpret normed_vectorspace_with_fn_norm E norm B fn_norm
+    by (auto simp: B_def fn_norm_def) intro_locales
+  interpret subspace F E by fact
+  interpret linearform F f by fact
+  interpret continuous F norm f by fact
+  have E: "vectorspace E" by intro_locales
+  have F: "vectorspace F" by rule intro_locales
+  have F_norm: "normed_vectorspace F norm"
+    using FE E_norm by (rule subspace_normed_vs)
+  have ge_zero: "0 \<le> \<parallel>f\<parallel>\<hyphen>F"
+    by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
+      [OF normed_vectorspace_with_fn_norm.intro,
+       OF F_norm `continuous F norm f` , folded B_def fn_norm_def])
+  txt {* We define a function @{text p} on @{text E} as follows:
+    @{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"} *}
+  def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+
+  txt {* @{text p} is a seminorm on @{text E}: *}
+  have q: "seminorm E p"
+  proof
+    fix x y a assume x: "x \<in> E" and y: "y \<in> E"
+    
+    txt {* @{text p} is positive definite: *}
+    have "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
+    moreover from x have "0 \<le> \<parallel>x\<parallel>" ..
+    ultimately show "0 \<le> p x"  
+      by (simp add: p_def zero_le_mult_iff)
+
+    txt {* @{text p} is absolutely homogenous: *}
+
+    show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
+    proof -
+      have "p (a \<cdot> x) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>a \<cdot> x\<parallel>" by (simp only: p_def)
+      also from x have "\<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>" by (rule abs_homogenous)
+      also have "\<parallel>f\<parallel>\<hyphen>F * (\<bar>a\<bar> * \<parallel>x\<parallel>) = \<bar>a\<bar> * (\<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>)" by simp
+      also have "\<dots> = \<bar>a\<bar> * p x" by (simp only: p_def)
+      finally show ?thesis .
+    qed
+
+    txt {* Furthermore, @{text p} is subadditive: *}
+
+    show "p (x + y) \<le> p x + p y"
+    proof -
+      have "p (x + y) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel>" by (simp only: p_def)
+      also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
+      from x y have "\<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>" ..
+      with a have " \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>F * (\<parallel>x\<parallel> + \<parallel>y\<parallel>)"
+        by (simp add: mult_left_mono)
+      also have "\<dots> = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel> + \<parallel>f\<parallel>\<hyphen>F * \<parallel>y\<parallel>" by (simp only: right_distrib)
+      also have "\<dots> = p x + p y" by (simp only: p_def)
+      finally show ?thesis .
+    qed
+  qed
+
+  txt {* @{text f} is bounded by @{text p}. *}
+
+  have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
+  proof
+    fix x assume "x \<in> F"
+    with `continuous F norm f` and linearform
+    show "\<bar>f x\<bar> \<le> p x"
+      unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
+        [OF normed_vectorspace_with_fn_norm.intro,
+         OF F_norm, folded B_def fn_norm_def])
+  qed
+
+  txt {* Using the fact that @{text p} is a seminorm and @{text f} is bounded
+    by @{text p} we can apply the Hahn-Banach Theorem for real vector
+    spaces. So @{text f} can be extended in a norm-preserving way to
+    some function @{text g} on the whole vector space @{text E}. *}
+
+  with E FE linearform q obtain g where
+      linearformE: "linearform E g"
+    and a: "\<forall>x \<in> F. g x = f x"
+    and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
+    by (rule abs_HahnBanach [elim_format]) iprover
+
+  txt {* We furthermore have to show that @{text g} is also continuous: *}
+
+  have g_cont: "continuous E norm g" using linearformE
+  proof
+    fix x assume "x \<in> E"
+    with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+      by (simp only: p_def)
+  qed
+
+  txt {* To complete the proof, we show that @{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. *}
+
+  have "\<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
+  proof (rule order_antisym)
+    txt {*
+      First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}.  The function norm @{text
+      "\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
+      \begin{center}
+      \begin{tabular}{l}
+      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
+      \end{tabular}
+      \end{center}
+      \noindent Furthermore holds
+      \begin{center}
+      \begin{tabular}{l}
+      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
+      \end{tabular}
+      \end{center}
+    *}
+
+    have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+    proof
+      fix x assume "x \<in> E"
+      with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+        by (simp only: p_def)
+    qed
+    from g_cont this ge_zero
+    show "\<parallel>g\<parallel>\<hyphen>E \<le> \<parallel>f\<parallel>\<hyphen>F"
+      by (rule fn_norm_least [of g, folded B_def fn_norm_def])
+
+    txt {* The other direction is achieved by a similar argument. *}
+
+    show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
+    proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
+	[OF normed_vectorspace_with_fn_norm.intro,
+	 OF F_norm, folded B_def fn_norm_def])
+      show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
+      proof
+	fix x assume x: "x \<in> F"
+	from a x have "g x = f x" ..
+	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
+	also from g_cont
+	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
+	proof (rule fn_norm_le_cong [OF _ linearformE, folded B_def fn_norm_def])
+	  from FE x show "x \<in> E" ..
+	qed
+	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
+      qed
+      show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
+	using g_cont
+	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
+      show "continuous F norm f" by fact
+    qed
+  qed
+  with linearformE a g_cont show ?thesis by blast
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/HahnBanachExtLemmas.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,280 @@
+(*  Title:      HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Extending non-maximal functions *}
+
+theory HahnBanachExtLemmas
+imports FunctionNorm
+begin
+
+text {*
+  In this section the following context is presumed.  Let @{text E} be
+  a real vector space with a seminorm @{text q} on @{text E}. @{text
+  F} is a subspace of @{text E} and @{text f} a linear function on
+  @{text F}. We consider a subspace @{text H} of @{text E} that is a
+  superspace of @{text F} and a linear form @{text h} on @{text
+  H}. @{text H} is a not equal to @{text E} and @{text "x\<^sub>0"} is
+  an element in @{text "E - H"}.  @{text H} is extended to the direct
+  sum @{text "H' = H + lin x\<^sub>0"}, so for any @{text "x \<in> H'"}
+  the decomposition of @{text "x = y + a \<cdot> x"} with @{text "y \<in> H"} is
+  unique. @{text h'} is defined on @{text H'} by @{text "h' x = h y +
+  a \<cdot> \<xi>"} for a certain @{text \<xi>}.
+
+  Subsequently we show some properties of this extension @{text h'} of
+  @{text h}.
+
+  \medskip This lemma will be used to show the existence of a linear
+  extension of @{text f} (see page \pageref{ex-xi-use}). It is a
+  consequence of the completeness of @{text \<real>}. To show
+  \begin{center}
+  \begin{tabular}{l}
+  @{text "\<exists>\<xi>. \<forall>y \<in> F. a y \<le> \<xi> \<and> \<xi> \<le> b y"}
+  \end{tabular}
+  \end{center}
+  \noindent it suffices to show that
+  \begin{center}
+  \begin{tabular}{l}
+  @{text "\<forall>u \<in> F. \<forall>v \<in> F. a u \<le> b v"}
+  \end{tabular}
+  \end{center}
+*}
+
+lemma ex_xi:
+  assumes "vectorspace F"
+  assumes r: "\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> b v"
+  shows "\<exists>xi::real. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
+proof -
+  interpret vectorspace F by fact
+  txt {* From the completeness of the reals follows:
+    The set @{text "S = {a u. u \<in> F}"} has a supremum, if it is
+    non-empty and has an upper bound. *}
+
+  let ?S = "{a u | u. u \<in> F}"
+  have "\<exists>xi. lub ?S xi"
+  proof (rule real_complete)
+    have "a 0 \<in> ?S" by blast
+    then show "\<exists>X. X \<in> ?S" ..
+    have "\<forall>y \<in> ?S. y \<le> b 0"
+    proof
+      fix y assume y: "y \<in> ?S"
+      then obtain u where u: "u \<in> F" and y: "y = a u" by blast
+      from u and zero have "a u \<le> b 0" by (rule r)
+      with y show "y \<le> b 0" by (simp only:)
+    qed
+    then show "\<exists>u. \<forall>y \<in> ?S. y \<le> u" ..
+  qed
+  then obtain xi where xi: "lub ?S xi" ..
+  {
+    fix y assume "y \<in> F"
+    then have "a y \<in> ?S" by blast
+    with xi have "a y \<le> xi" by (rule lub.upper)
+  } moreover {
+    fix y assume y: "y \<in> F"
+    from xi have "xi \<le> b y"
+    proof (rule lub.least)
+      fix au assume "au \<in> ?S"
+      then obtain u where u: "u \<in> F" and au: "au = a u" by blast
+      from u y have "a u \<le> b y" by (rule r)
+      with au show "au \<le> b y" by (simp only:)
+    qed
+  } ultimately show "\<exists>xi. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y" by blast
+qed
+
+text {*
+  \medskip The function @{text h'} is defined as a @{text "h' x = h y
+  + a \<cdot> \<xi>"} where @{text "x = y + a \<cdot> \<xi>"} is a linear extension of
+  @{text h} to @{text H'}.
+*}
+
+lemma h'_lf:
+  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
+      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
+    and H'_def: "H' \<equiv> H + lin x0"
+    and HE: "H \<unlhd> E"
+  assumes "linearform H h"
+  assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
+  assumes E: "vectorspace E"
+  shows "linearform H' h'"
+proof -
+  interpret linearform H h by fact
+  interpret vectorspace E by fact
+  show ?thesis
+  proof
+    note E = `vectorspace E`
+    have H': "vectorspace H'"
+    proof (unfold H'_def)
+      from `x0 \<in> E`
+      have "lin x0 \<unlhd> E" ..
+      with HE show "vectorspace (H + lin x0)" using E ..
+    qed
+    {
+      fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
+      show "h' (x1 + x2) = h' x1 + h' x2"
+      proof -
+	from H' x1 x2 have "x1 + x2 \<in> H'"
+          by (rule vectorspace.add_closed)
+	with x1 x2 obtain y y1 y2 a a1 a2 where
+          x1x2: "x1 + x2 = y + a \<cdot> x0" and y: "y \<in> H"
+          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
+          and x2_rep: "x2 = y2 + a2 \<cdot> x0" and y2: "y2 \<in> H"
+          unfolding H'_def sum_def lin_def by blast
+	
+	have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
+	proof (rule decomp_H') txt_raw {* \label{decomp-H-use} *}
+          from HE y1 y2 show "y1 + y2 \<in> H"
+            by (rule subspace.add_closed)
+          from x0 and HE y y1 y2
+          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E"  "y2 \<in> E" by auto
+          with x1_rep x2_rep have "(y1 + y2) + (a1 + a2) \<cdot> x0 = x1 + x2"
+            by (simp add: add_ac add_mult_distrib2)
+          also note x1x2
+          finally show "(y1 + y2) + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0" .
+	qed
+	
+	from h'_def x1x2 E HE y x0
+	have "h' (x1 + x2) = h y + a * xi"
+          by (rule h'_definite)
+	also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
+          by (simp only: ya)
+	also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
+          by simp
+	also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
+          by (simp add: left_distrib)
+	also from h'_def x1_rep E HE y1 x0
+	have "h y1 + a1 * xi = h' x1"
+          by (rule h'_definite [symmetric])
+	also from h'_def x2_rep E HE y2 x0
+	have "h y2 + a2 * xi = h' x2"
+          by (rule h'_definite [symmetric])
+	finally show ?thesis .
+      qed
+    next
+      fix x1 c assume x1: "x1 \<in> H'"
+      show "h' (c \<cdot> x1) = c * (h' x1)"
+      proof -
+	from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
+          by (rule vectorspace.mult_closed)
+	with x1 obtain y a y1 a1 where
+            cx1_rep: "c \<cdot> x1 = y + a \<cdot> x0" and y: "y \<in> H"
+          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
+          unfolding H'_def sum_def lin_def by blast
+	
+	have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
+	proof (rule decomp_H')
+          from HE y1 show "c \<cdot> y1 \<in> H"
+            by (rule subspace.mult_closed)
+          from x0 and HE y y1
+          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E" by auto
+          with x1_rep have "c \<cdot> y1 + (c * a1) \<cdot> x0 = c \<cdot> x1"
+            by (simp add: mult_assoc add_mult_distrib1)
+          also note cx1_rep
+          finally show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0" .
+	qed
+	
+	from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
+          by (rule h'_definite)
+	also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
+          by (simp only: ya)
+	also from y1 have "h (c \<cdot> y1) = c * h y1"
+          by simp
+	also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
+          by (simp only: right_distrib)
+	also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
+          by (rule h'_definite [symmetric])
+	finally show ?thesis .
+      qed
+    }
+  qed
+qed
+
+text {* \medskip The linear extension @{text h'} of @{text h}
+  is bounded by the seminorm @{text p}. *}
+
+lemma h'_norm_pres:
+  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
+      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
+    and H'_def: "H' \<equiv> H + lin x0"
+    and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
+  assumes E: "vectorspace E" and HE: "subspace H E"
+    and "seminorm E p" and "linearform H h"
+  assumes a: "\<forall>y \<in> H. h y \<le> p y"
+    and a': "\<forall>y \<in> H. - p (y + x0) - h y \<le> xi \<and> xi \<le> p (y + x0) - h y"
+  shows "\<forall>x \<in> H'. h' x \<le> p x"
+proof -
+  interpret vectorspace E by fact
+  interpret subspace H E by fact
+  interpret seminorm E p by fact
+  interpret linearform H h by fact
+  show ?thesis
+  proof
+    fix x assume x': "x \<in> H'"
+    show "h' x \<le> p x"
+    proof -
+      from a' have a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
+	and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
+      from x' obtain y a where
+          x_rep: "x = y + a \<cdot> x0" and y: "y \<in> H"
+	unfolding H'_def sum_def lin_def by blast
+      from y have y': "y \<in> E" ..
+      from y have ay: "inverse a \<cdot> y \<in> H" by simp
+      
+      from h'_def x_rep E HE y x0 have "h' x = h y + a * xi"
+	by (rule h'_definite)
+      also have "\<dots> \<le> p (y + a \<cdot> x0)"
+      proof (rule linorder_cases)
+	assume z: "a = 0"
+	then have "h y + a * xi = h y" by simp
+	also from a y have "\<dots> \<le> p y" ..
+	also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
+	finally show ?thesis .
+      next
+	txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
+          with @{text ya} taken as @{text "y / a"}: *}
+	assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
+	from a1 ay
+	have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
+	with lz have "a * xi \<le>
+          a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
+          by (simp add: mult_left_mono_neg order_less_imp_le)
+	
+	also have "\<dots> =
+          - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
+	  by (simp add: right_diff_distrib)
+	also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
+          p (a \<cdot> (inverse a \<cdot> y + x0))"
+          by (simp add: abs_homogenous)
+	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
+          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
+	also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
+          by simp
+	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
+	then show ?thesis by simp
+      next
+	txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
+          with @{text ya} taken as @{text "y / a"}: *}
+	assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
+	from a2 ay
+	have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
+	with gz have "a * xi \<le>
+          a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
+          by simp
+	also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
+	  by (simp add: right_diff_distrib)
+	also from gz x0 y'
+	have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
+          by (simp add: abs_homogenous)
+	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
+          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
+	also from nz y have "a * h (inverse a \<cdot> y) = h y"
+          by simp
+	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
+	then show ?thesis by simp
+      qed
+      also from x_rep have "\<dots> = p x" by (simp only:)
+      finally show ?thesis .
+    qed
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/HahnBanachLemmas.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,4 @@
+(*<*)
+theory HahnBanachLemmas imports HahnBanachSupLemmas HahnBanachExtLemmas begin
+end
+(*>*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/HahnBanachSupLemmas.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,446 @@
+(*  Title:      HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The supremum w.r.t.~the function order *}
+
+theory HahnBanachSupLemmas
+imports FunctionNorm ZornLemma
+begin
+
+text {*
+  This section contains some lemmas that will be used in the proof of
+  the Hahn-Banach Theorem.  In this section the following context is
+  presumed.  Let @{text E} be a real vector space with a seminorm
+  @{text p} on @{text E}.  @{text F} is a subspace of @{text E} and
+  @{text f} a linear form on @{text F}. We consider a chain @{text c}
+  of norm-preserving extensions of @{text f}, such that @{text "\<Union>c =
+  graph H h"}.  We will show some properties about the limit function
+  @{text h}, i.e.\ the supremum of the chain @{text c}.
+
+  \medskip Let @{text c} be a chain of norm-preserving extensions of
+  the function @{text f} and let @{text "graph H h"} be the supremum
+  of @{text c}.  Every element in @{text H} is member of one of the
+  elements of the chain.
+*}
+lemmas [dest?] = chainD
+lemmas chainE2 [elim?] = chainD2 [elim_format, standard]
+
+lemma some_H'h't:
+  assumes M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and u: "graph H h = \<Union>c"
+    and x: "x \<in> H"
+  shows "\<exists>H' h'. graph H' h' \<in> c
+    \<and> (x, h x) \<in> graph H' h'
+    \<and> linearform H' h' \<and> H' \<unlhd> E
+    \<and> F \<unlhd> H' \<and> graph F f \<subseteq> graph H' h'
+    \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
+proof -
+  from x have "(x, h x) \<in> graph H h" ..
+  also from u have "\<dots> = \<Union>c" .
+  finally obtain g where gc: "g \<in> c" and gh: "(x, h x) \<in> g" by blast
+
+  from cM have "c \<subseteq> M" ..
+  with gc have "g \<in> M" ..
+  also from M have "\<dots> = norm_pres_extensions E p F f" .
+  finally obtain H' and h' where g: "g = graph H' h'"
+    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
+      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x" ..
+
+  from gc and g have "graph H' h' \<in> c" by (simp only:)
+  moreover from gh and g have "(x, h x) \<in> graph H' h'" by (simp only:)
+  ultimately show ?thesis using * by blast
+qed
+
+text {*
+  \medskip Let @{text c} be a chain of norm-preserving extensions of
+  the function @{text f} and let @{text "graph H h"} be the supremum
+  of @{text c}.  Every element in the domain @{text H} of the supremum
+  function is member of the domain @{text H'} of some function @{text
+  h'}, such that @{text h} extends @{text h'}.
+*}
+
+lemma some_H'h':
+  assumes M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and u: "graph H h = \<Union>c"
+    and x: "x \<in> H"
+  shows "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
+    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
+    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
+proof -
+  from M cM u x obtain H' h' where
+      x_hx: "(x, h x) \<in> graph H' h'"
+    and c: "graph H' h' \<in> c"
+    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
+      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
+    by (rule some_H'h't [elim_format]) blast
+  from x_hx have "x \<in> H'" ..
+  moreover from cM u c have "graph H' h' \<subseteq> graph H h"
+    by (simp only: chain_ball_Union_upper)
+  ultimately show ?thesis using * by blast
+qed
+
+text {*
+  \medskip Any two elements @{text x} and @{text y} in the domain
+  @{text H} of the supremum function @{text h} are both in the domain
+  @{text H'} of some function @{text h'}, such that @{text h} extends
+  @{text h'}.
+*}
+
+lemma some_H'h'2:
+  assumes M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and u: "graph H h = \<Union>c"
+    and x: "x \<in> H"
+    and y: "y \<in> H"
+  shows "\<exists>H' h'. x \<in> H' \<and> y \<in> H'
+    \<and> graph H' h' \<subseteq> graph H h
+    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
+    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
+proof -
+  txt {* @{text y} is in the domain @{text H''} of some function @{text h''},
+  such that @{text h} extends @{text h''}. *}
+
+  from M cM u and y obtain H' h' where
+      y_hy: "(y, h y) \<in> graph H' h'"
+    and c': "graph H' h' \<in> c"
+    and * :
+      "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
+      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
+    by (rule some_H'h't [elim_format]) blast
+
+  txt {* @{text x} is in the domain @{text H'} of some function @{text h'},
+    such that @{text h} extends @{text h'}. *}
+
+  from M cM u and x obtain H'' h'' where
+      x_hx: "(x, h x) \<in> graph H'' h''"
+    and c'': "graph H'' h'' \<in> c"
+    and ** :
+      "linearform H'' h''"  "H'' \<unlhd> E"  "F \<unlhd> H''"
+      "graph F f \<subseteq> graph H'' h''"  "\<forall>x \<in> H''. h'' x \<le> p x"
+    by (rule some_H'h't [elim_format]) blast
+
+  txt {* Since both @{text h'} and @{text h''} are elements of the chain,
+    @{text h''} is an extension of @{text h'} or vice versa. Thus both
+    @{text x} and @{text y} are contained in the greater
+    one. \label{cases1}*}
+
+  from cM c'' c' have "graph H'' h'' \<subseteq> graph H' h' \<or> graph H' h' \<subseteq> graph H'' h''"
+    (is "?case1 \<or> ?case2") ..
+  then show ?thesis
+  proof
+    assume ?case1
+    have "(x, h x) \<in> graph H'' h''" by fact
+    also have "\<dots> \<subseteq> graph H' h'" by fact
+    finally have xh:"(x, h x) \<in> graph H' h'" .
+    then have "x \<in> H'" ..
+    moreover from y_hy have "y \<in> H'" ..
+    moreover from cM u and c' have "graph H' h' \<subseteq> graph H h"
+      by (simp only: chain_ball_Union_upper)
+    ultimately show ?thesis using * by blast
+  next
+    assume ?case2
+    from x_hx have "x \<in> H''" ..
+    moreover {
+      have "(y, h y) \<in> graph H' h'" by (rule y_hy)
+      also have "\<dots> \<subseteq> graph H'' h''" by fact
+      finally have "(y, h y) \<in> graph H'' h''" .
+    } then have "y \<in> H''" ..
+    moreover from cM u and c'' have "graph H'' h'' \<subseteq> graph H h"
+      by (simp only: chain_ball_Union_upper)
+    ultimately show ?thesis using ** by blast
+  qed
+qed
+
+text {*
+  \medskip The relation induced by the graph of the supremum of a
+  chain @{text c} is definite, i.~e.~t is the graph of a function.
+*}
+
+lemma sup_definite:
+  assumes M_def: "M \<equiv> norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and xy: "(x, y) \<in> \<Union>c"
+    and xz: "(x, z) \<in> \<Union>c"
+  shows "z = y"
+proof -
+  from cM have c: "c \<subseteq> M" ..
+  from xy obtain G1 where xy': "(x, y) \<in> G1" and G1: "G1 \<in> c" ..
+  from xz obtain G2 where xz': "(x, z) \<in> G2" and G2: "G2 \<in> c" ..
+
+  from G1 c have "G1 \<in> M" ..
+  then obtain H1 h1 where G1_rep: "G1 = graph H1 h1"
+    unfolding M_def by blast
+
+  from G2 c have "G2 \<in> M" ..
+  then obtain H2 h2 where G2_rep: "G2 = graph H2 h2"
+    unfolding M_def by blast
+
+  txt {* @{text "G\<^sub>1"} is contained in @{text "G\<^sub>2"}
+    or vice versa, since both @{text "G\<^sub>1"} and @{text
+    "G\<^sub>2"} are members of @{text c}. \label{cases2}*}
+
+  from cM G1 G2 have "G1 \<subseteq> G2 \<or> G2 \<subseteq> G1" (is "?case1 \<or> ?case2") ..
+  then show ?thesis
+  proof
+    assume ?case1
+    with xy' G2_rep have "(x, y) \<in> graph H2 h2" by blast
+    then have "y = h2 x" ..
+    also
+    from xz' G2_rep have "(x, z) \<in> graph H2 h2" by (simp only:)
+    then have "z = h2 x" ..
+    finally show ?thesis .
+  next
+    assume ?case2
+    with xz' G1_rep have "(x, z) \<in> graph H1 h1" by blast
+    then have "z = h1 x" ..
+    also
+    from xy' G1_rep have "(x, y) \<in> graph H1 h1" by (simp only:)
+    then have "y = h1 x" ..
+    finally show ?thesis ..
+  qed
+qed
+
+text {*
+  \medskip The limit function @{text h} is linear. Every element
+  @{text x} in the domain of @{text h} is in the domain of a function
+  @{text h'} in the chain of norm preserving extensions.  Furthermore,
+  @{text h} is an extension of @{text h'} so the function values of
+  @{text x} are identical for @{text h'} and @{text h}.  Finally, the
+  function @{text h'} is linear by construction of @{text M}.
+*}
+
+lemma sup_lf:
+  assumes M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and u: "graph H h = \<Union>c"
+  shows "linearform H h"
+proof
+  fix x y assume x: "x \<in> H" and y: "y \<in> H"
+  with M cM u obtain H' h' where
+        x': "x \<in> H'" and y': "y \<in> H'"
+      and b: "graph H' h' \<subseteq> graph H h"
+      and linearform: "linearform H' h'"
+      and subspace: "H' \<unlhd> E"
+    by (rule some_H'h'2 [elim_format]) blast
+
+  show "h (x + y) = h x + h y"
+  proof -
+    from linearform x' y' have "h' (x + y) = h' x + h' y"
+      by (rule linearform.add)
+    also from b x' have "h' x = h x" ..
+    also from b y' have "h' y = h y" ..
+    also from subspace x' y' have "x + y \<in> H'"
+      by (rule subspace.add_closed)
+    with b have "h' (x + y) = h (x + y)" ..
+    finally show ?thesis .
+  qed
+next
+  fix x a assume x: "x \<in> H"
+  with M cM u obtain H' h' where
+        x': "x \<in> H'"
+      and b: "graph H' h' \<subseteq> graph H h"
+      and linearform: "linearform H' h'"
+      and subspace: "H' \<unlhd> E"
+    by (rule some_H'h' [elim_format]) blast
+
+  show "h (a \<cdot> x) = a * h x"
+  proof -
+    from linearform x' have "h' (a \<cdot> x) = a * h' x"
+      by (rule linearform.mult)
+    also from b x' have "h' x = h x" ..
+    also from subspace x' have "a \<cdot> x \<in> H'"
+      by (rule subspace.mult_closed)
+    with b have "h' (a \<cdot> x) = h (a \<cdot> x)" ..
+    finally show ?thesis .
+  qed
+qed
+
+text {*
+  \medskip The limit of a non-empty chain of norm preserving
+  extensions of @{text f} is an extension of @{text f}, since every
+  element of the chain is an extension of @{text f} and the supremum
+  is an extension for every element of the chain.
+*}
+
+lemma sup_ext:
+  assumes graph: "graph H h = \<Union>c"
+    and M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and ex: "\<exists>x. x \<in> c"
+  shows "graph F f \<subseteq> graph H h"
+proof -
+  from ex obtain x where xc: "x \<in> c" ..
+  from cM have "c \<subseteq> M" ..
+  with xc have "x \<in> M" ..
+  with M have "x \<in> norm_pres_extensions E p F f"
+    by (simp only:)
+  then obtain G g where "x = graph G g" and "graph F f \<subseteq> graph G g" ..
+  then have "graph F f \<subseteq> x" by (simp only:)
+  also from xc have "\<dots> \<subseteq> \<Union>c" by blast
+  also from graph have "\<dots> = graph H h" ..
+  finally show ?thesis .
+qed
+
+text {*
+  \medskip The domain @{text H} of the limit function is a superspace
+  of @{text F}, since @{text F} is a subset of @{text H}. The
+  existence of the @{text 0} element in @{text F} and the closure
+  properties follow from the fact that @{text F} is a vector space.
+*}
+
+lemma sup_supF:
+  assumes graph: "graph H h = \<Union>c"
+    and M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and ex: "\<exists>x. x \<in> c"
+    and FE: "F \<unlhd> E"
+  shows "F \<unlhd> H"
+proof
+  from FE show "F \<noteq> {}" by (rule subspace.non_empty)
+  from graph M cM ex have "graph F f \<subseteq> graph H h" by (rule sup_ext)
+  then show "F \<subseteq> H" ..
+  fix x y assume "x \<in> F" and "y \<in> F"
+  with FE show "x + y \<in> F" by (rule subspace.add_closed)
+next
+  fix x a assume "x \<in> F"
+  with FE show "a \<cdot> x \<in> F" by (rule subspace.mult_closed)
+qed
+
+text {*
+  \medskip The domain @{text H} of the limit function is a subspace of
+  @{text E}.
+*}
+
+lemma sup_subE:
+  assumes graph: "graph H h = \<Union>c"
+    and M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and ex: "\<exists>x. x \<in> c"
+    and FE: "F \<unlhd> E"
+    and E: "vectorspace E"
+  shows "H \<unlhd> E"
+proof
+  show "H \<noteq> {}"
+  proof -
+    from FE E have "0 \<in> F" by (rule subspace.zero)
+    also from graph M cM ex FE have "F \<unlhd> H" by (rule sup_supF)
+    then have "F \<subseteq> H" ..
+    finally show ?thesis by blast
+  qed
+  show "H \<subseteq> E"
+  proof
+    fix x assume "x \<in> H"
+    with M cM graph
+    obtain H' h' where x: "x \<in> H'" and H'E: "H' \<unlhd> E"
+      by (rule some_H'h' [elim_format]) blast
+    from H'E have "H' \<subseteq> E" ..
+    with x show "x \<in> E" ..
+  qed
+  fix x y assume x: "x \<in> H" and y: "y \<in> H"
+  show "x + y \<in> H"
+  proof -
+    from M cM graph x y obtain H' h' where
+          x': "x \<in> H'" and y': "y \<in> H'" and H'E: "H' \<unlhd> E"
+        and graphs: "graph H' h' \<subseteq> graph H h"
+      by (rule some_H'h'2 [elim_format]) blast
+    from H'E x' y' have "x + y \<in> H'"
+      by (rule subspace.add_closed)
+    also from graphs have "H' \<subseteq> H" ..
+    finally show ?thesis .
+  qed
+next
+  fix x a assume x: "x \<in> H"
+  show "a \<cdot> x \<in> H"
+  proof -
+    from M cM graph x
+    obtain H' h' where x': "x \<in> H'" and H'E: "H' \<unlhd> E"
+        and graphs: "graph H' h' \<subseteq> graph H h"
+      by (rule some_H'h' [elim_format]) blast
+    from H'E x' have "a \<cdot> x \<in> H'" by (rule subspace.mult_closed)
+    also from graphs have "H' \<subseteq> H" ..
+    finally show ?thesis .
+  qed
+qed
+
+text {*
+  \medskip The limit function is bounded by the norm @{text p} as
+  well, since all elements in the chain are bounded by @{text p}.
+*}
+
+lemma sup_norm_pres:
+  assumes graph: "graph H h = \<Union>c"
+    and M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+  shows "\<forall>x \<in> H. h x \<le> p x"
+proof
+  fix x assume "x \<in> H"
+  with M cM graph obtain H' h' where x': "x \<in> H'"
+      and graphs: "graph H' h' \<subseteq> graph H h"
+      and a: "\<forall>x \<in> H'. h' x \<le> p x"
+    by (rule some_H'h' [elim_format]) blast
+  from graphs x' have [symmetric]: "h' x = h x" ..
+  also from a x' have "h' x \<le> p x " ..
+  finally show "h x \<le> p x" .
+qed
+
+text {*
+  \medskip The following lemma is a property of linear forms on real
+  vector spaces. It will be used for the lemma @{text abs_HahnBanach}
+  (see page \pageref{abs-HahnBanach}). \label{abs-ineq-iff} For real
+  vector spaces the following inequations are equivalent:
+  \begin{center}
+  \begin{tabular}{lll}
+  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
+  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
+  \end{tabular}
+  \end{center}
+*}
+
+lemma abs_ineq_iff:
+  assumes "subspace H E" and "vectorspace E" and "seminorm E p"
+    and "linearform H h"
+  shows "(\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" (is "?L = ?R")
+proof
+  interpret subspace H E by fact
+  interpret vectorspace E by fact
+  interpret seminorm E p by fact
+  interpret linearform H h by fact
+  have H: "vectorspace H" using `vectorspace E` ..
+  {
+    assume l: ?L
+    show ?R
+    proof
+      fix x assume x: "x \<in> H"
+      have "h x \<le> \<bar>h x\<bar>" by arith
+      also from l x have "\<dots> \<le> p x" ..
+      finally show "h x \<le> p x" .
+    qed
+  next
+    assume r: ?R
+    show ?L
+    proof
+      fix x assume x: "x \<in> H"
+      show "\<And>a b :: real. - a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> \<bar>b\<bar> \<le> a"
+        by arith
+      from `linearform H h` and H x
+      have "- h x = h (- x)" by (rule linearform.neg [symmetric])
+      also
+      from H x have "- x \<in> H" by (rule vectorspace.neg_closed)
+      with r have "h (- x) \<le> p (- x)" ..
+      also have "\<dots> = p x"
+	using `seminorm E p` `vectorspace E`
+      proof (rule seminorm.minus)
+        from x show "x \<in> E" ..
+      qed
+      finally have "- h x \<le> p x" .
+      then show "- p x \<le> h x" by simp
+      from r x show "h x \<le> p x" ..
+    qed
+  }
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/Linearform.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,60 @@
+(*  Title:      HOL/Real/HahnBanach/Linearform.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Linearforms *}
+
+theory Linearform
+imports VectorSpace
+begin
+
+text {*
+  A \emph{linear form} is a function on a vector space into the reals
+  that is additive and multiplicative.
+*}
+
+locale linearform =
+  fixes V :: "'a\<Colon>{minus, plus, zero, uminus} set" and f
+  assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
+    and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x"
+
+declare linearform.intro [intro?]
+
+lemma (in linearform) neg [iff]:
+  assumes "vectorspace V"
+  shows "x \<in> V \<Longrightarrow> f (- x) = - f x"
+proof -
+  interpret vectorspace V by fact
+  assume x: "x \<in> V"
+  then have "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1)
+  also from x have "\<dots> = (- 1) * (f x)" by (rule mult)
+  also from x have "\<dots> = - (f x)" by simp
+  finally show ?thesis .
+qed
+
+lemma (in linearform) diff [iff]:
+  assumes "vectorspace V"
+  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y"
+proof -
+  interpret vectorspace V by fact
+  assume x: "x \<in> V" and y: "y \<in> V"
+  then have "x - y = x + - y" by (rule diff_eq1)
+  also have "f \<dots> = f x + f (- y)" by (rule add) (simp_all add: x y)
+  also have "f (- y) = - f y" using `vectorspace V` y by (rule neg)
+  finally show ?thesis by simp
+qed
+
+text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
+
+lemma (in linearform) zero [iff]:
+  assumes "vectorspace V"
+  shows "f 0 = 0"
+proof -
+  interpret vectorspace V by fact
+  have "f 0 = f (0 - 0)" by simp
+  also have "\<dots> = f 0 - f 0" using `vectorspace V` by (rule diff) simp_all
+  also have "\<dots> = 0" by simp
+  finally show ?thesis .
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/NormedSpace.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,117 @@
+(*  Title:      HOL/Real/HahnBanach/NormedSpace.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Normed vector spaces *}
+
+theory NormedSpace
+imports Subspace
+begin
+
+subsection {* Quasinorms *}
+
+text {*
+  A \emph{seminorm} @{text "\<parallel>\<cdot>\<parallel>"} is a function on a real vector space
+  into the reals that has the following properties: it is positive
+  definite, absolute homogenous and subadditive.
+*}
+
+locale norm_syntax =
+  fixes norm :: "'a \<Rightarrow> real"    ("\<parallel>_\<parallel>")
+
+locale seminorm = var_V + norm_syntax +
+  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
+  assumes ge_zero [iff?]: "x \<in> V \<Longrightarrow> 0 \<le> \<parallel>x\<parallel>"
+    and abs_homogenous [iff?]: "x \<in> V \<Longrightarrow> \<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>"
+    and subadditive [iff?]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
+
+declare seminorm.intro [intro?]
+
+lemma (in seminorm) diff_subadditive:
+  assumes "vectorspace V"
+  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x - y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
+proof -
+  interpret vectorspace V by fact
+  assume x: "x \<in> V" and y: "y \<in> V"
+  then have "x - y = x + - 1 \<cdot> y"
+    by (simp add: diff_eq2 negate_eq2a)
+  also from x y have "\<parallel>\<dots>\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>- 1 \<cdot> y\<parallel>"
+    by (simp add: subadditive)
+  also from y have "\<parallel>- 1 \<cdot> y\<parallel> = \<bar>- 1\<bar> * \<parallel>y\<parallel>"
+    by (rule abs_homogenous)
+  also have "\<dots> = \<parallel>y\<parallel>" by simp
+  finally show ?thesis .
+qed
+
+lemma (in seminorm) minus:
+  assumes "vectorspace V"
+  shows "x \<in> V \<Longrightarrow> \<parallel>- x\<parallel> = \<parallel>x\<parallel>"
+proof -
+  interpret vectorspace V by fact
+  assume x: "x \<in> V"
+  then have "- x = - 1 \<cdot> x" by (simp only: negate_eq1)
+  also from x have "\<parallel>\<dots>\<parallel> = \<bar>- 1\<bar> * \<parallel>x\<parallel>"
+    by (rule abs_homogenous)
+  also have "\<dots> = \<parallel>x\<parallel>" by simp
+  finally show ?thesis .
+qed
+
+
+subsection {* Norms *}
+
+text {*
+  A \emph{norm} @{text "\<parallel>\<cdot>\<parallel>"} is a seminorm that maps only the
+  @{text 0} vector to @{text 0}.
+*}
+
+locale norm = seminorm +
+  assumes zero_iff [iff]: "x \<in> V \<Longrightarrow> (\<parallel>x\<parallel> = 0) = (x = 0)"
+
+
+subsection {* Normed vector spaces *}
+
+text {*
+  A vector space together with a norm is called a \emph{normed
+  space}.
+*}
+
+locale normed_vectorspace = vectorspace + norm
+
+declare normed_vectorspace.intro [intro?]
+
+lemma (in normed_vectorspace) gt_zero [intro?]:
+  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> 0 < \<parallel>x\<parallel>"
+proof -
+  assume x: "x \<in> V" and neq: "x \<noteq> 0"
+  from x have "0 \<le> \<parallel>x\<parallel>" ..
+  also have [symmetric]: "\<dots> \<noteq> 0"
+  proof
+    assume "\<parallel>x\<parallel> = 0"
+    with x have "x = 0" by simp
+    with neq show False by contradiction
+  qed
+  finally show ?thesis .
+qed
+
+text {*
+  Any subspace of a normed vector space is again a normed vectorspace.
+*}
+
+lemma subspace_normed_vs [intro?]:
+  fixes F E norm
+  assumes "subspace F E" "normed_vectorspace E norm"
+  shows "normed_vectorspace F norm"
+proof -
+  interpret subspace F E by fact
+  interpret normed_vectorspace E norm by fact
+  show ?thesis
+  proof
+    show "vectorspace F" by (rule vectorspace) unfold_locales
+  next
+    have "NormedSpace.norm E norm" ..
+    with subset show "NormedSpace.norm F norm"
+      by (simp add: norm_def seminorm_def norm_axioms_def)
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/README.html	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,38 @@
+<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
+
+<!-- $Id$ -->
+
+<HTML>
+
+<HEAD>
+  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
+  <TITLE>HOL/Real/HahnBanach/README</TITLE>
+</HEAD>
+
+<BODY>
+
+<H3>The Hahn-Banach Theorem for Real Vector Spaces (Isabelle/Isar)</H3>
+
+Author: Gertrud Bauer, Technische Universit&auml;t M&uuml;nchen<P>
+
+This directory contains the proof of the Hahn-Banach theorem for real vectorspaces,
+following H. Heuser, Funktionalanalysis, p. 228 -232.
+The Hahn-Banach theorem is one of the fundamental theorems of functioal analysis.
+It is a conclusion of Zorn's lemma.<P>
+
+Two different formaulations of the theorem are presented, one for general real vectorspaces
+and its application to normed vectorspaces. <P>
+
+The theorem says, that every continous linearform, defined on arbitrary subspaces
+(not only one-dimensional subspaces), can be extended to a continous linearform on
+the whole vectorspace.
+
+
+<HR>
+
+<ADDRESS>
+<A NAME="bauerg@in.tum.de" HREF="mailto:bauerg@in.tum.de">bauerg@in.tum.de</A>
+</ADDRESS>
+
+</BODY>
+</HTML>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/ROOT.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,8 @@
+(*  Title:      HOL/Real/HahnBanach/ROOT.ML
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+
+The Hahn-Banach theorem for real vector spaces (Isabelle/Isar).
+*)
+
+time_use_thy "HahnBanach";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/Subspace.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,513 @@
+(*  Title:      HOL/Real/HahnBanach/Subspace.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Subspaces *}
+
+theory Subspace
+imports VectorSpace
+begin
+
+subsection {* Definition *}
+
+text {*
+  A non-empty subset @{text U} of a vector space @{text V} is a
+  \emph{subspace} of @{text V}, iff @{text U} is closed under addition
+  and scalar multiplication.
+*}
+
+locale subspace =
+  fixes U :: "'a\<Colon>{minus, plus, zero, uminus} set" and V
+  assumes non_empty [iff, intro]: "U \<noteq> {}"
+    and subset [iff]: "U \<subseteq> V"
+    and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U"
+    and mult_closed [iff]: "x \<in> U \<Longrightarrow> a \<cdot> x \<in> U"
+
+notation (symbols)
+  subspace  (infix "\<unlhd>" 50)
+
+declare vectorspace.intro [intro?] subspace.intro [intro?]
+
+lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V"
+  by (rule subspace.subset)
+
+lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V"
+  using subset by blast
+
+lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V"
+  by (rule subspace.subsetD)
+
+lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V"
+  by (rule subspace.subsetD)
+
+lemma (in subspace) diff_closed [iff]:
+  assumes "vectorspace V"
+  assumes x: "x \<in> U" and y: "y \<in> U"
+  shows "x - y \<in> U"
+proof -
+  interpret vectorspace V by fact
+  from x y show ?thesis by (simp add: diff_eq1 negate_eq1)
+qed
+
+text {*
+  \medskip Similar as for linear spaces, the existence of the zero
+  element in every subspace follows from the non-emptiness of the
+  carrier set and by vector space laws.
+*}
+
+lemma (in subspace) zero [intro]:
+  assumes "vectorspace V"
+  shows "0 \<in> U"
+proof -
+  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
+  also from `vectorspace V` x x have "\<dots> \<in> U" by (rule diff_closed)
+  finally show ?thesis .
+qed
+
+lemma (in subspace) neg_closed [iff]:
+  assumes "vectorspace V"
+  assumes x: "x \<in> U"
+  shows "- x \<in> U"
+proof -
+  interpret vectorspace V by fact
+  from x show ?thesis by (simp add: negate_eq1)
+qed
+
+text {* \medskip Further derived laws: every subspace is a vector space. *}
+
+lemma (in subspace) vectorspace [iff]:
+  assumes "vectorspace V"
+  shows "vectorspace U"
+proof -
+  interpret vectorspace V by fact
+  show ?thesis
+  proof
+    show "U \<noteq> {}" ..
+    fix x y z assume x: "x \<in> U" and y: "y \<in> U" and z: "z \<in> U"
+    fix a b :: real
+    from x y show "x + y \<in> U" by simp
+    from x show "a \<cdot> x \<in> U" by simp
+    from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac)
+    from x y show "x + y = y + x" by (simp add: add_ac)
+    from x show "x - x = 0" by simp
+    from x show "0 + x = x" by simp
+    from x y show "a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" by (simp add: distrib)
+    from x show "(a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" by (simp add: distrib)
+    from x show "(a * b) \<cdot> x = a \<cdot> b \<cdot> x" by (simp add: mult_assoc)
+    from x show "1 \<cdot> x = x" by simp
+    from x show "- x = - 1 \<cdot> x" by (simp add: negate_eq1)
+    from x y show "x - y = x + - y" by (simp add: diff_eq1)
+  qed
+qed
+
+
+text {* The subspace relation is reflexive. *}
+
+lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
+proof
+  show "V \<noteq> {}" ..
+  show "V \<subseteq> V" ..
+  fix x y assume x: "x \<in> V" and y: "y \<in> V"
+  fix a :: real
+  from x y show "x + y \<in> V" by simp
+  from x show "a \<cdot> x \<in> V" by simp
+qed
+
+text {* The subspace relation is transitive. *}
+
+lemma (in vectorspace) subspace_trans [trans]:
+  "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W"
+proof
+  assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W"
+  from uv show "U \<noteq> {}" by (rule subspace.non_empty)
+  show "U \<subseteq> W"
+  proof -
+    from uv have "U \<subseteq> V" by (rule subspace.subset)
+    also from vw have "V \<subseteq> W" by (rule subspace.subset)
+    finally show ?thesis .
+  qed
+  fix x y assume x: "x \<in> U" and y: "y \<in> U"
+  from uv and x y show "x + y \<in> U" by (rule subspace.add_closed)
+  from uv and x show "\<And>a. a \<cdot> x \<in> U" by (rule subspace.mult_closed)
+qed
+
+
+subsection {* Linear closure *}
+
+text {*
+  The \emph{linear closure} of a vector @{text x} is the set of all
+  scalar multiples of @{text x}.
+*}
+
+definition
+  lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set" where
+  "lin x = {a \<cdot> x | a. True}"
+
+lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
+  unfolding lin_def by blast
+
+lemma linI' [iff]: "a \<cdot> x \<in> lin x"
+  unfolding lin_def by blast
+
+lemma linE [elim]: "x \<in> lin v \<Longrightarrow> (\<And>a::real. x = a \<cdot> v \<Longrightarrow> C) \<Longrightarrow> C"
+  unfolding lin_def by blast
+
+
+text {* Every vector is contained in its linear closure. *}
+
+lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x"
+proof -
+  assume "x \<in> V"
+  then have "x = 1 \<cdot> x" by simp
+  also have "\<dots> \<in> lin x" ..
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x"
+proof
+  assume "x \<in> V"
+  then show "0 = 0 \<cdot> x" by simp
+qed
+
+text {* Any linear closure is a subspace. *}
+
+lemma (in vectorspace) lin_subspace [intro]:
+  "x \<in> V \<Longrightarrow> lin x \<unlhd> V"
+proof
+  assume x: "x \<in> V"
+  then show "lin x \<noteq> {}" by (auto simp add: x_lin_x)
+  show "lin x \<subseteq> V"
+  proof
+    fix x' assume "x' \<in> lin x"
+    then obtain a where "x' = a \<cdot> x" ..
+    with x show "x' \<in> V" by simp
+  qed
+  fix x' x'' assume x': "x' \<in> lin x" and x'': "x'' \<in> lin x"
+  show "x' + x'' \<in> lin x"
+  proof -
+    from x' obtain a' where "x' = a' \<cdot> x" ..
+    moreover from x'' obtain a'' where "x'' = a'' \<cdot> x" ..
+    ultimately have "x' + x'' = (a' + a'') \<cdot> x"
+      using x by (simp add: distrib)
+    also have "\<dots> \<in> lin x" ..
+    finally show ?thesis .
+  qed
+  fix a :: real
+  show "a \<cdot> x' \<in> lin x"
+  proof -
+    from x' obtain a' where "x' = a' \<cdot> x" ..
+    with x have "a \<cdot> x' = (a * a') \<cdot> x" by (simp add: mult_assoc)
+    also have "\<dots> \<in> lin x" ..
+    finally show ?thesis .
+  qed
+qed
+
+
+text {* Any linear closure is a vector space. *}
+
+lemma (in vectorspace) lin_vectorspace [intro]:
+  assumes "x \<in> V"
+  shows "vectorspace (lin x)"
+proof -
+  from `x \<in> V` have "subspace (lin x) V"
+    by (rule lin_subspace)
+  from this and vectorspace_axioms show ?thesis
+    by (rule subspace.vectorspace)
+qed
+
+
+subsection {* Sum of two vectorspaces *}
+
+text {*
+  The \emph{sum} of two vectorspaces @{text U} and @{text V} is the
+  set of all sums of elements from @{text U} and @{text V}.
+*}
+
+instantiation "fun" :: (type, type) plus
+begin
+
+definition
+  sum_def: "plus_fun U V = {u + v | u v. u \<in> U \<and> v \<in> V}"  (* FIXME not fully general!? *)
+
+instance ..
+
+end
+
+lemma sumE [elim]:
+    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
+  unfolding sum_def by blast
+
+lemma sumI [intro]:
+    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
+  unfolding sum_def by blast
+
+lemma sumI' [intro]:
+    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
+  unfolding sum_def by blast
+
+text {* @{text U} is a subspace of @{text "U + V"}. *}
+
+lemma subspace_sum1 [iff]:
+  assumes "vectorspace U" "vectorspace V"
+  shows "U \<unlhd> U + V"
+proof -
+  interpret vectorspace U by fact
+  interpret vectorspace V by fact
+  show ?thesis
+  proof
+    show "U \<noteq> {}" ..
+    show "U \<subseteq> U + V"
+    proof
+      fix x assume x: "x \<in> U"
+      moreover have "0 \<in> V" ..
+      ultimately have "x + 0 \<in> U + V" ..
+      with x show "x \<in> U + V" by simp
+    qed
+    fix x y assume x: "x \<in> U" and "y \<in> U"
+    then show "x + y \<in> U" by simp
+    from x show "\<And>a. a \<cdot> x \<in> U" by simp
+  qed
+qed
+
+text {* The sum of two subspaces is again a subspace. *}
+
+lemma sum_subspace [intro?]:
+  assumes "subspace U E" "vectorspace E" "subspace V E"
+  shows "U + V \<unlhd> E"
+proof -
+  interpret subspace U E by fact
+  interpret vectorspace E by fact
+  interpret subspace V E by fact
+  show ?thesis
+  proof
+    have "0 \<in> U + V"
+    proof
+      show "0 \<in> U" using `vectorspace E` ..
+      show "0 \<in> V" using `vectorspace E` ..
+      show "(0::'a) = 0 + 0" by simp
+    qed
+    then show "U + V \<noteq> {}" by blast
+    show "U + V \<subseteq> E"
+    proof
+      fix x assume "x \<in> U + V"
+      then obtain u v where "x = u + v" and
+	"u \<in> U" and "v \<in> V" ..
+      then show "x \<in> E" by simp
+    qed
+    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
+    show "x + y \<in> U + V"
+    proof -
+      from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
+      moreover
+      from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" ..
+      ultimately
+      have "ux + uy \<in> U"
+	and "vx + vy \<in> V"
+	and "x + y = (ux + uy) + (vx + vy)"
+	using x y by (simp_all add: add_ac)
+      then show ?thesis ..
+    qed
+    fix a show "a \<cdot> x \<in> U + V"
+    proof -
+      from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
+      then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
+	and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
+      then show ?thesis ..
+    qed
+  qed
+qed
+
+text{* The sum of two subspaces is a vectorspace. *}
+
+lemma sum_vs [intro?]:
+    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
+  by (rule subspace.vectorspace) (rule sum_subspace)
+
+
+subsection {* Direct sums *}
+
+text {*
+  The sum of @{text U} and @{text V} is called \emph{direct}, iff the
+  zero element is the only common element of @{text U} and @{text
+  V}. For every element @{text x} of the direct sum of @{text U} and
+  @{text V} the decomposition in @{text "x = u + v"} with
+  @{text "u \<in> U"} and @{text "v \<in> V"} is unique.
+*}
+
+lemma decomp:
+  assumes "vectorspace E" "subspace U E" "subspace V E"
+  assumes direct: "U \<inter> V = {0}"
+    and u1: "u1 \<in> U" and u2: "u2 \<in> U"
+    and v1: "v1 \<in> V" and v2: "v2 \<in> V"
+    and sum: "u1 + v1 = u2 + v2"
+  shows "u1 = u2 \<and> v1 = v2"
+proof -
+  interpret vectorspace E by fact
+  interpret subspace U E by fact
+  interpret subspace V E by fact
+  show ?thesis
+  proof
+    have U: "vectorspace U"  (* FIXME: use interpret *)
+      using `subspace U E` `vectorspace E` by (rule subspace.vectorspace)
+    have V: "vectorspace V"
+      using `subspace V E` `vectorspace E` by (rule subspace.vectorspace)
+    from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1"
+      by (simp add: add_diff_swap)
+    from u1 u2 have u: "u1 - u2 \<in> U"
+      by (rule vectorspace.diff_closed [OF U])
+    with eq have v': "v2 - v1 \<in> U" by (simp only:)
+    from v2 v1 have v: "v2 - v1 \<in> V"
+      by (rule vectorspace.diff_closed [OF V])
+    with eq have u': " u1 - u2 \<in> V" by (simp only:)
+    
+    show "u1 = u2"
+    proof (rule add_minus_eq)
+      from u1 show "u1 \<in> E" ..
+      from u2 show "u2 \<in> E" ..
+      from u u' and direct show "u1 - u2 = 0" by blast
+    qed
+    show "v1 = v2"
+    proof (rule add_minus_eq [symmetric])
+      from v1 show "v1 \<in> E" ..
+      from v2 show "v2 \<in> E" ..
+      from v v' and direct show "v2 - v1 = 0" by blast
+    qed
+  qed
+qed
+
+text {*
+  An application of the previous lemma will be used in the proof of
+  the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
+  element @{text "y + a \<cdot> x\<^sub>0"} of the direct sum of a
+  vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"}
+  the components @{text "y \<in> H"} and @{text a} are uniquely
+  determined.
+*}
+
+lemma decomp_H':
+  assumes "vectorspace E" "subspace H E"
+  assumes y1: "y1 \<in> H" and y2: "y2 \<in> H"
+    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
+    and eq: "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'"
+  shows "y1 = y2 \<and> a1 = a2"
+proof -
+  interpret vectorspace E by fact
+  interpret subspace H E by fact
+  show ?thesis
+  proof
+    have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'"
+    proof (rule decomp)
+      show "a1 \<cdot> x' \<in> lin x'" ..
+      show "a2 \<cdot> x' \<in> lin x'" ..
+      show "H \<inter> lin x' = {0}"
+      proof
+	show "H \<inter> lin x' \<subseteq> {0}"
+	proof
+          fix x assume x: "x \<in> H \<inter> lin x'"
+          then obtain a where xx': "x = a \<cdot> x'"
+            by blast
+          have "x = 0"
+          proof cases
+            assume "a = 0"
+            with xx' and x' show ?thesis by simp
+          next
+            assume a: "a \<noteq> 0"
+            from x have "x \<in> H" ..
+            with xx' have "inverse a \<cdot> a \<cdot> x' \<in> H" by simp
+            with a and x' have "x' \<in> H" by (simp add: mult_assoc2)
+            with `x' \<notin> H` show ?thesis by contradiction
+          qed
+          then show "x \<in> {0}" ..
+	qed
+	show "{0} \<subseteq> H \<inter> lin x'"
+	proof -
+          have "0 \<in> H" using `vectorspace E` ..
+          moreover have "0 \<in> lin x'" using `x' \<in> E` ..
+          ultimately show ?thesis by blast
+	qed
+      qed
+      show "lin x' \<unlhd> E" using `x' \<in> E` ..
+    qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq)
+    then show "y1 = y2" ..
+    from c have "a1 \<cdot> x' = a2 \<cdot> x'" ..
+    with x' show "a1 = a2" by (simp add: mult_right_cancel)
+  qed
+qed
+
+text {*
+  Since for any element @{text "y + a \<cdot> x'"} of the direct sum of a
+  vectorspace @{text H} and the linear closure of @{text x'} the
+  components @{text "y \<in> H"} and @{text a} are unique, it follows from
+  @{text "y \<in> H"} that @{text "a = 0"}.
+*}
+
+lemma decomp_H'_H:
+  assumes "vectorspace E" "subspace H E"
+  assumes t: "t \<in> H"
+    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
+  shows "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
+proof -
+  interpret vectorspace E by fact
+  interpret subspace H E by fact
+  show ?thesis
+  proof (rule, simp_all only: split_paired_all split_conv)
+    from t x' show "t = t + 0 \<cdot> x' \<and> t \<in> H" by simp
+    fix y and a assume ya: "t = y + a \<cdot> x' \<and> y \<in> H"
+    have "y = t \<and> a = 0"
+    proof (rule decomp_H')
+      from ya x' show "y + a \<cdot> x' = t + 0 \<cdot> x'" by simp
+      from ya show "y \<in> H" ..
+    qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+)
+    with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp
+  qed
+qed
+
+text {*
+  The components @{text "y \<in> H"} and @{text a} in @{text "y + a \<cdot> x'"}
+  are unique, so the function @{text h'} defined by
+  @{text "h' (y + a \<cdot> x') = h y + a \<cdot> \<xi>"} is definite.
+*}
+
+lemma h'_definite:
+  fixes H
+  assumes h'_def:
+    "h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
+                in (h y) + a * xi)"
+    and x: "x = y + a \<cdot> x'"
+  assumes "vectorspace E" "subspace H E"
+  assumes y: "y \<in> H"
+    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
+  shows "h' x = h y + a * xi"
+proof -
+  interpret vectorspace E by fact
+  interpret subspace H E by fact
+  from x y x' have "x \<in> H + lin x'" by auto
+  have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p")
+  proof (rule ex_ex1I)
+    from x y show "\<exists>p. ?P p" by blast
+    fix p q assume p: "?P p" and q: "?P q"
+    show "p = q"
+    proof -
+      from p have xp: "x = fst p + snd p \<cdot> x' \<and> fst p \<in> H"
+        by (cases p) simp
+      from q have xq: "x = fst q + snd q \<cdot> x' \<and> fst q \<in> H"
+        by (cases q) simp
+      have "fst p = fst q \<and> snd p = snd q"
+      proof (rule decomp_H')
+        from xp show "fst p \<in> H" ..
+        from xq show "fst q \<in> H" ..
+        from xp and xq show "fst p + snd p \<cdot> x' = fst q + snd q \<cdot> x'"
+          by simp
+      qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+)
+      then show ?thesis by (cases p, cases q) simp
+    qed
+  qed
+  then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)"
+    by (rule some1_equality) (simp add: x y)
+  with h'_def show "h' x = h y + a * xi" by (simp add: Let_def)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/VectorSpace.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,419 @@
+(*  Title:      HOL/Real/HahnBanach/VectorSpace.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Vector spaces *}
+
+theory VectorSpace
+imports Real Bounds Zorn
+begin
+
+subsection {* Signature *}
+
+text {*
+  For the definition of real vector spaces a type @{typ 'a} of the
+  sort @{text "{plus, minus, zero}"} is considered, on which a real
+  scalar multiplication @{text \<cdot>} is declared.
+*}
+
+consts
+  prod  :: "real \<Rightarrow> 'a::{plus, minus, zero} \<Rightarrow> 'a"     (infixr "'(*')" 70)
+
+notation (xsymbols)
+  prod  (infixr "\<cdot>" 70)
+notation (HTML output)
+  prod  (infixr "\<cdot>" 70)
+
+
+subsection {* Vector space laws *}
+
+text {*
+  A \emph{vector space} is a non-empty set @{text V} of elements from
+  @{typ 'a} with the following vector space laws: The set @{text V} is
+  closed under addition and scalar multiplication, addition is
+  associative and commutative; @{text "- x"} is the inverse of @{text
+  x} w.~r.~t.~addition and @{text 0} is the neutral element of
+  addition.  Addition and multiplication are distributive; scalar
+  multiplication is associative and the real number @{text "1"} is
+  the neutral element of scalar multiplication.
+*}
+
+locale var_V = fixes V
+
+locale vectorspace = var_V +
+  assumes non_empty [iff, intro?]: "V \<noteq> {}"
+    and add_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y \<in> V"
+    and mult_closed [iff]: "x \<in> V \<Longrightarrow> a \<cdot> x \<in> V"
+    and add_assoc: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y) + z = x + (y + z)"
+    and add_commute: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = y + x"
+    and diff_self [simp]: "x \<in> V \<Longrightarrow> x - x = 0"
+    and add_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 + x = x"
+    and add_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y"
+    and add_mult_distrib2: "x \<in> V \<Longrightarrow> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x"
+    and mult_assoc: "x \<in> V \<Longrightarrow> (a * b) \<cdot> x = a \<cdot> (b \<cdot> x)"
+    and mult_1 [simp]: "x \<in> V \<Longrightarrow> 1 \<cdot> x = x"
+    and negate_eq1: "x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x"
+    and diff_eq1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = x + - y"
+
+lemma (in vectorspace) negate_eq2: "x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x"
+  by (rule negate_eq1 [symmetric])
+
+lemma (in vectorspace) negate_eq2a: "x \<in> V \<Longrightarrow> -1 \<cdot> x = - x"
+  by (simp add: negate_eq1)
+
+lemma (in vectorspace) diff_eq2: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + - y = x - y"
+  by (rule diff_eq1 [symmetric])
+
+lemma (in vectorspace) diff_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y \<in> V"
+  by (simp add: diff_eq1 negate_eq1)
+
+lemma (in vectorspace) neg_closed [iff]: "x \<in> V \<Longrightarrow> - x \<in> V"
+  by (simp add: negate_eq1)
+
+lemma (in vectorspace) add_left_commute:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> x + (y + z) = y + (x + z)"
+proof -
+  assume xyz: "x \<in> V"  "y \<in> V"  "z \<in> V"
+  then have "x + (y + z) = (x + y) + z"
+    by (simp only: add_assoc)
+  also from xyz have "\<dots> = (y + x) + z" by (simp only: add_commute)
+  also from xyz have "\<dots> = y + (x + z)" by (simp only: add_assoc)
+  finally show ?thesis .
+qed
+
+theorems (in vectorspace) add_ac =
+  add_assoc add_commute add_left_commute
+
+
+text {* The existence of the zero element of a vector space
+  follows from the non-emptiness of carrier set. *}
+
+lemma (in vectorspace) zero [iff]: "0 \<in> V"
+proof -
+  from non_empty obtain x where x: "x \<in> V" by blast
+  then have "0 = x - x" by (rule diff_self [symmetric])
+  also from x x have "\<dots> \<in> V" by (rule diff_closed)
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) add_zero_right [simp]:
+  "x \<in> V \<Longrightarrow>  x + 0 = x"
+proof -
+  assume x: "x \<in> V"
+  from this and zero have "x + 0 = 0 + x" by (rule add_commute)
+  also from x have "\<dots> = x" by (rule add_zero_left)
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) mult_assoc2:
+    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = (a * b) \<cdot> x"
+  by (simp only: mult_assoc)
+
+lemma (in vectorspace) diff_mult_distrib1:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x - y) = a \<cdot> x - a \<cdot> y"
+  by (simp add: diff_eq1 negate_eq1 add_mult_distrib1 mult_assoc2)
+
+lemma (in vectorspace) diff_mult_distrib2:
+  "x \<in> V \<Longrightarrow> (a - b) \<cdot> x = a \<cdot> x - (b \<cdot> x)"
+proof -
+  assume x: "x \<in> V"
+  have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
+    by (simp add: real_diff_def)
+  also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
+    by (rule add_mult_distrib2)
+  also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
+    by (simp add: negate_eq1 mult_assoc2)
+  also from x have "\<dots> = a \<cdot> x - (b \<cdot> x)"
+    by (simp add: diff_eq1)
+  finally show ?thesis .
+qed
+
+lemmas (in vectorspace) distrib =
+  add_mult_distrib1 add_mult_distrib2
+  diff_mult_distrib1 diff_mult_distrib2
+
+
+text {* \medskip Further derived laws: *}
+
+lemma (in vectorspace) mult_zero_left [simp]:
+  "x \<in> V \<Longrightarrow> 0 \<cdot> x = 0"
+proof -
+  assume x: "x \<in> V"
+  have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp
+  also have "\<dots> = (1 + - 1) \<cdot> x" by simp
+  also from x have "\<dots> =  1 \<cdot> x + (- 1) \<cdot> x"
+    by (rule add_mult_distrib2)
+  also from x have "\<dots> = x + (- 1) \<cdot> x" by simp
+  also from x have "\<dots> = x + - x" by (simp add: negate_eq2a)
+  also from x have "\<dots> = x - x" by (simp add: diff_eq2)
+  also from x have "\<dots> = 0" by simp
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) mult_zero_right [simp]:
+  "a \<cdot> 0 = (0::'a)"
+proof -
+  have "a \<cdot> 0 = a \<cdot> (0 - (0::'a))" by simp
+  also have "\<dots> =  a \<cdot> 0 - a \<cdot> 0"
+    by (rule diff_mult_distrib1) simp_all
+  also have "\<dots> = 0" by simp
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) minus_mult_cancel [simp]:
+    "x \<in> V \<Longrightarrow> (- a) \<cdot> - x = a \<cdot> x"
+  by (simp add: negate_eq1 mult_assoc2)
+
+lemma (in vectorspace) add_minus_left_eq_diff:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + y = y - x"
+proof -
+  assume xy: "x \<in> V"  "y \<in> V"
+  then have "- x + y = y + - x" by (simp add: add_commute)
+  also from xy have "\<dots> = y - x" by (simp add: diff_eq1)
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) add_minus [simp]:
+    "x \<in> V \<Longrightarrow> x + - x = 0"
+  by (simp add: diff_eq2)
+
+lemma (in vectorspace) add_minus_left [simp]:
+    "x \<in> V \<Longrightarrow> - x + x = 0"
+  by (simp add: diff_eq2 add_commute)
+
+lemma (in vectorspace) minus_minus [simp]:
+    "x \<in> V \<Longrightarrow> - (- x) = x"
+  by (simp add: negate_eq1 mult_assoc2)
+
+lemma (in vectorspace) minus_zero [simp]:
+    "- (0::'a) = 0"
+  by (simp add: negate_eq1)
+
+lemma (in vectorspace) minus_zero_iff [simp]:
+  "x \<in> V \<Longrightarrow> (- x = 0) = (x = 0)"
+proof
+  assume x: "x \<in> V"
+  {
+    from x have "x = - (- x)" by (simp add: minus_minus)
+    also assume "- x = 0"
+    also have "- \<dots> = 0" by (rule minus_zero)
+    finally show "x = 0" .
+  next
+    assume "x = 0"
+    then show "- x = 0" by simp
+  }
+qed
+
+lemma (in vectorspace) add_minus_cancel [simp]:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + (- x + y) = y"
+  by (simp add: add_assoc [symmetric] del: add_commute)
+
+lemma (in vectorspace) minus_add_cancel [simp]:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + (x + y) = y"
+  by (simp add: add_assoc [symmetric] del: add_commute)
+
+lemma (in vectorspace) minus_add_distrib [simp]:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - (x + y) = - x + - y"
+  by (simp add: negate_eq1 add_mult_distrib1)
+
+lemma (in vectorspace) diff_zero [simp]:
+    "x \<in> V \<Longrightarrow> x - 0 = x"
+  by (simp add: diff_eq1)
+
+lemma (in vectorspace) diff_zero_right [simp]:
+    "x \<in> V \<Longrightarrow> 0 - x = - x"
+  by (simp add: diff_eq1)
+
+lemma (in vectorspace) add_left_cancel:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y = x + z) = (y = z)"
+proof
+  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
+  {
+    from y have "y = 0 + y" by simp
+    also from x y have "\<dots> = (- x + x) + y" by simp
+    also from x y have "\<dots> = - x + (x + y)"
+      by (simp add: add_assoc neg_closed)
+    also assume "x + y = x + z"
+    also from x z have "- x + (x + z) = - x + x + z"
+      by (simp add: add_assoc [symmetric] neg_closed)
+    also from x z have "\<dots> = z" by simp
+    finally show "y = z" .
+  next
+    assume "y = z"
+    then show "x + y = x + z" by (simp only:)
+  }
+qed
+
+lemma (in vectorspace) add_right_cancel:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (y + x = z + x) = (y = z)"
+  by (simp only: add_commute add_left_cancel)
+
+lemma (in vectorspace) add_assoc_cong:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x' \<in> V \<Longrightarrow> y' \<in> V \<Longrightarrow> z \<in> V
+    \<Longrightarrow> x + y = x' + y' \<Longrightarrow> x + (y + z) = x' + (y' + z)"
+  by (simp only: add_assoc [symmetric])
+
+lemma (in vectorspace) mult_left_commute:
+    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = b \<cdot> a \<cdot> x"
+  by (simp add: real_mult_commute mult_assoc2)
+
+lemma (in vectorspace) mult_zero_uniq:
+  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> a \<cdot> x = 0 \<Longrightarrow> a = 0"
+proof (rule classical)
+  assume a: "a \<noteq> 0"
+  assume x: "x \<in> V"  "x \<noteq> 0" and ax: "a \<cdot> x = 0"
+  from x a have "x = (inverse a * a) \<cdot> x" by simp
+  also from `x \<in> V` have "\<dots> = inverse a \<cdot> (a \<cdot> x)" by (rule mult_assoc)
+  also from ax have "\<dots> = inverse a \<cdot> 0" by simp
+  also have "\<dots> = 0" by simp
+  finally have "x = 0" .
+  with `x \<noteq> 0` show "a = 0" by contradiction
+qed
+
+lemma (in vectorspace) mult_left_cancel:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (a \<cdot> x = a \<cdot> y) = (x = y)"
+proof
+  assume x: "x \<in> V" and y: "y \<in> V" and a: "a \<noteq> 0"
+  from x have "x = 1 \<cdot> x" by simp
+  also from a have "\<dots> = (inverse a * a) \<cdot> x" by simp
+  also from x have "\<dots> = inverse a \<cdot> (a \<cdot> x)"
+    by (simp only: mult_assoc)
+  also assume "a \<cdot> x = a \<cdot> y"
+  also from a y have "inverse a \<cdot> \<dots> = y"
+    by (simp add: mult_assoc2)
+  finally show "x = y" .
+next
+  assume "x = y"
+  then show "a \<cdot> x = a \<cdot> y" by (simp only:)
+qed
+
+lemma (in vectorspace) mult_right_cancel:
+  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (a \<cdot> x = b \<cdot> x) = (a = b)"
+proof
+  assume x: "x \<in> V" and neq: "x \<noteq> 0"
+  {
+    from x have "(a - b) \<cdot> x = a \<cdot> x - b \<cdot> x"
+      by (simp add: diff_mult_distrib2)
+    also assume "a \<cdot> x = b \<cdot> x"
+    with x have "a \<cdot> x - b \<cdot> x = 0" by simp
+    finally have "(a - b) \<cdot> x = 0" .
+    with x neq have "a - b = 0" by (rule mult_zero_uniq)
+    then show "a = b" by simp
+  next
+    assume "a = b"
+    then show "a \<cdot> x = b \<cdot> x" by (simp only:)
+  }
+qed
+
+lemma (in vectorspace) eq_diff_eq:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x = z - y) = (x + y = z)"
+proof
+  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
+  {
+    assume "x = z - y"
+    then have "x + y = z - y + y" by simp
+    also from y z have "\<dots> = z + - y + y"
+      by (simp add: diff_eq1)
+    also have "\<dots> = z + (- y + y)"
+      by (rule add_assoc) (simp_all add: y z)
+    also from y z have "\<dots> = z + 0"
+      by (simp only: add_minus_left)
+    also from z have "\<dots> = z"
+      by (simp only: add_zero_right)
+    finally show "x + y = z" .
+  next
+    assume "x + y = z"
+    then have "z - y = (x + y) - y" by simp
+    also from x y have "\<dots> = x + y + - y"
+      by (simp add: diff_eq1)
+    also have "\<dots> = x + (y + - y)"
+      by (rule add_assoc) (simp_all add: x y)
+    also from x y have "\<dots> = x" by simp
+    finally show "x = z - y" ..
+  }
+qed
+
+lemma (in vectorspace) add_minus_eq_minus:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = 0 \<Longrightarrow> x = - y"
+proof -
+  assume x: "x \<in> V" and y: "y \<in> V"
+  from x y have "x = (- y + y) + x" by simp
+  also from x y have "\<dots> = - y + (x + y)" by (simp add: add_ac)
+  also assume "x + y = 0"
+  also from y have "- y + 0 = - y" by simp
+  finally show "x = - y" .
+qed
+
+lemma (in vectorspace) add_minus_eq:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = 0 \<Longrightarrow> x = y"
+proof -
+  assume x: "x \<in> V" and y: "y \<in> V"
+  assume "x - y = 0"
+  with x y have eq: "x + - y = 0" by (simp add: diff_eq1)
+  with _ _ have "x = - (- y)"
+    by (rule add_minus_eq_minus) (simp_all add: x y)
+  with x y show "x = y" by simp
+qed
+
+lemma (in vectorspace) add_diff_swap:
+  "a \<in> V \<Longrightarrow> b \<in> V \<Longrightarrow> c \<in> V \<Longrightarrow> d \<in> V \<Longrightarrow> a + b = c + d
+    \<Longrightarrow> a - c = d - b"
+proof -
+  assume vs: "a \<in> V"  "b \<in> V"  "c \<in> V"  "d \<in> V"
+    and eq: "a + b = c + d"
+  then have "- c + (a + b) = - c + (c + d)"
+    by (simp add: add_left_cancel)
+  also have "\<dots> = d" using `c \<in> V` `d \<in> V` by (rule minus_add_cancel)
+  finally have eq: "- c + (a + b) = d" .
+  from vs have "a - c = (- c + (a + b)) + - b"
+    by (simp add: add_ac diff_eq1)
+  also from vs eq have "\<dots>  = d + - b"
+    by (simp add: add_right_cancel)
+  also from vs have "\<dots> = d - b" by (simp add: diff_eq2)
+  finally show "a - c = d - b" .
+qed
+
+lemma (in vectorspace) vs_add_cancel_21:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> u \<in> V
+    \<Longrightarrow> (x + (y + z) = y + u) = (x + z = u)"
+proof
+  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"  "u \<in> V"
+  {
+    from vs have "x + z = - y + y + (x + z)" by simp
+    also have "\<dots> = - y + (y + (x + z))"
+      by (rule add_assoc) (simp_all add: vs)
+    also from vs have "y + (x + z) = x + (y + z)"
+      by (simp add: add_ac)
+    also assume "x + (y + z) = y + u"
+    also from vs have "- y + (y + u) = u" by simp
+    finally show "x + z = u" .
+  next
+    assume "x + z = u"
+    with vs show "x + (y + z) = y + u"
+      by (simp only: add_left_commute [of x])
+  }
+qed
+
+lemma (in vectorspace) add_cancel_end:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + (y + z) = y) = (x = - z)"
+proof
+  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"
+  {
+    assume "x + (y + z) = y"
+    with vs have "(x + z) + y = 0 + y"
+      by (simp add: add_ac)
+    with vs have "x + z = 0"
+      by (simp only: add_right_cancel add_closed zero)
+    with vs show "x = - z" by (simp add: add_minus_eq_minus)
+  next
+    assume eq: "x = - z"
+    then have "x + (y + z) = - z + (y + z)" by simp
+    also have "\<dots> = y + (- z + z)"
+      by (rule add_left_commute) (simp_all add: vs)
+    also from vs have "\<dots> = y"  by simp
+    finally show "x + (y + z) = y" .
+  }
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/ZornLemma.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,57 @@
+(*  Title:      HOL/Real/HahnBanach/ZornLemma.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Zorn's Lemma *}
+
+theory ZornLemma
+imports Zorn
+begin
+
+text {*
+  Zorn's Lemmas states: if every linear ordered subset of an ordered
+  set @{text S} has an upper bound in @{text S}, then there exists a
+  maximal element in @{text S}.  In our application, @{text S} is a
+  set of sets ordered by set inclusion. Since the union of a chain of
+  sets is an upper bound for all elements of the chain, the conditions
+  of Zorn's lemma can be modified: if @{text S} is non-empty, it
+  suffices to show that for every non-empty chain @{text c} in @{text
+  S} the union of @{text c} also lies in @{text S}.
+*}
+
+theorem Zorn's_Lemma:
+  assumes r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
+    and aS: "a \<in> S"
+  shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
+proof (rule Zorn_Lemma2)
+  show "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
+  proof
+    fix c assume "c \<in> chain S"
+    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
+    proof cases
+
+      txt {* If @{text c} is an empty chain, then every element in
+	@{text S} is an upper bound of @{text c}. *}
+
+      assume "c = {}"
+      with aS show ?thesis by fast
+
+      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
+	bound of @{text c}, lying in @{text S}. *}
+
+    next
+      assume "c \<noteq> {}"
+      show ?thesis
+      proof
+        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
+        show "\<Union>c \<in> S"
+        proof (rule r)
+          from `c \<noteq> {}` show "\<exists>x. x \<in> c" by fast
+	  show "c \<in> chain S" by fact
+        qed
+      qed
+    qed
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/document/root.bib	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,27 @@
+
+@Book{Heuser:1986,
+  author = 	 {H. Heuser},
+  title = 	 {Funktionalanalysis: Theorie und Anwendung},
+  publisher = 	 {Teubner},
+  year = 	 1986
+}
+
+@InCollection{Narici:1996,
+  author = 	 {L. Narici and E. Beckenstein},
+  title = 	 {The {Hahn-Banach Theorem}: The Life and Times},
+  booktitle = 	 {Topology Atlas},
+  publisher =	 {York University, Toronto, Ontario, Canada},
+  year =	 1996,
+  note =	 {\url{http://at.yorku.ca/topology/preprint.htm} and
+                  \url{http://at.yorku.ca/p/a/a/a/16.htm}}
+}
+
+@Article{Nowak:1993,
+  author =       {B. Nowak and A. Trybulec},
+  title =	 {{Hahn-Banach} Theorem},
+  journal =      {Journal of Formalized Mathematics},
+  year =         {1993},
+  volume =       {5},
+  institution =  {University of Bialystok},
+  note =         {\url{http://mizar.uwb.edu.pl/JFM/Vol5/hahnban.html}}
+}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/document/root.tex	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,83 @@
+\documentclass[10pt,a4paper,twoside]{article}
+\usepackage{graphicx}
+\usepackage{latexsym,theorem}
+\usepackage{isabelle,isabellesym}
+\usepackage{pdfsetup} %last one!
+
+\isabellestyle{it}
+\urlstyle{rm}
+
+\newcommand{\isasymsup}{\isamath{\sup\,}}
+\newcommand{\skp}{\smallskip}
+
+
+\begin{document}
+
+\pagestyle{headings}
+\pagenumbering{arabic}
+
+\title{The Hahn-Banach Theorem \\ for Real Vector Spaces}
+\author{Gertrud Bauer \\ \url{http://www.in.tum.de/~bauerg/}}
+\maketitle
+
+\begin{abstract}
+  The Hahn-Banach Theorem is one of the most fundamental results in functional
+  analysis. We present a fully formal proof of two versions of the theorem,
+  one for general linear spaces and another for normed spaces.  This
+  development is based on simply-typed classical set-theory, as provided by
+  Isabelle/HOL.
+\end{abstract}
+
+
+\tableofcontents
+\parindent 0pt \parskip 0.5ex
+
+\clearpage
+\section{Preface}
+
+This is a fully formal proof of the Hahn-Banach Theorem. It closely follows
+the informal presentation given in Heuser's textbook \cite[{\S} 36]{Heuser:1986}.
+Another formal proof of the same theorem has been done in Mizar
+\cite{Nowak:1993}.  A general overview of the relevance and history of the
+Hahn-Banach Theorem is given by Narici and Beckenstein \cite{Narici:1996}.
+
+\medskip The document is structured as follows.  The first part contains
+definitions of basic notions of linear algebra: vector spaces, subspaces,
+normed spaces, continuous linear-forms, norm of functions and an order on
+functions by domain extension.  The second part contains some lemmas about the
+supremum (w.r.t.\ the function order) and extension of non-maximal functions.
+With these preliminaries, the main proof of the theorem (in its two versions)
+is conducted in the third part.  The dependencies of individual theories are
+as follows.
+
+\begin{center}
+  \includegraphics[scale=0.5]{session_graph}  
+\end{center}
+
+\clearpage
+\part {Basic Notions}
+
+\input{Bounds}
+\input{VectorSpace}
+\input{Subspace}
+\input{NormedSpace}
+\input{Linearform}
+\input{FunctionOrder}
+\input{FunctionNorm}
+\input{ZornLemma}
+
+\clearpage
+\part {Lemmas for the Proof}
+
+\input{HahnBanachSupLemmas}
+\input{HahnBanachExtLemmas}
+\input{HahnBanachLemmas}
+
+\clearpage
+\part {The Main Proof}
+
+\input{HahnBanach}
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
--- a/src/HOL/Hyperreal/SEQ.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1136 +0,0 @@
-(*  Title       : SEQ.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Convergence of sequences and series
-    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
-    Additional contributions by Jeremy Avigad and Brian Huffman
-*)
-
-header {* Sequences and Convergence *}
-
-theory SEQ
-imports "../Real/RealVector" "../RComplete"
-begin
-
-definition
-  Zseq :: "[nat \<Rightarrow> 'a::real_normed_vector] \<Rightarrow> bool" where
-    --{*Standard definition of sequence converging to zero*}
-  [code del]: "Zseq X = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. norm (X n) < r)"
-
-definition
-  LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
-    ("((_)/ ----> (_))" [60, 60] 60) where
-    --{*Standard definition of convergence of sequence*}
-  [code del]: "X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))"
-
-definition
-  lim :: "(nat => 'a::real_normed_vector) => 'a" where
-    --{*Standard definition of limit using choice operator*}
-  "lim X = (THE L. X ----> L)"
-
-definition
-  convergent :: "(nat => 'a::real_normed_vector) => bool" where
-    --{*Standard definition of convergence*}
-  "convergent X = (\<exists>L. X ----> L)"
-
-definition
-  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
-    --{*Standard definition for bounded sequence*}
-  [code del]: "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
-
-definition
-  monoseq :: "(nat=>real)=>bool" where
-    --{*Definition for monotonicity*}
-  [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
-  subseq :: "(nat => nat) => bool" where
-    --{*Definition of subsequence*}
-  [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
-
-definition
-  Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
-    --{*Standard definition of the Cauchy condition*}
-  [code del]: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)"
-
-
-subsection {* Bounded Sequences *}
-
-lemma BseqI': assumes K: "\<And>n. norm (X n) \<le> K" shows "Bseq X"
-unfolding Bseq_def
-proof (intro exI conjI allI)
-  show "0 < max K 1" by simp
-next
-  fix n::nat
-  have "norm (X n) \<le> K" by (rule K)
-  thus "norm (X n) \<le> max K 1" by simp
-qed
-
-lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-unfolding Bseq_def by auto
-
-lemma BseqI2': assumes K: "\<forall>n\<ge>N. norm (X n) \<le> K" shows "Bseq X"
-proof (rule BseqI')
-  let ?A = "norm ` X ` {..N}"
-  have 1: "finite ?A" by simp
-  fix n::nat
-  show "norm (X n) \<le> max K (Max ?A)"
-  proof (cases rule: linorder_le_cases)
-    assume "n \<ge> N"
-    hence "norm (X n) \<le> K" using K by simp
-    thus "norm (X n) \<le> max K (Max ?A)" by simp
-  next
-    assume "n \<le> N"
-    hence "norm (X n) \<in> ?A" by simp
-    with 1 have "norm (X n) \<le> Max ?A" by (rule Max_ge)
-    thus "norm (X n) \<le> max K (Max ?A)" by simp
-  qed
-qed
-
-lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
-unfolding Bseq_def by auto
-
-lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
-apply (erule BseqE)
-apply (rule_tac N="k" and K="K" in BseqI2')
-apply clarify
-apply (drule_tac x="n - k" in spec, simp)
-done
-
-
-subsection {* Sequences That Converge to Zero *}
-
-lemma ZseqI:
-  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r) \<Longrightarrow> Zseq X"
-unfolding Zseq_def by simp
-
-lemma ZseqD:
-  "\<lbrakk>Zseq X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r"
-unfolding Zseq_def by simp
-
-lemma Zseq_zero: "Zseq (\<lambda>n. 0)"
-unfolding Zseq_def by simp
-
-lemma Zseq_const_iff: "Zseq (\<lambda>n. k) = (k = 0)"
-unfolding Zseq_def by force
-
-lemma Zseq_norm_iff: "Zseq (\<lambda>n. norm (X n)) = Zseq (\<lambda>n. X n)"
-unfolding Zseq_def by simp
-
-lemma Zseq_imp_Zseq:
-  assumes X: "Zseq X"
-  assumes Y: "\<And>n. norm (Y n) \<le> norm (X n) * K"
-  shows "Zseq (\<lambda>n. Y n)"
-proof (cases)
-  assume K: "0 < K"
-  show ?thesis
-  proof (rule ZseqI)
-    fix r::real assume "0 < r"
-    hence "0 < r / K"
-      using K by (rule divide_pos_pos)
-    then obtain N where "\<forall>n\<ge>N. norm (X n) < r / K"
-      using ZseqD [OF X] by fast
-    hence "\<forall>n\<ge>N. norm (X n) * K < r"
-      by (simp add: pos_less_divide_eq K)
-    hence "\<forall>n\<ge>N. norm (Y n) < r"
-      by (simp add: order_le_less_trans [OF Y])
-    thus "\<exists>N. \<forall>n\<ge>N. norm (Y n) < r" ..
-  qed
-next
-  assume "\<not> 0 < K"
-  hence K: "K \<le> 0" by (simp only: linorder_not_less)
-  {
-    fix n::nat
-    have "norm (Y n) \<le> norm (X n) * K" by (rule Y)
-    also have "\<dots> \<le> norm (X n) * 0"
-      using K norm_ge_zero by (rule mult_left_mono)
-    finally have "norm (Y n) = 0" by simp
-  }
-  thus ?thesis by (simp add: Zseq_zero)
-qed
-
-lemma Zseq_le: "\<lbrakk>Zseq Y; \<forall>n. norm (X n) \<le> norm (Y n)\<rbrakk> \<Longrightarrow> Zseq X"
-by (erule_tac K="1" in Zseq_imp_Zseq, simp)
-
-lemma Zseq_add:
-  assumes X: "Zseq X"
-  assumes Y: "Zseq Y"
-  shows "Zseq (\<lambda>n. X n + Y n)"
-proof (rule ZseqI)
-  fix r::real assume "0 < r"
-  hence r: "0 < r / 2" by simp
-  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r/2"
-    using ZseqD [OF X r] by fast
-  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < r/2"
-    using ZseqD [OF Y r] by fast
-  show "\<exists>N. \<forall>n\<ge>N. norm (X n + Y n) < r"
-  proof (intro exI allI impI)
-    fix n assume n: "max M N \<le> n"
-    have "norm (X n + Y n) \<le> norm (X n) + norm (Y n)"
-      by (rule norm_triangle_ineq)
-    also have "\<dots> < r/2 + r/2"
-    proof (rule add_strict_mono)
-      from M n show "norm (X n) < r/2" by simp
-      from N n show "norm (Y n) < r/2" by simp
-    qed
-    finally show "norm (X n + Y n) < r" by simp
-  qed
-qed
-
-lemma Zseq_minus: "Zseq X \<Longrightarrow> Zseq (\<lambda>n. - X n)"
-unfolding Zseq_def by simp
-
-lemma Zseq_diff: "\<lbrakk>Zseq X; Zseq Y\<rbrakk> \<Longrightarrow> Zseq (\<lambda>n. X n - Y n)"
-by (simp only: diff_minus Zseq_add Zseq_minus)
-
-lemma (in bounded_linear) Zseq:
-  assumes X: "Zseq X"
-  shows "Zseq (\<lambda>n. f (X n))"
-proof -
-  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
-    using bounded by fast
-  with X show ?thesis
-    by (rule Zseq_imp_Zseq)
-qed
-
-lemma (in bounded_bilinear) Zseq:
-  assumes X: "Zseq X"
-  assumes Y: "Zseq Y"
-  shows "Zseq (\<lambda>n. X n ** Y n)"
-proof (rule ZseqI)
-  fix r::real assume r: "0 < r"
-  obtain K where K: "0 < K"
-    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
-    using pos_bounded by fast
-  from K have K': "0 < inverse K"
-    by (rule positive_imp_inverse_positive)
-  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r"
-    using ZseqD [OF X r] by fast
-  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < inverse K"
-    using ZseqD [OF Y K'] by fast
-  show "\<exists>N. \<forall>n\<ge>N. norm (X n ** Y n) < r"
-  proof (intro exI allI impI)
-    fix n assume n: "max M N \<le> n"
-    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
-      by (rule norm_le)
-    also have "norm (X n) * norm (Y n) * K < r * inverse K * K"
-    proof (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero K)
-      from M n show Xn: "norm (X n) < r" by simp
-      from N n show Yn: "norm (Y n) < inverse K" by simp
-    qed
-    also from K have "r * inverse K * K = r" by simp
-    finally show "norm (X n ** Y n) < r" .
-  qed
-qed
-
-lemma (in bounded_bilinear) Zseq_prod_Bseq:
-  assumes X: "Zseq X"
-  assumes Y: "Bseq Y"
-  shows "Zseq (\<lambda>n. X n ** Y n)"
-proof -
-  obtain K where K: "0 \<le> K"
-    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
-    using nonneg_bounded by fast
-  obtain B where B: "0 < B"
-    and norm_Y: "\<And>n. norm (Y n) \<le> B"
-    using Y [unfolded Bseq_def] by fast
-  from X show ?thesis
-  proof (rule Zseq_imp_Zseq)
-    fix n::nat
-    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
-      by (rule norm_le)
-    also have "\<dots> \<le> norm (X n) * B * K"
-      by (intro mult_mono' order_refl norm_Y norm_ge_zero
-                mult_nonneg_nonneg K)
-    also have "\<dots> = norm (X n) * (B * K)"
-      by (rule mult_assoc)
-    finally show "norm (X n ** Y n) \<le> norm (X n) * (B * K)" .
-  qed
-qed
-
-lemma (in bounded_bilinear) Bseq_prod_Zseq:
-  assumes X: "Bseq X"
-  assumes Y: "Zseq Y"
-  shows "Zseq (\<lambda>n. X n ** Y n)"
-proof -
-  obtain K where K: "0 \<le> K"
-    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
-    using nonneg_bounded by fast
-  obtain B where B: "0 < B"
-    and norm_X: "\<And>n. norm (X n) \<le> B"
-    using X [unfolded Bseq_def] by fast
-  from Y show ?thesis
-  proof (rule Zseq_imp_Zseq)
-    fix n::nat
-    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
-      by (rule norm_le)
-    also have "\<dots> \<le> B * norm (Y n) * K"
-      by (intro mult_mono' order_refl norm_X norm_ge_zero
-                mult_nonneg_nonneg K)
-    also have "\<dots> = norm (Y n) * (B * K)"
-      by (simp only: mult_ac)
-    finally show "norm (X n ** Y n) \<le> norm (Y n) * (B * K)" .
-  qed
-qed
-
-lemma (in bounded_bilinear) Zseq_left:
-  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. X n ** a)"
-by (rule bounded_linear_left [THEN bounded_linear.Zseq])
-
-lemma (in bounded_bilinear) Zseq_right:
-  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. a ** X n)"
-by (rule bounded_linear_right [THEN bounded_linear.Zseq])
-
-lemmas Zseq_mult = mult.Zseq
-lemmas Zseq_mult_right = mult.Zseq_right
-lemmas Zseq_mult_left = mult.Zseq_left
-
-
-subsection {* Limits of Sequences *}
-
-lemma LIMSEQ_iff:
-      "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
-by (rule LIMSEQ_def)
-
-lemma LIMSEQ_Zseq_iff: "((\<lambda>n. X n) ----> L) = Zseq (\<lambda>n. X n - L)"
-by (simp only: LIMSEQ_def Zseq_def)
-
-lemma LIMSEQ_I:
-  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_D:
-  "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_const_iff: "(\<lambda>n. k) ----> l = (k = l)"
-by (simp add: LIMSEQ_Zseq_iff Zseq_const_iff)
-
-lemma LIMSEQ_norm: "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
-apply (simp add: LIMSEQ_def, safe)
-apply (drule_tac x="r" in spec, safe)
-apply (rule_tac x="no" in exI, safe)
-apply (drule_tac x="n" in spec, safe)
-apply (erule order_le_less_trans [OF norm_triangle_ineq3])
-done
-
-lemma LIMSEQ_ignore_initial_segment:
-  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
-apply (rule LIMSEQ_I)
-apply (drule (1) LIMSEQ_D)
-apply (erule exE, rename_tac N)
-apply (rule_tac x=N in exI)
-apply simp
-done
-
-lemma LIMSEQ_offset:
-  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
-apply (rule LIMSEQ_I)
-apply (drule (1) LIMSEQ_D)
-apply (erule exE, rename_tac N)
-apply (rule_tac x="N + k" in exI)
-apply clarify
-apply (drule_tac x="n - k" in spec)
-apply (simp add: le_diff_conv2)
-done
-
-lemma LIMSEQ_Suc: "f ----> l \<Longrightarrow> (\<lambda>n. f (Suc n)) ----> l"
-by (drule_tac k="1" in LIMSEQ_ignore_initial_segment, simp)
-
-lemma LIMSEQ_imp_Suc: "(\<lambda>n. f (Suc n)) ----> l \<Longrightarrow> f ----> l"
-by (rule_tac k="1" in LIMSEQ_offset, simp)
-
-lemma LIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l"
-by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
-
-lemma add_diff_add:
-  fixes a b c d :: "'a::ab_group_add"
-  shows "(a + c) - (b + d) = (a - b) + (c - d)"
-by simp
-
-lemma minus_diff_minus:
-  fixes a b :: "'a::ab_group_add"
-  shows "(- a) - (- b) = - (a - b)"
-by simp
-
-lemma LIMSEQ_add: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
-by (simp only: LIMSEQ_Zseq_iff add_diff_add Zseq_add)
-
-lemma LIMSEQ_minus: "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
-by (simp only: LIMSEQ_Zseq_iff minus_diff_minus Zseq_minus)
-
-lemma LIMSEQ_minus_cancel: "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
-by (drule LIMSEQ_minus, simp)
-
-lemma LIMSEQ_diff: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
-by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus)
-
-lemma LIMSEQ_unique: "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
-by (drule (1) LIMSEQ_diff, simp add: LIMSEQ_const_iff)
-
-lemma (in bounded_linear) LIMSEQ:
-  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
-by (simp only: LIMSEQ_Zseq_iff diff [symmetric] Zseq)
-
-lemma (in bounded_bilinear) LIMSEQ:
-  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n ** Y n) ----> a ** b"
-by (simp only: LIMSEQ_Zseq_iff prod_diff_prod
-               Zseq_add Zseq Zseq_left Zseq_right)
-
-lemma LIMSEQ_mult:
-  fixes a b :: "'a::real_normed_algebra"
-  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
-by (rule mult.LIMSEQ)
-
-lemma inverse_diff_inverse:
-  "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
-   \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
-by (simp add: ring_simps)
-
-lemma Bseq_inverse_lemma:
-  fixes x :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
-apply (subst nonzero_norm_inverse, clarsimp)
-apply (erule (1) le_imp_inverse_le)
-done
-
-lemma Bseq_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  assumes X: "X ----> a"
-  assumes a: "a \<noteq> 0"
-  shows "Bseq (\<lambda>n. inverse (X n))"
-proof -
-  from a have "0 < norm a" by simp
-  hence "\<exists>r>0. r < norm a" by (rule dense)
-  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
-  obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (X n - a) < r"
-    using LIMSEQ_D [OF X r1] by fast
-  show ?thesis
-  proof (rule BseqI2' [rule_format])
-    fix n assume n: "N \<le> n"
-    hence 1: "norm (X n - a) < r" by (rule N)
-    hence 2: "X n \<noteq> 0" using r2 by auto
-    hence "norm (inverse (X n)) = inverse (norm (X n))"
-      by (rule nonzero_norm_inverse)
-    also have "\<dots> \<le> inverse (norm a - r)"
-    proof (rule le_imp_inverse_le)
-      show "0 < norm a - r" using r2 by simp
-    next
-      have "norm a - norm (X n) \<le> norm (a - X n)"
-        by (rule norm_triangle_ineq2)
-      also have "\<dots> = norm (X n - a)"
-        by (rule norm_minus_commute)
-      also have "\<dots> < r" using 1 .
-      finally show "norm a - r \<le> norm (X n)" by simp
-    qed
-    finally show "norm (inverse (X n)) \<le> inverse (norm a - r)" .
-  qed
-qed
-
-lemma LIMSEQ_inverse_lemma:
-  fixes a :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>X ----> a; a \<noteq> 0; \<forall>n. X n \<noteq> 0\<rbrakk>
-         \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
-apply (subst LIMSEQ_Zseq_iff)
-apply (simp add: inverse_diff_inverse nonzero_imp_inverse_nonzero)
-apply (rule Zseq_minus)
-apply (rule Zseq_mult_left)
-apply (rule mult.Bseq_prod_Zseq)
-apply (erule (1) Bseq_inverse)
-apply (simp add: LIMSEQ_Zseq_iff)
-done
-
-lemma LIMSEQ_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  assumes X: "X ----> a"
-  assumes a: "a \<noteq> 0"
-  shows "(\<lambda>n. inverse (X n)) ----> inverse a"
-proof -
-  from a have "0 < norm a" by simp
-  then obtain k where "\<forall>n\<ge>k. norm (X n - a) < norm a"
-    using LIMSEQ_D [OF X] by fast
-  hence "\<forall>n\<ge>k. X n \<noteq> 0" by auto
-  hence k: "\<forall>n. X (n + k) \<noteq> 0" by simp
-
-  from X have "(\<lambda>n. X (n + k)) ----> a"
-    by (rule LIMSEQ_ignore_initial_segment)
-  hence "(\<lambda>n. inverse (X (n + k))) ----> inverse a"
-    using a k by (rule LIMSEQ_inverse_lemma)
-  thus "(\<lambda>n. inverse (X n)) ----> inverse a"
-    by (rule LIMSEQ_offset)
-qed
-
-lemma LIMSEQ_divide:
-  fixes a b :: "'a::real_normed_field"
-  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
-by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)
-
-lemma LIMSEQ_pow:
-  fixes a :: "'a::{real_normed_algebra,recpower}"
-  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
-by (induct m) (simp_all add: power_Suc LIMSEQ_const LIMSEQ_mult)
-
-lemma LIMSEQ_setsum:
-  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
-  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
-proof (cases "finite S")
-  case True
-  thus ?thesis using n
-  proof (induct)
-    case empty
-    show ?case
-      by (simp add: LIMSEQ_const)
-  next
-    case insert
-    thus ?case
-      by (simp add: LIMSEQ_add)
-  qed
-next
-  case False
-  thus ?thesis
-    by (simp add: LIMSEQ_const)
-qed
-
-lemma LIMSEQ_setprod:
-  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
-  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
-  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)"
-proof (cases "finite S")
-  case True
-  thus ?thesis using n
-  proof (induct)
-    case empty
-    show ?case
-      by (simp add: LIMSEQ_const)
-  next
-    case insert
-    thus ?case
-      by (simp add: LIMSEQ_mult)
-  qed
-next
-  case False
-  thus ?thesis
-    by (simp add: setprod_def LIMSEQ_const)
-qed
-
-lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b"
-by (simp add: LIMSEQ_add LIMSEQ_const)
-
-(* FIXME: delete *)
-lemma LIMSEQ_add_minus:
-     "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
-by (simp only: LIMSEQ_add LIMSEQ_minus)
-
-lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n  - b)) ----> a - b"
-by (simp add: LIMSEQ_diff LIMSEQ_const)
-
-lemma LIMSEQ_diff_approach_zero: 
-  "g ----> L ==> (%x. f x - g x) ----> 0  ==>
-     f ----> L"
-  apply (drule LIMSEQ_add)
-  apply assumption
-  apply simp
-done
-
-lemma LIMSEQ_diff_approach_zero2: 
-  "f ----> L ==> (%x. f x - g x) ----> 0  ==>
-     g ----> L";
-  apply (drule LIMSEQ_diff)
-  apply assumption
-  apply simp
-done
-
-text{*A sequence tends to zero iff its abs does*}
-lemma LIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----> 0) = (X ----> 0)"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
-by (drule LIMSEQ_norm, simp)
-
-text{*An unbounded sequence's inverse tends to 0*}
-
-lemma LIMSEQ_inverse_zero:
-  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
-apply (rule LIMSEQ_I)
-apply (drule_tac x="inverse r" in spec, safe)
-apply (rule_tac x="N" in exI, safe)
-apply (drule_tac x="n" in spec, safe)
-apply (frule positive_imp_inverse_positive)
-apply (frule (1) less_imp_inverse_less)
-apply (subgoal_tac "0 < X n", simp)
-apply (erule (1) order_less_trans)
-done
-
-text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
-
-lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
-apply (rule LIMSEQ_inverse_zero, safe)
-apply (cut_tac x = r in reals_Archimedean2)
-apply (safe, rule_tac x = n in exI)
-apply (auto simp add: real_of_nat_Suc)
-done
-
-text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
-infinity is now easily proved*}
-
-lemma LIMSEQ_inverse_real_of_nat_add:
-     "(%n. r + inverse(real(Suc n))) ----> r"
-by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
-
-lemma LIMSEQ_inverse_real_of_nat_add_minus:
-     "(%n. r + -inverse(real(Suc n))) ----> r"
-by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
-
-lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
-     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
-by (cut_tac b=1 in
-        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)
-
-lemma LIMSEQ_le_const:
-  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
-apply (rule ccontr, simp only: linorder_not_le)
-apply (drule_tac r="a - x" in LIMSEQ_D, simp)
-apply clarsimp
-apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
-apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
-apply simp
-done
-
-lemma LIMSEQ_le_const2:
-  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
-apply (subgoal_tac "- a \<le> - x", simp)
-apply (rule LIMSEQ_le_const)
-apply (erule LIMSEQ_minus)
-apply simp
-done
-
-lemma LIMSEQ_le:
-  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
-apply (subgoal_tac "0 \<le> y - x", simp)
-apply (rule LIMSEQ_le_const)
-apply (erule (1) LIMSEQ_diff)
-apply (simp add: le_diff_eq)
-done
-
-
-subsection {* Convergence *}
-
-lemma limI: "X ----> L ==> lim X = L"
-apply (simp add: lim_def)
-apply (blast intro: LIMSEQ_unique)
-done
-
-lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
-by (simp add: convergent_def)
-
-lemma convergentI: "(X ----> L) ==> convergent X"
-by (auto simp add: convergent_def)
-
-lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
-by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def)
-
-lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))"
-apply (simp add: convergent_def)
-apply (auto dest: LIMSEQ_minus)
-apply (drule LIMSEQ_minus, auto)
-done
-
-
-subsection {* Bounded Monotonic Sequences *}
-
-text{*Subsequence (alternative definition, (e.g. Hoskins)*}
-
-lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
-apply (simp add: subseq_def)
-apply (auto dest!: less_imp_Suc_add)
-apply (induct_tac k)
-apply (auto intro: less_trans)
-done
-
-lemma monoseq_Suc:
-   "monoseq X = ((\<forall>n. X n \<le> X (Suc n))
-                 | (\<forall>n. X (Suc n) \<le> X n))"
-apply (simp add: monoseq_def)
-apply (auto dest!: le_imp_less_or_eq)
-apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
-apply (induct_tac "ka")
-apply (auto intro: order_trans)
-apply (erule contrapos_np)
-apply (induct_tac "k")
-apply (auto intro: order_trans)
-done
-
-lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
-by (simp add: monoseq_def)
-
-lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
-by (simp add: monoseq_def)
-
-lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
-by (simp add: monoseq_Suc)
-
-lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
-by (simp add: monoseq_Suc)
-
-text{*Bounded Sequence*}
-
-lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
-by (simp add: Bseq_def)
-
-lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
-by (auto simp add: Bseq_def)
-
-lemma lemma_NBseq_def:
-     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) =
-      (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
-apply auto
- prefer 2 apply force
-apply (cut_tac x = K in reals_Archimedean2, clarify)
-apply (rule_tac x = n in exI, clarify)
-apply (drule_tac x = na in spec)
-apply (auto simp add: real_of_nat_Suc)
-done
-
-text{* alternative definition for Bseq *}
-lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
-apply (simp add: Bseq_def)
-apply (simp (no_asm) add: lemma_NBseq_def)
-done
-
-lemma lemma_NBseq_def2:
-     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
-apply (subst lemma_NBseq_def, auto)
-apply (rule_tac x = "Suc N" in exI)
-apply (rule_tac [2] x = N in exI)
-apply (auto simp add: real_of_nat_Suc)
- prefer 2 apply (blast intro: order_less_imp_le)
-apply (drule_tac x = n in spec, simp)
-done
-
-(* yet another definition for Bseq *)
-lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
-by (simp add: Bseq_def lemma_NBseq_def2)
-
-subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
-
-lemma Bseq_isUb:
-  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
-by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
-
-
-text{* Use completeness of reals (supremum property)
-   to show that any bounded sequence has a least upper bound*}
-
-lemma Bseq_isLub:
-  "!!(X::nat=>real). Bseq X ==>
-   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
-by (blast intro: reals_complete Bseq_isUb)
-
-subsubsection{*A Bounded and Monotonic Sequence Converges*}
-
-lemma lemma_converg1:
-     "!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n;
-                  isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma)
-               |] ==> \<forall>n \<ge> ma. X n = X ma"
-apply safe
-apply (drule_tac y = "X n" in isLubD2)
-apply (blast dest: order_antisym)+
-done
-
-text{* The best of both worlds: Easier to prove this result as a standard
-   theorem and then use equivalence to "transfer" it into the
-   equivalent nonstandard form if needed!*}
-
-lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
-apply (simp add: LIMSEQ_def)
-apply (rule_tac x = "X m" in exI, safe)
-apply (rule_tac x = m in exI, safe)
-apply (drule spec, erule impE, auto)
-done
-
-lemma lemma_converg2:
-   "!!(X::nat=>real).
-    [| \<forall>m. X m ~= U;  isLub UNIV {x. \<exists>n. X n = x} U |] ==> \<forall>m. X m < U"
-apply safe
-apply (drule_tac y = "X m" in isLubD2)
-apply (auto dest!: order_le_imp_less_or_eq)
-done
-
-lemma lemma_converg3: "!!(X ::nat=>real). \<forall>m. X m \<le> U ==> isUb UNIV {x. \<exists>n. X n = x} U"
-by (rule setleI [THEN isUbI], auto)
-
-text{* FIXME: @{term "U - T < U"} is redundant *}
-lemma lemma_converg4: "!!(X::nat=> real).
-               [| \<forall>m. X m ~= U;
-                  isLub UNIV {x. \<exists>n. X n = x} U;
-                  0 < T;
-                  U + - T < U
-               |] ==> \<exists>m. U + -T < X m & X m < U"
-apply (drule lemma_converg2, assumption)
-apply (rule ccontr, simp)
-apply (simp add: linorder_not_less)
-apply (drule lemma_converg3)
-apply (drule isLub_le_isUb, assumption)
-apply (auto dest: order_less_le_trans)
-done
-
-text{*A standard proof of the theorem for monotone increasing sequence*}
-
-lemma Bseq_mono_convergent:
-     "[| Bseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> convergent (X::nat=>real)"
-apply (simp add: convergent_def)
-apply (frule Bseq_isLub, safe)
-apply (case_tac "\<exists>m. X m = U", auto)
-apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ)
-(* second case *)
-apply (rule_tac x = U in exI)
-apply (subst LIMSEQ_iff, safe)
-apply (frule lemma_converg2, assumption)
-apply (drule lemma_converg4, auto)
-apply (rule_tac x = m in exI, safe)
-apply (subgoal_tac "X m \<le> X n")
- prefer 2 apply blast
-apply (drule_tac x=n and P="%m. X m < U" in spec, arith)
-done
-
-lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
-by (simp add: Bseq_def)
-
-text{*Main monotonicity theorem*}
-lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X"
-apply (simp add: monoseq_def, safe)
-apply (rule_tac [2] convergent_minus_iff [THEN ssubst])
-apply (drule_tac [2] Bseq_minus_iff [THEN ssubst])
-apply (auto intro!: Bseq_mono_convergent)
-done
-
-subsubsection{*A Few More Equivalence Theorems for Boundedness*}
-
-text{*alternative formulation for boundedness*}
-lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
-apply (unfold Bseq_def, safe)
-apply (rule_tac [2] x = "k + norm x" in exI)
-apply (rule_tac x = K in exI, simp)
-apply (rule exI [where x = 0], auto)
-apply (erule order_less_le_trans, simp)
-apply (drule_tac x=n in spec, fold diff_def)
-apply (drule order_trans [OF norm_triangle_ineq2])
-apply simp
-done
-
-text{*alternative formulation for boundedness*}
-lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
-apply safe
-apply (simp add: Bseq_def, safe)
-apply (rule_tac x = "K + norm (X N)" in exI)
-apply auto
-apply (erule order_less_le_trans, simp)
-apply (rule_tac x = N in exI, safe)
-apply (drule_tac x = n in spec)
-apply (rule order_trans [OF norm_triangle_ineq], simp)
-apply (auto simp add: Bseq_iff2)
-done
-
-lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
-apply (simp add: Bseq_def)
-apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
-apply (drule_tac x = n in spec, arith)
-done
-
-
-subsection {* Cauchy Sequences *}
-
-lemma CauchyI:
-  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
-by (simp add: Cauchy_def)
-
-lemma CauchyD:
-  "\<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)
-
-subsubsection {* Cauchy Sequences are Bounded *}
-
-text{*A Cauchy sequence is bounded -- this is the standard
-  proof mechanization rather than the nonstandard proof*}
-
-lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
-          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
-apply (clarify, drule spec, drule (1) mp)
-apply (simp only: norm_minus_commute)
-apply (drule order_le_less_trans [OF norm_triangle_ineq2])
-apply simp
-done
-
-lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
-apply (simp add: Cauchy_def)
-apply (drule spec, drule mp, rule zero_less_one, safe)
-apply (drule_tac x="M" in spec, simp)
-apply (drule lemmaCauchy)
-apply (rule_tac k="M" in Bseq_offset)
-apply (simp add: Bseq_def)
-apply (rule_tac x="1 + norm (X M)" in exI)
-apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp)
-apply (simp add: order_less_imp_le)
-done
-
-subsubsection {* Cauchy Sequences are Convergent *}
-
-axclass banach \<subseteq> real_normed_vector
-  Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
-
-theorem LIMSEQ_imp_Cauchy:
-  assumes X: "X ----> a" shows "Cauchy X"
-proof (rule CauchyI)
-  fix e::real assume "0 < e"
-  hence "0 < e/2" by simp
-  with X have "\<exists>N. \<forall>n\<ge>N. norm (X n - a) < e/2" by (rule LIMSEQ_D)
-  then obtain N where N: "\<forall>n\<ge>N. norm (X n - a) < e/2" ..
-  show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < e"
-  proof (intro exI allI impI)
-    fix m assume "N \<le> m"
-    hence m: "norm (X m - a) < e/2" using N by fast
-    fix n assume "N \<le> n"
-    hence n: "norm (X n - a) < e/2" using N by fast
-    have "norm (X m - X n) = norm ((X m - a) - (X n - a))" by simp
-    also have "\<dots> \<le> norm (X m - a) + norm (X n - a)"
-      by (rule norm_triangle_ineq4)
-    also from m n have "\<dots> < e" by(simp add:field_simps)
-    finally show "norm (X m - X n) < e" .
-  qed
-qed
-
-lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
-unfolding convergent_def
-by (erule exE, erule LIMSEQ_imp_Cauchy)
-
-text {*
-Proof that Cauchy sequences converge based on the one from
-http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
-*}
-
-text {*
-  If sequence @{term "X"} is Cauchy, then its limit is the lub of
-  @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
-*}
-
-lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
-by (simp add: isUbI setleI)
-
-lemma real_abs_diff_less_iff:
-  "(\<bar>x - a\<bar> < (r::real)) = (a - r < x \<and> x < a + r)"
-by auto
-
-locale real_Cauchy =
-  fixes X :: "nat \<Rightarrow> real"
-  assumes X: "Cauchy X"
-  fixes S :: "real set"
-  defines S_def: "S \<equiv> {x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
-
-lemma real_CauchyI:
-  assumes "Cauchy X"
-  shows "real_Cauchy X"
-  proof qed (fact assms)
-
-lemma (in real_Cauchy) mem_S: "\<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S"
-by (unfold S_def, auto)
-
-lemma (in real_Cauchy) bound_isUb:
-  assumes N: "\<forall>n\<ge>N. X n < x"
-  shows "isUb UNIV S x"
-proof (rule isUb_UNIV_I)
-  fix y::real assume "y \<in> S"
-  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
-    by (simp add: S_def)
-  then obtain M where "\<forall>n\<ge>M. y < X n" ..
-  hence "y < X (max M N)" by simp
-  also have "\<dots> < x" using N by simp
-  finally show "y \<le> x"
-    by (rule order_less_imp_le)
-qed
-
-lemma (in real_Cauchy) isLub_ex: "\<exists>u. isLub UNIV S u"
-proof (rule reals_complete)
-  obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < 1"
-    using CauchyD [OF X zero_less_one] by fast
-  hence N: "\<forall>n\<ge>N. norm (X n - X N) < 1" by simp
-  show "\<exists>x. x \<in> S"
-  proof
-    from N have "\<forall>n\<ge>N. X N - 1 < X n"
-      by (simp add: real_abs_diff_less_iff)
-    thus "X N - 1 \<in> S" by (rule mem_S)
-  qed
-  show "\<exists>u. isUb UNIV S u"
-  proof
-    from N have "\<forall>n\<ge>N. X n < X N + 1"
-      by (simp add: real_abs_diff_less_iff)
-    thus "isUb UNIV S (X N + 1)"
-      by (rule bound_isUb)
-  qed
-qed
-
-lemma (in real_Cauchy) isLub_imp_LIMSEQ:
-  assumes x: "isLub UNIV S x"
-  shows "X ----> x"
-proof (rule LIMSEQ_I)
-  fix r::real assume "0 < r"
-  hence r: "0 < r/2" by simp
-  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. norm (X n - X m) < r/2"
-    using CauchyD [OF X r] by fast
-  hence "\<forall>n\<ge>N. norm (X n - X N) < r/2" by simp
-  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
-    by (simp only: real_norm_def real_abs_diff_less_iff)
-
-  from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
-  hence "X N - r/2 \<in> S" by (rule mem_S)
-  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
-
-  from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
-  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
-  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
-
-  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
-  proof (intro exI allI impI)
-    fix n assume n: "N \<le> n"
-    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
-    thus "norm (X n - x) < r" using 1 2
-      by (simp add: real_abs_diff_less_iff)
-  qed
-qed
-
-lemma (in real_Cauchy) LIMSEQ_ex: "\<exists>x. X ----> x"
-proof -
-  obtain x where "isLub UNIV S x"
-    using isLub_ex by fast
-  hence "X ----> x"
-    by (rule isLub_imp_LIMSEQ)
-  thus ?thesis ..
-qed
-
-lemma real_Cauchy_convergent:
-  fixes X :: "nat \<Rightarrow> real"
-  shows "Cauchy X \<Longrightarrow> convergent X"
-unfolding convergent_def
-by (rule real_Cauchy.LIMSEQ_ex)
- (rule real_CauchyI)
-
-instance real :: banach
-by intro_classes (rule real_Cauchy_convergent)
-
-lemma Cauchy_convergent_iff:
-  fixes X :: "nat \<Rightarrow> 'a::banach"
-  shows "Cauchy X = convergent X"
-by (fast intro: Cauchy_convergent convergent_Cauchy)
-
-
-subsection {* Power Sequences *}
-
-text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
-"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
-  also fact that bounded and monotonic sequence converges.*}
-
-lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
-apply (simp add: Bseq_def)
-apply (rule_tac x = 1 in exI)
-apply (simp add: power_abs)
-apply (auto dest: power_mono)
-done
-
-lemma monoseq_realpow: "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
-apply (clarify intro!: mono_SucI2)
-apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
-done
-
-lemma convergent_realpow:
-  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
-by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
-
-lemma LIMSEQ_inverse_realpow_zero_lemma:
-  fixes x :: real
-  assumes x: "0 \<le> x"
-  shows "real n * x + 1 \<le> (x + 1) ^ n"
-apply (induct n)
-apply simp
-apply simp
-apply (rule order_trans)
-prefer 2
-apply (erule mult_left_mono)
-apply (rule add_increasing [OF x], simp)
-apply (simp add: real_of_nat_Suc)
-apply (simp add: ring_distribs)
-apply (simp add: mult_nonneg_nonneg x)
-done
-
-lemma LIMSEQ_inverse_realpow_zero:
-  "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
-proof (rule LIMSEQ_inverse_zero [rule_format])
-  fix y :: real
-  assume x: "1 < x"
-  hence "0 < x - 1" by simp
-  hence "\<forall>y. \<exists>N::nat. y < real N * (x - 1)"
-    by (rule reals_Archimedean3)
-  hence "\<exists>N::nat. y < real N * (x - 1)" ..
-  then obtain N::nat where "y < real N * (x - 1)" ..
-  also have "\<dots> \<le> real N * (x - 1) + 1" by simp
-  also have "\<dots> \<le> (x - 1 + 1) ^ N"
-    by (rule LIMSEQ_inverse_realpow_zero_lemma, cut_tac x, simp)
-  also have "\<dots> = x ^ N" by simp
-  finally have "y < x ^ N" .
-  hence "\<forall>n\<ge>N. y < x ^ n"
-    apply clarify
-    apply (erule order_less_le_trans)
-    apply (erule power_increasing)
-    apply (rule order_less_imp_le [OF x])
-    done
-  thus "\<exists>N. \<forall>n\<ge>N. y < x ^ n" ..
-qed
-
-lemma LIMSEQ_realpow_zero:
-  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
-proof (cases)
-  assume "x = 0"
-  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
-  thus ?thesis by (rule LIMSEQ_imp_Suc)
-next
-  assume "0 \<le> x" and "x \<noteq> 0"
-  hence x0: "0 < x" by simp
-  assume x1: "x < 1"
-  from x0 x1 have "1 < inverse x"
-    by (rule real_inverse_gt_one)
-  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
-    by (rule LIMSEQ_inverse_realpow_zero)
-  thus ?thesis by (simp add: power_inverse)
-qed
-
-lemma LIMSEQ_power_zero:
-  fixes x :: "'a::{real_normed_algebra_1,recpower}"
-  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
-apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
-apply (simp only: LIMSEQ_Zseq_iff, erule Zseq_le)
-apply (simp add: power_abs norm_power_ineq)
-done
-
-lemma LIMSEQ_divide_realpow_zero:
-  "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0"
-apply (cut_tac a = a and x1 = "inverse x" in
-        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero])
-apply (auto simp add: divide_inverse power_inverse)
-apply (simp add: inverse_eq_divide pos_divide_less_eq)
-done
-
-text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
-
-lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----> 0"
-by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
-
-lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----> 0"
-apply (rule LIMSEQ_rabs_zero [THEN iffD1])
-apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs)
-done
-
-end
--- a/src/HOL/Import/proof_kernel.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Import/proof_kernel.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*  Title:      HOL/Import/proof_kernel.ML
-    ID:         $Id$
-    Author:     Sebastian Skalberg (TU Muenchen), Steven Obua
+    Author:     Sebastian Skalberg and Steven Obua, TU Muenchen
 *)
 
 signature ProofKernel =
@@ -281,14 +280,12 @@
           | itr (a::rst) = i=a orelse itr rst
     in itr L end;
 
-fun insert i L = if i mem L then L else i::L
-
 fun mk_set [] = []
-  | mk_set (a::rst) = insert a (mk_set rst)
+  | mk_set (a::rst) = insert (op =) a (mk_set rst)
 
 fun [] union S = S
   | S union [] = S
-  | (a::rst) union S2 = rst union (insert a S2)
+  | (a::rst) union S2 = rst union (insert (op =) a S2)
 
 fun implode_subst [] = []
   | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest))
@@ -1149,9 +1146,9 @@
       val ct = cprop_of thm
       val t = term_of ct
       val thy = theory_of_cterm ct
-      val frees = term_frees t
-      val freenames = add_term_free_names (t, [])
-      fun is_old_name n = n mem_string freenames
+      val frees = OldTerm.term_frees t
+      val freenames = Term.add_free_names t []
+      val is_old_name = member (op =) freenames
       fun name_of (Free (n, _)) = n
         | name_of _ = ERR "name_of"
       fun new_name' bump map n =
@@ -1214,24 +1211,23 @@
               | NONE => NONE
         end
 
-fun non_trivial_term_consts tm =
-    List.filter (fn c => not (c = "Trueprop" orelse
-                         c = "All" orelse
-                         c = "op -->" orelse
-                         c = "op &" orelse
-                         c = "op =")) (Term.term_consts tm)
+fun non_trivial_term_consts t = fold_aterms
+  (fn Const (c, _) =>
+      if c = "Trueprop" orelse c = "All" orelse c = "op -->" orelse c = "op &" orelse c = "op ="
+      then I else insert (op =) c
+    | _ => I) t [];
 
 fun match_consts t (* th *) =
     let
         fun add_consts (Const (c, _), cs) =
             (case c of
-                 "op =" => Library.insert (op =) "==" cs
-               | "op -->" => Library.insert (op =) "==>" cs
+                 "op =" => insert (op =) "==" cs
+               | "op -->" => insert (op =) "==>" cs
                | "All" => cs
                | "all" => cs
                | "op &" => cs
                | "Trueprop" => cs
-               | _ => Library.insert (op =) c cs)
+               | _ => insert (op =) c cs)
           | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
           | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
           | add_consts (_, cs) = cs
@@ -1298,7 +1294,7 @@
             let
                 val _ = (message "Looking for conclusion:";
                          if_debug prin isaconc)
-                val cs = non_trivial_term_consts isaconc
+                val cs = non_trivial_term_consts isaconc;
                 val _ = (message "Looking for consts:";
                          message (commas cs))
                 val pot_thms = Shuffler.find_potential thy isaconc
@@ -1424,9 +1420,9 @@
     let
         val _ = message "INST_TYPE:"
         val _ = if_debug pth hth
-        val tys_before = add_term_tfrees (prop_of th,[])
+        val tys_before = OldTerm.add_term_tfrees (prop_of th,[])
         val th1 = Thm.varifyT th
-        val tys_after = add_term_tvars (prop_of th1,[])
+        val tys_after = OldTerm.add_term_tvars (prop_of th1,[])
         val tyinst = map (fn (bef, iS) =>
                              (case try (Lib.assoc (TFree bef)) lambda of
                                   SOME ty => (ctyp_of thy (TVar iS), ctyp_of thy ty)
@@ -2093,7 +2089,7 @@
             val c = case concl_of th2 of
                         _ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
                       | _ => raise ERR "new_type_definition" "Bad type definition theorem"
-            val tfrees = term_tfrees c
+            val tfrees = OldTerm.term_tfrees c
             val tnames = map fst tfrees
             val tsyn = mk_syn thy tycname
             val typ = (tycname,tnames,tsyn)
@@ -2179,7 +2175,7 @@
             val c = case concl_of th2 of
                         _ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
                       | _ => raise ERR "type_introduction" "Bad type definition theorem"
-            val tfrees = term_tfrees c
+            val tfrees = OldTerm.term_tfrees c
             val tnames = sort string_ord (map fst tfrees)
             val tsyn = mk_syn thy tycname
             val typ = (tycname,tnames,tsyn)
--- a/src/HOL/Import/shuffler.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Import/shuffler.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Import/shuffler.ML
-    ID:         $Id$
     Author:     Sebastian Skalberg, TU Muenchen
 
 Package for proving two terms equal by normalizing (hence the
@@ -248,8 +247,8 @@
 
 fun freeze_thaw_term t =
     let
-        val tvars = term_tvars t
-        val tfree_names = add_term_tfree_names(t,[])
+        val tvars = OldTerm.term_tvars t
+        val tfree_names = OldTerm.add_term_tfree_names(t,[])
         val (type_inst,_) =
             Library.foldl (fn ((inst,used),(w as (v,_),S)) =>
                       let
@@ -268,7 +267,7 @@
   | inst_tfrees thy ((name,U)::rest) thm =
     let
         val cU = ctyp_of thy U
-        val tfrees = add_term_tfrees (prop_of thm,[])
+        val tfrees = OldTerm.add_term_tfrees (prop_of thm,[])
         val (rens, thm') = Thm.varifyT'
     (remove (op = o apsnd fst) name tfrees) thm
         val mid =
@@ -322,7 +321,7 @@
               then
                   let
                       val cert = cterm_of thy
-                      val v = Free(Name.variant (add_term_free_names(t,[])) "v",xT)
+                      val v = Free (Name.variant (Term.add_free_names t []) "v", xT)
                       val cv = cert v
                       val ct = cert t
                       val th = (assume ct)
@@ -354,7 +353,7 @@
 
 fun equals_fun thy assume t =
     case t of
-        Const("op ==",_) $ u $ v => if Term.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE
+        Const("op ==",_) $ u $ v => if TermOrd.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE
       | _ => NONE
 
 fun eta_expand thy assumes origt =
@@ -385,7 +384,7 @@
                       Type("fun",[aT,bT]) =>
                       let
                           val cert = cterm_of thy
-                          val vname = Name.variant (add_term_free_names(t,[])) "v"
+                          val vname = Name.variant (Term.add_free_names t []) "v"
                           val v = Free(vname,aT)
                           val cv = cert v
                           val ct = cert t
@@ -520,7 +519,7 @@
     let
         val thy = Thm.theory_of_thm th
         val c = prop_of th
-        val vars = add_term_frees (c,add_term_vars(c,[]))
+        val vars = OldTerm.add_term_frees (c, OldTerm.add_term_vars(c,[]))
     in
         Drule.forall_intr_list (map (cterm_of thy) vars) th
     end
@@ -573,8 +572,8 @@
     fold_rev (fn thm => fn cs =>
               let
                   val (lhs,rhs) = Logic.dest_equals (prop_of thm)
-                  val ignore_lhs = term_consts lhs \\ term_consts rhs
-                  val ignore_rhs = term_consts rhs \\ term_consts lhs
+                  val ignore_lhs = Term.add_const_names lhs [] \\ Term.add_const_names rhs []
+                  val ignore_rhs = Term.add_const_names rhs [] \\ Term.add_const_names lhs []
               in
                   fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs)
               end)
@@ -585,7 +584,7 @@
 
 fun set_prop thy t =
     let
-        val vars = add_term_frees (t,add_term_vars (t,[]))
+        val vars = OldTerm.add_term_frees (t, OldTerm.add_term_vars (t,[]))
         val closed_t = fold_rev Logic.all vars t
         val rew_th = norm_term thy closed_t
         val rhs = Thm.rhs_of rew_th
--- a/src/HOL/Inductive.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Inductive.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Inductive.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
@@ -363,7 +362,7 @@
 let
   fun fun_tr ctxt [cs] =
     let
-      val x = Free (Name.variant (add_term_free_names (cs, [])) "x", dummyT);
+      val x = Free (Name.variant (Term.add_free_names cs []) "x", dummyT);
       val ft = DatatypeCase.case_tr true DatatypePackage.datatype_of_constr
                  ctxt [x, cs]
     in lambda x ft end
--- a/src/HOL/IsaMakefile	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/IsaMakefile	Mon Jan 05 07:54:16 2009 -0800
@@ -261,7 +261,7 @@
 $(OUT)/HOL: ROOT.ML $(MAIN_DEPENDENCIES) \
   Complex_Main.thy \
   Complex.thy \
-  Complex/Fundamental_Theorem_Algebra.thy \
+  Fundamental_Theorem_Algebra.thy \
   Deriv.thy \
   Fact.thy \
   FrechetDeriv.thy \
@@ -271,11 +271,11 @@
   Log.thy \
   MacLaurin.thy \
   NthRoot.thy \
-  Hyperreal/SEQ.thy \
+  SEQ.thy \
   Series.thy \
   Taylor.thy \
   Transcendental.thy \
-  Library/Dense_Linear_Order.thy \
+  Dense_Linear_Order.thy \
   GCD.thy \
   Order_Relation.thy \
   Parity.thy \
@@ -287,7 +287,7 @@
   RealDef.thy \
   RealPow.thy \
   Real.thy \
-  Real/RealVector.thy \
+  RealVector.thy \
   Tools/float_syntax.ML \
   Tools/rat_arith.ML \
   Tools/real_arith.ML \
@@ -337,16 +337,16 @@
 HOL-HahnBanach: HOL $(LOG)/HOL-HahnBanach.gz
 
 $(LOG)/HOL-HahnBanach.gz: $(OUT)/HOL			\
-  Real/HahnBanach/Bounds.thy Real/HahnBanach/FunctionNorm.thy		\
-  Real/HahnBanach/FunctionOrder.thy Real/HahnBanach/HahnBanach.thy	\
-  Real/HahnBanach/HahnBanachExtLemmas.thy				\
-  Real/HahnBanach/HahnBanachSupLemmas.thy				\
-  Real/HahnBanach/Linearform.thy Real/HahnBanach/NormedSpace.thy	\
-  Real/HahnBanach/README.html Real/HahnBanach/ROOT.ML			\
-  Real/HahnBanach/Subspace.thy Real/HahnBanach/VectorSpace.thy		\
-  Real/HahnBanach/ZornLemma.thy Real/HahnBanach/document/root.bib	\
-  Real/HahnBanach/document/root.tex
-	@cd Real; $(ISABELLE_TOOL) usedir -g true $(OUT)/HOL HahnBanach
+  HahnBanach/Bounds.thy HahnBanach/FunctionNorm.thy		\
+  HahnBanach/FunctionOrder.thy HahnBanach/HahnBanach.thy	\
+  HahnBanach/HahnBanachExtLemmas.thy				\
+  HahnBanach/HahnBanachSupLemmas.thy				\
+  HahnBanach/Linearform.thy HahnBanach/NormedSpace.thy	\
+  HahnBanach/README.html HahnBanach/ROOT.ML			\
+  HahnBanach/Subspace.thy HahnBanach/VectorSpace.thy		\
+  HahnBanach/ZornLemma.thy HahnBanach/document/root.bib	\
+  HahnBanach/document/root.tex
+	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL HahnBanach
 
 
 ## HOL-Subst
--- a/src/HOL/Lattices.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Lattices.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -300,7 +300,7 @@
   by auto
 qed (auto simp add: min_def max_def not_le less_imp_le)
 
-interpretation min_max:
+class_interpretation min_max:
   distrib_lattice ["op \<le> \<Colon> 'a\<Colon>linorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" min max]
   by (rule distrib_lattice_min_max)
 
--- a/src/HOL/Library/Dense_Linear_Order.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,880 +0,0 @@
-(*
-    ID:         $Id$
-    Author:     Amine Chaieb, TU Muenchen
-*)
-
-header {* Dense linear order without endpoints
-  and a quantifier elimination procedure in Ferrante and Rackoff style *}
-
-theory Dense_Linear_Order
-imports Plain "~~/src/HOL/Groebner_Basis"
-uses
-  "~~/src/HOL/Tools/Qelim/langford_data.ML"
-  "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML"
-  ("~~/src/HOL/Tools/Qelim/langford.ML")
-  ("~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML")
-begin
-
-setup {* Langford_Data.setup #> Ferrante_Rackoff_Data.setup *}
-
-context linorder
-begin
-
-lemma less_not_permute: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear)
-
-lemma gather_simps: 
-  shows 
-  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)"
-  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)"
-  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))"
-  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))"  by auto
-
-lemma 
-  gather_start: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
-  by simp
-
-text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
-lemma minf_lt:  "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto
-lemma minf_gt: "\<exists>z . \<forall>x. x < z \<longrightarrow>  (t < x \<longleftrightarrow>  False)"
-  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
-
-lemma minf_le: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le)
-lemma minf_ge: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)"
-  by (auto simp add: less_le not_less not_le)
-lemma minf_eq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
-lemma minf_neq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
-lemma minf_P: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
-
-text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
-lemma pinf_gt:  "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
-lemma pinf_lt: "\<exists>z . \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
-  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
-
-lemma pinf_ge: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le)
-lemma pinf_le: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)"
-  by (auto simp add: less_le not_less not_le)
-lemma pinf_eq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
-lemma pinf_neq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
-lemma pinf_P: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast
-
-lemma nmi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)"
-  by (auto simp add: le_less)
-lemma  nmi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x\<le> t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma  nmi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t\<le> x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma  nmi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma  nmi_neq: "t \<in> U \<Longrightarrow>\<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma  nmi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma  nmi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
-  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
-  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma  nmi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
-  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
-  \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-
-lemma  npi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x < t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by (auto simp add: le_less)
-lemma  npi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-lemma  npi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x \<le> t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-lemma  npi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t \<le> x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-lemma  npi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-lemma  npi_neq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u )" by auto
-lemma  npi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-lemma  npi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ;  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
-  \<Longrightarrow>  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-lemma  npi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
-  \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-
-lemma lin_dense_lt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x < t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y < t)"
-proof(clarsimp)
-  fix x l u y  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x"
-    and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "t < y"
-    from less_trans[OF lx px] less_trans[OF H yu]
-    have "l < t \<and> t < u"  by simp
-    with tU noU have "False" by auto}
-  hence "\<not> t < y"  by auto hence "y \<le> t" by (simp add: not_less)
-  thus "y < t" using tny by (simp add: less_le)
-qed
-
-lemma lin_dense_gt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> t < x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t < y)"
-proof(clarsimp)
-  fix x l u y
-  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
-  and px: "t < x" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "y< t"
-    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> y<t"  by auto hence "t \<le> y" by (auto simp add: not_less)
-  thus "t < y" using tny by (simp add:less_le)
-qed
-
-lemma lin_dense_le: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<le> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<le> t)"
-proof(clarsimp)
-  fix x l u y
-  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
-  and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "t < y"
-    from less_le_trans[OF lx px] less_trans[OF H yu]
-    have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> t < y"  by auto thus "y \<le> t" by (simp add: not_less)
-qed
-
-lemma lin_dense_ge: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> t \<le> x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t \<le> y)"
-proof(clarsimp)
-  fix x l u y
-  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
-  and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "y< t"
-    from less_trans[OF ly H] le_less_trans[OF px xu]
-    have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> y<t"  by auto thus "t \<le> y" by (simp add: not_less)
-qed
-lemma lin_dense_eq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x = t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y= t)"  by auto
-lemma lin_dense_neq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<noteq> t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<noteq> t)"  by auto
-lemma lin_dense_P: "\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P)"  by auto
-
-lemma lin_dense_conj:
-  "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
-  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
-  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<and> P2 x)
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))"
-  by blast
-lemma lin_dense_disj:
-  "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
-  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
-  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<or> P2 x)
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))"
-  by blast
-
-lemma npmibnd: "\<lbrakk>\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<le> x); \<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk>
-  \<Longrightarrow> \<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<le> x \<and> x \<le> u')"
-by auto
-
-lemma finite_set_intervals:
-  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
-  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
-  shows "\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x"
-proof-
-  let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
-  let ?xM = "{y. y\<in> S \<and> x \<le> y}"
-  let ?a = "Max ?Mx"
-  let ?b = "Min ?xM"
-  have MxS: "?Mx \<subseteq> S" by blast
-  hence fMx: "finite ?Mx" using fS finite_subset by auto
-  from lx linS have linMx: "l \<in> ?Mx" by blast
-  hence Mxne: "?Mx \<noteq> {}" by blast
-  have xMS: "?xM \<subseteq> S" by blast
-  hence fxM: "finite ?xM" using fS finite_subset by auto
-  from xu uinS have linxM: "u \<in> ?xM" by blast
-  hence xMne: "?xM \<noteq> {}" by blast
-  have ax:"?a \<le> x" using Mxne fMx by auto
-  have xb:"x \<le> ?b" using xMne fxM by auto
-  have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
-  have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
-  have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
-  proof(clarsimp)
-    fix y   assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
-    from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
-    moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
-    moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
-    ultimately show "False" by blast
-  qed
-  from ainS binS noy ax xb px show ?thesis by blast
-qed
-
-lemma finite_set_intervals2:
-  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
-  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
-  shows "(\<exists> s\<in> S. P s) \<or> (\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a < x \<and> x < b \<and> P x)"
-proof-
-  from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
-  obtain a and b where
-    as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
-    and axb: "a \<le> x \<and> x \<le> b \<and> P x"  by auto
-  from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
-  thus ?thesis using px as bs noS by blast
-qed
-
-end
-
-section {* The classical QE after Langford for dense linear orders *}
-
-context dense_linear_order
-begin
-
-lemma interval_empty_iff:
-  "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
-  by (auto dest: dense)
-
-lemma dlo_qe_bnds: 
-  assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite U"
-  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> (\<forall> l \<in> L. \<forall>u \<in> U. l < u)"
-proof (simp only: atomize_eq, rule iffI)
-  assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)"
-  then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
-  {fix l u assume l: "l \<in> L" and u: "u \<in> U"
-    have "l < x" using xL l by blast
-    also have "x < u" using xU u by blast
-    finally (less_trans) have "l < u" .}
-  thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
-next
-  assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u"
-  let ?ML = "Max L"
-  let ?MU = "Min U"  
-  from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" by auto
-  from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" by auto
-  from th1 th2 H have "?ML < ?MU" by auto
-  with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast
-  from th3 th1' have "\<forall>l \<in> L. l < w" by auto
-  moreover from th4 th2' have "\<forall>u \<in> U. w < u" by auto
-  ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" by auto
-qed
-
-lemma dlo_qe_noub: 
-  assumes ne: "L \<noteq> {}" and fL: "finite L"
-  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
-proof(simp add: atomize_eq)
-  from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast
-  from ne fL have "\<forall>x \<in> L. x \<le> Max L" by simp
-  with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans)
-  thus "\<exists>x. \<forall>y\<in>L. y < x" by blast
-qed
-
-lemma dlo_qe_nolb: 
-  assumes ne: "U \<noteq> {}" and fU: "finite U"
-  shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
-proof(simp add: atomize_eq)
-  from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast
-  from ne fU have "\<forall>x \<in> U. Min U \<le> x" by simp
-  with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans)
-  thus "\<exists>x. \<forall>y\<in>U. x < y" by blast
-qed
-
-lemma exists_neq: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x" 
-  using gt_ex[of t] by auto
-
-lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq 
-  le_less neq_iff linear less_not_permute
-
-lemma axiom: "dense_linear_order (op \<le>) (op <)" by (rule dense_linear_order_axioms)
-lemma atoms:
-  shows "TERM (less :: 'a \<Rightarrow> _)"
-    and "TERM (less_eq :: 'a \<Rightarrow> _)"
-    and "TERM (op = :: 'a \<Rightarrow> _)" .
-
-declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms]
-declare dlo_simps[langfordsimp]
-
-end
-
-(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
-lemma dnf:
-  "(P & (Q | R)) = ((P&Q) | (P&R))" 
-  "((Q | R) & P) = ((Q&P) | (R&P))"
-  by blast+
-
-lemmas weak_dnf_simps = simp_thms dnf
-
-lemma nnf_simps:
-    "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
-    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
-  by blast+
-
-lemma ex_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" by blast
-
-lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib
-
-use "~~/src/HOL/Tools/Qelim/langford.ML"
-method_setup dlo = {*
-  Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac)
-*} "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"} *}
-
-text {* Linear order without upper bounds *}
-
-locale linorder_stupid_syntax = linorder
-begin
-notation
-  less_eq  ("op \<sqsubseteq>") and
-  less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
-  less  ("op \<sqsubset>") and
-  less  ("(_/ \<sqsubset> _)"  [51, 51] 50)
-
-end
-
-locale linorder_no_ub = linorder_stupid_syntax +
-  assumes gt_ex: "\<exists>y. less x y"
-begin
-lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
-
-text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
-lemma pinf_conj:
-  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
-  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
-proof-
-  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
-  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
-  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
-  {fix x assume H: "z \<sqsubset> x"
-    from less_trans[OF zz1 H] less_trans[OF zz2 H]
-    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
-  }
-  thus ?thesis by blast
-qed
-
-lemma pinf_disj:
-  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
-  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
-proof-
-  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
-  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
-  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
-  {fix x assume H: "z \<sqsubset> x"
-    from less_trans[OF zz1 H] less_trans[OF zz2 H]
-    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
-  }
-  thus ?thesis by blast
-qed
-
-lemma pinf_ex: assumes ex:"\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
-proof-
-  from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
-  from gt_ex obtain x where x: "z \<sqsubset> x" by blast
-  from z x p1 show ?thesis by blast
-qed
-
-end
-
-text {* Linear order without upper bounds *}
-
-locale linorder_no_lb = linorder_stupid_syntax +
-  assumes lt_ex: "\<exists>y. less y x"
-begin
-lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
-
-
-text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
-lemma minf_conj:
-  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
-  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
-proof-
-  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
-  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
-  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
-  {fix x assume H: "x \<sqsubset> z"
-    from less_trans[OF H zz1] less_trans[OF H zz2]
-    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
-  }
-  thus ?thesis by blast
-qed
-
-lemma minf_disj:
-  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
-  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
-proof-
-  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
-  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
-  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
-  {fix x assume H: "x \<sqsubset> z"
-    from less_trans[OF H zz1] less_trans[OF H zz2]
-    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
-  }
-  thus ?thesis by blast
-qed
-
-lemma minf_ex: assumes ex:"\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
-proof-
-  from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
-  from lt_ex obtain x where x: "x \<sqsubset> z" by blast
-  from z x p1 show ?thesis by blast
-qed
-
-end
-
-
-locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
-  fixes between
-  assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
-     and  between_same: "between x x = x"
-
-interpretation  constr_dense_linear_order < dense_linear_order 
-  apply unfold_locales
-  using gt_ex lt_ex between_less
-    by (auto, rule_tac x="between x y" in exI, simp)
-
-context  constr_dense_linear_order
-begin
-
-lemma rinf_U:
-  assumes fU: "finite U"
-  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
-  \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
-  and nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')"
-  and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  P x"
-  shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
-proof-
-  from ex obtain x where px: "P x" by blast
-  from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto
-  then obtain u and u' where uU:"u\<in> U" and uU': "u' \<in> U" and ux:"u \<sqsubseteq> x" and xu':"x \<sqsubseteq> u'" by auto
-  from uU have Une: "U \<noteq> {}" by auto
-  term "linorder.Min less_eq"
-  let ?l = "linorder.Min less_eq U"
-  let ?u = "linorder.Max less_eq U"
-  have linM: "?l \<in> U" using fU Une by simp
-  have uinM: "?u \<in> U" using fU Une by simp
-  have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU by auto
-  have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU by auto
-  have th:"?l \<sqsubseteq> u" using uU Une lM by auto
-  from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" .
-  have th: "u' \<sqsubseteq> ?u" using uU' Une Mu by simp
-  from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" .
-  from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
-  have "(\<exists> s\<in> U. P s) \<or>
-      (\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x)" .
-  moreover { fix u assume um: "u\<in>U" and pu: "P u"
-    have "between u u = u" by (simp add: between_same)
-    with um pu have "P (between u u)" by simp
-    with um have ?thesis by blast}
-  moreover{
-    assume "\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x"
-      then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
-        and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U" and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x"
-        by blast
-      from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
-      let ?u = "between t1 t2"
-      from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
-      from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
-      with t1M t2M have ?thesis by blast}
-    ultimately show ?thesis by blast
-  qed
-
-theorem fr_eq:
-  assumes fU: "finite U"
-  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
-   \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
-  and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
-  and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
-  and mi: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x = MP)"  and pi: "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x = PP)"
-  shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))"
-  (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
-proof-
- {
-   assume px: "\<exists> x. P x"
-   have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
-   moreover {assume "MP \<or> PP" hence "?D" by blast}
-   moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
-     from npmibnd[OF nmibnd npibnd]
-     have nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')" .
-     from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
-   ultimately have "?D" by blast}
- moreover
- { assume "?D"
-   moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
-   moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
-   moreover {assume f:"?F" hence "?E" by blast}
-   ultimately have "?E" by blast}
- ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
-qed
-
-lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
-lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P
-
-lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P
-lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P
-lemmas lin_dense_thms = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P
-
-lemma ferrack_axiom: "constr_dense_linear_order less_eq less between"
-  by (rule constr_dense_linear_order_axioms)
-lemma atoms:
-  shows "TERM (less :: 'a \<Rightarrow> _)"
-    and "TERM (less_eq :: 'a \<Rightarrow> _)"
-    and "TERM (op = :: 'a \<Rightarrow> _)" .
-
-declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms
-    nmi: nmi_thms npi: npi_thms lindense:
-    lin_dense_thms qe: fr_eq atoms: atoms]
-
-declaration {*
-let
-fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
-fun generic_whatis phi =
- let
-  val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
-  fun h x t =
-   case term_of t of
-     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
-                            else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
-                            else Ferrante_Rackoff_Data.Nox
-   | b$y$z => if Term.could_unify (b, lt) then
-                 if term_of x aconv y then Ferrante_Rackoff_Data.Lt
-                 else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
-                 else Ferrante_Rackoff_Data.Nox
-             else if Term.could_unify (b, le) then
-                 if term_of x aconv y then Ferrante_Rackoff_Data.Le
-                 else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
-                 else Ferrante_Rackoff_Data.Nox
-             else Ferrante_Rackoff_Data.Nox
-   | _ => Ferrante_Rackoff_Data.Nox
- in h end
- fun ss phi = HOL_ss addsimps (simps phi)
-in
- Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
-  {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
-end
-*}
-
-end
-
-use "~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML"
-
-method_setup ferrack = {*
-  Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
-*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
-
-subsection {* Ferrante and Rackoff algorithm over ordered fields *}
-
-lemma neg_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
-proof-
-  assume H: "c < 0"
-  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
-  also have "\<dots> = (0 < x)" by simp
-  finally show  "(c*x < 0) == (x > 0)" by simp
-qed
-
-lemma pos_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
-proof-
-  assume H: "c > 0"
-  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
-  also have "\<dots> = (0 > x)" by simp
-  finally show  "(c*x < 0) == (x < 0)" by simp
-qed
-
-lemma neg_prod_sum_lt: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t< 0) == (x > (- 1/c)*t))"
-proof-
-  assume H: "c < 0"
-  have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
-  also have "\<dots> = ((- 1/c)*t < x)" by simp
-  finally show  "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
-qed
-
-lemma pos_prod_sum_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t < 0) == (x < (- 1/c)*t))"
-proof-
-  assume H: "c > 0"
-  have "c*x + t< 0 = (c*x < -t)"  by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
-  also have "\<dots> = ((- 1/c)*t > x)" by simp
-  finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" by simp
-qed
-
-lemma sum_lt:"((x::'a::pordered_ab_group_add) + t < 0) == (x < - t)"
-  using less_diff_eq[where a= x and b=t and c=0] by simp
-
-lemma neg_prod_le:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
-proof-
-  assume H: "c < 0"
-  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
-  also have "\<dots> = (0 <= x)" by simp
-  finally show  "(c*x <= 0) == (x >= 0)" by simp
-qed
-
-lemma pos_prod_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
-proof-
-  assume H: "c > 0"
-  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
-  also have "\<dots> = (0 >= x)" by simp
-  finally show  "(c*x <= 0) == (x <= 0)" by simp
-qed
-
-lemma neg_prod_sum_le: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t <= 0) == (x >= (- 1/c)*t))"
-proof-
-  assume H: "c < 0"
-  have "c*x + t <= 0 = (c*x <= -t)"  by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
-  also have "\<dots> = ((- 1/c)*t <= x)" by simp
-  finally show  "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
-qed
-
-lemma pos_prod_sum_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t <= 0) == (x <= (- 1/c)*t))"
-proof-
-  assume H: "c > 0"
-  have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
-  also have "\<dots> = ((- 1/c)*t >= x)" by simp
-  finally show  "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
-qed
-
-lemma sum_le:"((x::'a::pordered_ab_group_add) + t <= 0) == (x <= - t)"
-  using le_diff_eq[where a= x and b=t and c=0] by simp
-
-lemma nz_prod_eq:"(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x = 0) == (x = 0))" by simp
-lemma nz_prod_sum_eq: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x + t = 0) == (x = (- 1/c)*t))"
-proof-
-  assume H: "c \<noteq> 0"
-  have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] ring_simps)
-  finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" by simp
-qed
-lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)"
-  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
- ["op <=" "op <"
-   "\<lambda> x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"]
-proof (unfold_locales, dlo, dlo, auto)
-  fix x y::'a assume lt: "x < y"
-  from  less_half_sum[OF lt] show "x < (x + y) /2" by simp
-next
-  fix x y::'a assume lt: "x < y"
-  from  gt_half_sum[OF lt] show "(x + y) /2 < y" by simp
-qed
-
-declaration{*
-let
-fun earlier [] x y = false
-        | earlier (h::t) x y =
-    if h aconvc y then false else if h aconvc x then true else earlier t x y;
-
-fun dest_frac ct = case term_of ct of
-   Const (@{const_name "HOL.divide"},_) $ a $ b=>
-    Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
- | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
-
-fun mk_frac phi cT x =
- let val (a, b) = Rat.quotient_of_rat x
- in if b = 1 then Numeral.mk_cnumber cT a
-    else Thm.capply
-         (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
-                     (Numeral.mk_cnumber cT a))
-         (Numeral.mk_cnumber cT b)
- end
-
-fun whatis x ct = case term_of ct of
-  Const(@{const_name "HOL.plus"}, _)$(Const(@{const_name "HOL.times"},_)$_$y)$_ =>
-     if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
-     else ("Nox",[])
-| Const(@{const_name "HOL.plus"}, _)$y$_ =>
-     if y aconv term_of x then ("x+t",[Thm.dest_arg ct])
-     else ("Nox",[])
-| Const(@{const_name "HOL.times"}, _)$_$y =>
-     if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct])
-     else ("Nox",[])
-| t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
-
-fun xnormalize_conv ctxt [] ct = reflexive ct
-| xnormalize_conv ctxt (vs as (x::_)) ct =
-   case term_of ct of
-   Const(@{const_name HOL.less},_)$_$Const(@{const_name "HOL.zero"},_) =>
-    (case whatis x (Thm.dest_arg1 ct) of
-    ("c*x+t",[c,t]) =>
-       let
-        val cr = dest_frac c
-        val clt = Thm.dest_fun2 ct
-        val cz = Thm.dest_arg ct
-        val neg = cr </ Rat.zero
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-               (Thm.capply @{cterm "Trueprop"}
-                  (if neg then Thm.capply (Thm.capply clt c) cz
-                    else Thm.capply (Thm.capply clt cz) c))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
-             (if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-      in rth end
-    | ("x+t",[t]) =>
-       let
-        val T = ctyp_of_term x
-        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-       in  rth end
-    | ("c*x",[c]) =>
-       let
-        val cr = dest_frac c
-        val clt = Thm.dest_fun2 ct
-        val cz = Thm.dest_arg ct
-        val neg = cr </ Rat.zero
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-               (Thm.capply @{cterm "Trueprop"}
-                  (if neg then Thm.capply (Thm.capply clt c) cz
-                    else Thm.capply (Thm.capply clt cz) c))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
-             (if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
-        val rth = th
-      in rth end
-    | _ => reflexive ct)
-
-
-|  Const(@{const_name HOL.less_eq},_)$_$Const(@{const_name "HOL.zero"},_) =>
-   (case whatis x (Thm.dest_arg1 ct) of
-    ("c*x+t",[c,t]) =>
-       let
-        val T = ctyp_of_term x
-        val cr = dest_frac c
-        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
-        val cz = Thm.dest_arg ct
-        val neg = cr </ Rat.zero
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-               (Thm.capply @{cterm "Trueprop"}
-                  (if neg then Thm.capply (Thm.capply clt c) cz
-                    else Thm.capply (Thm.capply clt cz) c))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim (instantiate' [SOME T] (map SOME [c,x,t])
-             (if neg then @{thm neg_prod_sum_le} else @{thm pos_prod_sum_le})) cth
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-      in rth end
-    | ("x+t",[t]) =>
-       let
-        val T = ctyp_of_term x
-        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"}
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-       in  rth end
-    | ("c*x",[c]) =>
-       let
-        val T = ctyp_of_term x
-        val cr = dest_frac c
-        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
-        val cz = Thm.dest_arg ct
-        val neg = cr </ Rat.zero
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-               (Thm.capply @{cterm "Trueprop"}
-                  (if neg then Thm.capply (Thm.capply clt c) cz
-                    else Thm.capply (Thm.capply clt cz) c))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
-             (if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
-        val rth = th
-      in rth end
-    | _ => reflexive ct)
-
-|  Const("op =",_)$_$Const(@{const_name "HOL.zero"},_) =>
-   (case whatis x (Thm.dest_arg1 ct) of
-    ("c*x+t",[c,t]) =>
-       let
-        val T = ctyp_of_term x
-        val cr = dest_frac c
-        val ceq = Thm.dest_fun2 ct
-        val cz = Thm.dest_arg ct
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-            (Thm.capply @{cterm "Trueprop"}
-             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim
-                 (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-      in rth end
-    | ("x+t",[t]) =>
-       let
-        val T = ctyp_of_term x
-        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-       in  rth end
-    | ("c*x",[c]) =>
-       let
-        val T = ctyp_of_term x
-        val cr = dest_frac c
-        val ceq = Thm.dest_fun2 ct
-        val cz = Thm.dest_arg ct
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-            (Thm.capply @{cterm "Trueprop"}
-             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val rth = implies_elim
-                 (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth
-      in rth end
-    | _ => reflexive ct);
-
-local
-  val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"}
-  val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"}
-  val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
-in
-fun field_isolate_conv phi ctxt vs ct = case term_of ct of
-  Const(@{const_name HOL.less},_)$a$b =>
-   let val (ca,cb) = Thm.dest_binop ct
-       val T = ctyp_of_term ca
-       val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0
-       val nth = Conv.fconv_rule
-         (Conv.arg_conv (Conv.arg1_conv
-              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
-       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
-   in rth end
-| Const(@{const_name HOL.less_eq},_)$a$b =>
-   let val (ca,cb) = Thm.dest_binop ct
-       val T = ctyp_of_term ca
-       val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0
-       val nth = Conv.fconv_rule
-         (Conv.arg_conv (Conv.arg1_conv
-              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
-       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
-   in rth end
-
-| Const("op =",_)$a$b =>
-   let val (ca,cb) = Thm.dest_binop ct
-       val T = ctyp_of_term ca
-       val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
-       val nth = Conv.fconv_rule
-         (Conv.arg_conv (Conv.arg1_conv
-              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
-       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
-   in rth end
-| @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
-| _ => reflexive ct
-end;
-
-fun classfield_whatis phi =
- let
-  fun h x t =
-   case term_of t of
-     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
-                            else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
-                            else Ferrante_Rackoff_Data.Nox
-   | Const(@{const_name HOL.less},_)$y$z =>
-       if term_of x aconv y then Ferrante_Rackoff_Data.Lt
-        else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
-        else Ferrante_Rackoff_Data.Nox
-   | Const (@{const_name HOL.less_eq},_)$y$z =>
-         if term_of x aconv y then Ferrante_Rackoff_Data.Le
-         else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
-         else Ferrante_Rackoff_Data.Nox
-   | _ => Ferrante_Rackoff_Data.Nox
- in h end;
-fun class_field_ss phi =
-   HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}])
-   addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}]
-
-in
-Ferrante_Rackoff_Data.funs @{thm "class_ordered_field_dense_linear_order.ferrack_axiom"}
-  {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
-end
-*}
-
-
-end 
--- a/src/HOL/Library/Efficient_Nat.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Library/Efficient_Nat.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/Efficient_Nat.thy
-    ID:         $Id$
     Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
 *)
 
@@ -127,7 +126,7 @@
 fun remove_suc thy thms =
   let
     val vname = Name.variant (map fst
-      (fold (Term.add_varnames o Thm.full_prop_of) thms [])) "x";
+      (fold (Term.add_var_names o Thm.full_prop_of) thms [])) "x";
     val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT));
     fun lhs_of th = snd (Thm.dest_comb
       (fst (Thm.dest_comb (snd (Thm.dest_comb (cprop_of th))))));
@@ -170,7 +169,7 @@
 fun eqn_suc_preproc thy ths =
   let
     val dest = fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of;
-    fun contains_suc t = member (op =) (term_consts t) @{const_name Suc};
+    val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc});
   in
     if forall (can dest) ths andalso
       exists (contains_suc o dest) ths
@@ -180,7 +179,7 @@
 fun remove_suc_clause thy thms =
   let
     val vname = Name.variant (map fst
-      (fold (Term.add_varnames o Thm.full_prop_of) thms [])) "x";
+      (fold (Term.add_var_names o Thm.full_prop_of) thms [])) "x";
     fun find_var (t as Const (@{const_name Suc}, _) $ (v as Var _)) = SOME (t, v)
       | find_var (t $ u) = (case find_var t of NONE => find_var u | x => x)
       | find_var _ = NONE;
@@ -212,8 +211,8 @@
     val dest = fst o HOLogic.dest_mem o HOLogic.dest_Trueprop
   in
     if forall (can (dest o concl_of)) ths andalso
-      exists (fn th => member (op =) (foldr add_term_consts
-        [] (map_filter (try dest) (concl_of th :: prems_of th))) "Suc") ths
+      exists (fn th => exists (exists_Const (fn (c, _) => c = @{const_name Suc}))
+        (map_filter (try dest) (concl_of th :: prems_of th))) ths
     then remove_suc_clause thy ths else ths
   end;
 
--- a/src/HOL/Library/Library.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Library/Library.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -16,7 +16,6 @@
   Continuity
   ContNotDenum
   Countable
-  Dense_Linear_Order
   Efficient_Nat
   Enum
   Eval_Witness
--- a/src/HOL/Library/Multiset.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Library/Multiset.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1080,15 +1080,15 @@
 apply simp
 done
 
-interpretation mset_order: order ["op \<le>#" "op <#"]
+class_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:
+class_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:
+class_interpretation mset_order_semigroup_cancel:
   pordered_ab_semigroup_add_imp_le ["op +" "op \<le>#" "op <#"]
 proof qed simp
 
@@ -1404,7 +1404,7 @@
   assumes "left_commutative g"
   shows "(\<And>x y. h (g x y) = f x (h y)) \<Longrightarrow> h (fold_mset g w A) = fold_mset f (h w) A" (is "PROP ?P")
 proof -
-  interpret left_commutative [g] by fact
+  interpret left_commutative g by fact
   show "PROP ?P" by (induct A) auto
 qed
 
@@ -1436,7 +1436,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 {#} = {#}"
--- a/src/HOL/Library/Nat_Infinity.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Library/Nat_Infinity.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/Nat_Infinity.thy
-    ID:         $Id$
     Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
 *)
 
@@ -9,6 +8,17 @@
 imports Plain "~~/src/HOL/Presburger"
 begin
 
+text {* FIXME: move to Nat.thy *}
+
+instantiation nat :: bot
+begin
+
+definition bot_nat :: nat where
+  "bot_nat = 0"
+
+instance proof
+qed (simp add: bot_nat_def)
+
 subsection {* Type definition *}
 
 text {*
@@ -16,6 +26,8 @@
   infinity.
 *}
 
+end
+
 datatype inat = Fin nat | Infty
 
 notation (xsymbols)
@@ -353,6 +365,20 @@
 apply (erule (1) le_less_trans)
 done
 
+instantiation inat :: "{bot, top}"
+begin
+
+definition bot_inat :: inat where
+  "bot_inat = 0"
+
+definition top_inat :: inat where
+  "top_inat = \<infinity>"
+
+instance proof
+qed (simp_all add: bot_inat_def top_inat_def)
+
+end
+
 
 subsection {* Well-ordering *}
 
--- a/src/HOL/Library/SetsAndFunctions.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Library/SetsAndFunctions.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/SetsAndFunctions.thy
-    ID:         $Id$
     Author:     Jeremy Avigad and Kevin Donnelly
 *)
 
@@ -108,26 +107,26 @@
   apply simp
   done
 
-interpretation set_semigroup_add: semigroup_add ["op \<oplus> :: ('a::semigroup_add) set => 'a set => 'a set"]
+class_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"]
+class_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"]
+class_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"]
+class_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/comm_ring.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Library/comm_ring.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
-(*  ID:         $Id$
-    Author:     Amine Chaieb
+(*  Author:     Amine Chaieb
 
 Tactic for solving equalities over commutative rings.
 *)
@@ -71,7 +70,7 @@
 fun reif_eq thy (eq as Const("op =", Type("fun", [T, _])) $ lhs $ rhs) =
       if Sign.of_sort thy (T, cr_sort) then
         let
-          val fs = term_frees eq;
+          val fs = OldTerm.term_frees eq;
           val cvs = cterm_of thy (HOLogic.mk_list T fs);
           val clhs = cterm_of thy (reif_polex T fs lhs);
           val crhs = cterm_of thy (reif_polex T fs rhs);
--- a/src/HOL/Lim.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Lim.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -7,7 +7,7 @@
 header{* Limits and Continuity *}
 
 theory Lim
-imports "~~/src/HOL/Hyperreal/SEQ"
+imports SEQ
 begin
 
 text{*Standard Definitions*}
--- a/src/HOL/List.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/List.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/List.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
 *)
 
@@ -359,7 +358,7 @@
 
    fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
     let
-      val x = Free (Name.variant (add_term_free_names (p$e, [])) "x", dummyT);
+      val x = Free (Name.variant (fold Term.add_free_names [p, e] []) "x", dummyT);
       val e = if opti then singl e else e;
       val case1 = Syntax.const "_case1" $ p $ e;
       val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN
@@ -548,9 +547,9 @@
 lemma append_Nil2 [simp]: "xs @ [] = xs"
 by (induct xs) auto
 
-interpretation semigroup_append: semigroup_add ["op @"]
+class_interpretation semigroup_append: semigroup_add ["op @"]
   proof qed simp
-interpretation monoid_append: monoid_add ["[]" "op @"]
+class_interpretation monoid_append: monoid_add ["[]" "op @"]
   proof qed simp+
 
 lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])"
--- a/src/HOL/Main.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Main.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,13 +1,14 @@
-(*  Title:      HOL/Main.thy
-    ID:         $Id$
-*)
-
 header {* Main HOL *}
 
 theory Main
 imports Plain Code_Eval Map Nat_Int_Bij Recdef
 begin
 
+text {*
+  Classical Higher-order Logic -- only ``Main'', excluding real and
+  complex numbers etc.
+*}
+
 text {* See further \cite{Nipkow-et-al:2002:tutorial} *}
 
 end
--- a/src/HOL/Matrix/cplex/matrixlp.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Matrix/cplex/matrixlp.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Matrix/cplex/matrixlp.ML
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
@@ -20,7 +19,7 @@
 fun inst_real thm =
   let val certT = ctyp_of (Thm.theory_of_thm thm) in
     standard (Thm.instantiate
-      ([(certT (TVar (hd (term_tvars (prop_of thm)))), certT HOLogic.realT)], []) thm)
+      ([(certT (TVar (hd (OldTerm.term_tvars (prop_of thm)))), certT HOLogic.realT)], []) thm)
   end
 
 local
@@ -58,7 +57,7 @@
     let
         val certT = Thm.ctyp_of (Thm.theory_of_thm thm);
         val ord = prod_ord (prod_ord string_ord int_ord) (list_ord string_ord)
-        val v = TVar (hd (sort ord (term_tvars (prop_of thm))))
+        val v = TVar (hd (sort ord (OldTerm.term_tvars (prop_of thm))))
     in
         standard (Thm.instantiate ([(certT v, certT ty)], []) thm)
     end
--- a/src/HOL/MicroJava/BV/Kildall.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/MicroJava/BV/Kildall.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -321,7 +321,7 @@
   ss <[r] merges f qs ss \<or> 
   merges f qs ss = ss \<and> {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un (w-{p}) < w" (is "PROP ?P")
 proof -
-  interpret Semilat [A r f] using assms by (rule Semilat.intro)
+  interpret Semilat A r f using assms by (rule Semilat.intro)
   show "PROP ?P" apply(insert semilat)
     apply (unfold lesssub_def)
     apply (simp (no_asm_simp) add: merges_incr)
@@ -351,7 +351,7 @@
    (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss' <=[r] ts)"
   (is "PROP ?P")
 proof -
-  interpret Semilat [A r f] using assms by (rule Semilat.intro)
+  interpret Semilat A r f using assms by (rule Semilat.intro)
   show "PROP ?P" apply(insert semilat)
 apply (unfold iter_def stables_def)
 apply (rule_tac P = "%(ss,w).
@@ -457,7 +457,7 @@
                  kildall r f step ss0 <=[r] ts)"
   (is "PROP ?P")
 proof -
-  interpret Semilat [A r f] using assms by (rule Semilat.intro)
+  interpret Semilat A r f using assms by (rule Semilat.intro)
   show "PROP ?P"
 apply (unfold kildall_def)
 apply(case_tac "iter f step ss0 (unstables r step ss0)")
@@ -474,7 +474,7 @@
   \<Longrightarrow> is_bcv r T step n A (kildall r f step)"
   (is "PROP ?P")
 proof -
-  interpret Semilat [A r f] using assms by (rule Semilat.intro)
+  interpret Semilat A r f using assms by (rule Semilat.intro)
   show "PROP ?P"
 apply(unfold is_bcv_def wt_step_def)
 apply(insert semilat kildall_properties[of A])
--- a/src/HOL/MicroJava/BV/LBVComplete.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/MicroJava/BV/LBVComplete.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -197,7 +197,7 @@
   have "merge c pc ?step (c!Suc pc) =
     (if \<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'
     then map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc
-    else \<top>)" unfolding merge_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms])
+    else \<top>)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms])
   moreover {
     fix pc' s' assume s': "(pc', s') \<in> set ?step" and suc_pc: "pc' \<noteq> pc+1"
     with less have "s' <=_r \<phi>!pc'" by auto
--- a/src/HOL/MicroJava/BV/Listn.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/MicroJava/BV/Listn.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -380,7 +380,7 @@
 lemma Listn_sl_aux:
 assumes "semilat (A, r, f)" shows "semilat (Listn.sl n (A,r,f))"
 proof -
-  interpret Semilat [A r f] using assms by (rule Semilat.intro)
+  interpret Semilat A r f using assms by (rule Semilat.intro)
 show ?thesis
 apply (unfold Listn.sl_def)
 apply (simp (no_asm) only: semilat_Def split_conv)
--- a/src/HOL/MicroJava/BV/SemilatAlg.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/MicroJava/BV/SemilatAlg.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -67,7 +67,7 @@
 lemma plusplus_closed: assumes "semilat (A, r, f)" shows
   "\<And>y. \<lbrakk> set x \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> x ++_f y \<in> A" (is "PROP ?P")
 proof -
-  interpret Semilat [A r f] using assms by (rule Semilat.intro)
+  interpret Semilat A r f using assms by (rule Semilat.intro)
   show "PROP ?P" proof (induct x)
     show "\<And>y. y \<in> A \<Longrightarrow> [] ++_f y \<in> A" by simp
     fix y x xs
@@ -164,7 +164,7 @@
   shows "\<lbrakk>\<forall>(p,s) \<in> set S. s \<in> A; y \<in> A; (a,b) \<in> set S\<rbrakk> 
   \<Longrightarrow> b <=_r map snd [(p', t')\<leftarrow>S. p' = a] ++_f y" 
 proof -
-  interpret Semilat [A r f] using assms by (rule Semilat.intro)
+  interpret Semilat A r f using assms by (rule Semilat.intro)
 
   let "b <=_r ?map ++_f y" = ?thesis
 
--- a/src/HOL/Modelcheck/mucke_oracle.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Modelcheck/mucke_oracle.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,3 @@
-
-(* $Id$ *)
 
 val trace_mc = ref false; 
 
@@ -249,8 +247,8 @@
 (* replace Vars bei Frees, freeze_thaw shares code of tactic/freeze_thaw
    and thm.instantiate *)
 fun freeze_thaw t =
-  let val used = add_term_names (t, [])
-          and vars = term_vars t;
+  let val used = OldTerm.add_term_names (t, [])
+          and vars = OldTerm.term_vars t;
       fun newName (Var(ix,_), (pairs,used)) = 
           let val v = Name.variant used (string_of_indexname ix)
           in  ((ix,v)::pairs, v::used)  end;
--- a/src/HOL/NSA/StarDef.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/NSA/StarDef.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -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/Nominal/Examples/SOS.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Nominal/Examples/SOS.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -328,7 +328,7 @@
   have ih2:"\<And>t. e\<^isub>2 \<Down> t \<Longrightarrow> e\<^isub>2' = t" by fact
   have ih3: "\<And>t. e\<^isub>1'[x::=e\<^isub>2'] \<Down> t \<Longrightarrow> e' = t" by fact
   have app: "App e\<^isub>1 e\<^isub>2 \<Down> t\<^isub>2" by fact
-  have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" "x\<sharp>t\<^isub>2" by fact
+  have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" "x\<sharp>t\<^isub>2" by fact+
   then have "x\<sharp>App e\<^isub>1 e\<^isub>2" by auto
   from app vc obtain f\<^isub>1 f\<^isub>2 where x1: "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" and x2: "e\<^isub>2 \<Down> f\<^isub>2" and x3: "f\<^isub>1[x::=f\<^isub>2] \<Down> t\<^isub>2"
     by (auto elim!: b_App_elim)
@@ -515,7 +515,7 @@
   have ih:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T" by fact
   have as\<^isub>1: "\<theta> Vcloses \<Gamma>" by fact
   have as\<^isub>2: "\<Gamma> \<turnstile> Lam [x].e : T" by fact
-  have fs: "x\<sharp>\<Gamma>" "x\<sharp>\<theta>" by fact
+  have fs: "x\<sharp>\<Gamma>" "x\<sharp>\<theta>" by fact+
   from as\<^isub>2 fs obtain T\<^isub>1 T\<^isub>2 
     where "(i)": "(x,T\<^isub>1)#\<Gamma> \<turnstile> e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 \<rightarrow> T\<^isub>2" using fs
     by (auto elim: t_Lam_elim)
--- a/src/HOL/Nominal/nominal_fresh_fun.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,10 +1,8 @@
 (*  Title:      HOL/Nominal/nominal_fresh_fun.ML
-    ID:         $Id$
-    Authors:    Stefan Berghofer, Julien Narboux, TU Muenchen
+    Authors:    Stefan Berghofer and Julien Narboux, TU Muenchen
 
 Provides a tactic to generate fresh names and
 a tactic to analyse instances of the fresh_fun.
-
 *)
 
 (* First some functions that should be in the library *)
@@ -83,7 +81,7 @@
    val bvs = map_index (Bound o fst) ps;
 (* select variables of the right class *)
    val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t)))
-     (term_frees goal @ bvs);
+     (OldTerm.term_frees goal @ bvs);
 (* build the tuple *)
    val s = (Library.foldr1 (fn (v, s) =>
      HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s) vs) handle _ => HOLogic.unit ;  (* FIXME avoid handle _ *)
@@ -91,7 +89,7 @@
    val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
    val exists_fresh' = at_name_inst_thm RS at_exists_fresh';
 (* find the variable we want to instantiate *)
-   val x = hd (term_vars (prop_of exists_fresh'));
+   val x = hd (OldTerm.term_vars (prop_of exists_fresh'));
  in 
    (cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN
    rtac fs_name_thm 1 THEN
@@ -150,7 +148,7 @@
     val fresh_perm_app = PureThy.get_thms thy "fresh_perm_app";
     val ss' = ss addsimps fresh_prod::abs_fresh;
     val ss'' = ss' addsimps fresh_perm_app;
-    val x = hd (tl (term_vars (prop_of exI)));
+    val x = hd (tl (OldTerm.term_vars (prop_of exI)));
     val goal = nth (prems_of thm) (i-1);
     val atom_name_opt = get_inner_fresh_fun goal;
     val n = List.length (Logic.strip_params goal);
@@ -165,7 +163,7 @@
     val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename;
     val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
     fun inst_fresh vars params i st =
-   let val vars' = term_vars (prop_of st);
+   let val vars' = OldTerm.term_vars (prop_of st);
        val thy = theory_of_thm st;
    in case vars' \\ vars of 
      [x] => Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (list_abs (params,Bound 0)))]) st)
@@ -174,7 +172,7 @@
   in
   ((fn st =>
   let 
-    val vars = term_vars (prop_of st);
+    val vars = OldTerm.term_vars (prop_of st);
     val params = Logic.strip_params (nth (prems_of st) (i-1))
     (* The tactics which solve the subgoals generated 
        by the conditionnal rewrite rule. *)
--- a/src/HOL/Nominal/nominal_inductive.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Nominal/nominal_inductive.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Nominal/nominal_inductive.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Infrastructure for proving equivariance and strong induction theorems
@@ -153,7 +152,7 @@
     val elims = map (atomize_induct ctxt) elims;
     val monos = InductivePackage.get_monos ctxt;
     val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
-    val _ = (case names \\ foldl (apfst prop_of #> add_term_consts) [] eqvt_thms of
+    val _ = (case names \\ fold (Term.add_const_names o Thm.prop_of) eqvt_thms [] of
         [] => ()
       | xs => error ("Missing equivariance theorem for predicate(s): " ^
           commas_quote xs));
@@ -200,8 +199,8 @@
     val ind_sort = if null atomTs then HOLogic.typeS
       else Sign.certify_sort thy (map (fn T => Sign.intern_class thy
         ("fs_" ^ Sign.base_name (fst (dest_Type T)))) atomTs);
-    val fs_ctxt_tyname = Name.variant (map fst (term_tfrees raw_induct')) "'n";
-    val fs_ctxt_name = Name.variant (add_term_names (raw_induct', [])) "z";
+    val fs_ctxt_tyname = Name.variant (map fst (OldTerm.term_tfrees raw_induct')) "'n";
+    val fs_ctxt_name = Name.variant (OldTerm.add_term_names (raw_induct', [])) "z";
     val fsT = TFree (fs_ctxt_tyname, ind_sort);
 
     val inductive_forall_def' = Drule.instantiate'
@@ -238,10 +237,10 @@
         val prem = Logic.list_implies
           (map mk_fresh bvars @ mk_distinct bvars @
            map (fn prem =>
-             if null (term_frees prem inter ps) then prem
+             if null (OldTerm.term_frees prem inter ps) then prem
              else lift_prem prem) prems,
            HOLogic.mk_Trueprop (lift_pred p ts));
-        val vs = map (Var o apfst (rpair 0)) (rename_wrt_term prem params')
+        val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params')
       in
         (list_all (params', prem), (rev vs, subst_bounds (vs, prem)))
       end) prems);
@@ -264,7 +263,7 @@
     val vc_compat = map (fn (params, bvars, prems, (p, ts)) =>
       map (fn q => list_all (params, incr_boundvars ~1 (Logic.list_implies
           (List.mapPartial (fn prem =>
-             if null (ps inter term_frees prem) then SOME prem
+             if null (ps inter OldTerm.term_frees prem) then SOME prem
              else map_term (split_conj (K o I) names) prem prem) prems, q))))
         (mk_distinct bvars @
          maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
@@ -353,7 +352,7 @@
                          (rev pis' @ pis) th));
                  val (gprems1, gprems2) = split_list
                    (map (fn (th, t) =>
-                      if null (term_frees t inter ps) then (SOME th, mk_pi th)
+                      if null (OldTerm.term_frees t inter ps) then (SOME th, mk_pi th)
                       else
                         (map_thm ctxt (split_conj (K o I) names)
                            (etac conjunct1 1) monos NONE th,
@@ -412,7 +411,7 @@
       let
         val prem :: prems = Logic.strip_imp_prems rule;
         val concl = Logic.strip_imp_concl rule;
-        val used = add_term_free_names (rule, [])
+        val used = Term.add_free_names rule [];
       in
         (prem,
          List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params),
@@ -614,7 +613,7 @@
       [mk_perm_bool_simproc names,
        NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun];
     val t = Logic.unvarify (concl_of raw_induct);
-    val pi = Name.variant (add_term_names (t, [])) "pi";
+    val pi = Name.variant (OldTerm.add_term_names (t, [])) "pi";
     val ps = map (fst o HOLogic.dest_imp)
       (HOLogic.dest_conj (HOLogic.dest_Trueprop t));
     fun eqvt_tac pi (intr, vs) st =
--- a/src/HOL/Nominal/nominal_inductive2.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Nominal/nominal_inductive2.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Infrastructure for proving equivariance and strong induction theorems
@@ -139,7 +138,7 @@
   in Option.map prove (map_term f prop (the_default prop opt)) end;
 
 fun abs_params params t =
-  let val vs =  map (Var o apfst (rpair 0)) (rename_wrt_term t params)
+  let val vs =  map (Var o apfst (rpair 0)) (Term.rename_wrt_term t params)
   in (list_all (params, t), (rev vs, subst_bounds (vs, t))) end;
 
 fun inst_params thy (vs, p) th cts =
@@ -159,7 +158,7 @@
     val elims = map (atomize_induct ctxt) elims;
     val monos = InductivePackage.get_monos ctxt;
     val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
-    val _ = (case names \\ foldl (apfst prop_of #> add_term_consts) [] eqvt_thms of
+    val _ = (case names \\ fold (Term.add_const_names o Thm.prop_of) eqvt_thms [] of
         [] => ()
       | xs => error ("Missing equivariance theorem for predicate(s): " ^
           commas_quote xs));
@@ -222,8 +221,8 @@
     val ind_sort = if null atomTs then HOLogic.typeS
       else Sign.certify_sort thy (map (fn a => Sign.intern_class thy
         ("fs_" ^ Sign.base_name a)) atoms);
-    val fs_ctxt_tyname = Name.variant (map fst (term_tfrees raw_induct')) "'n";
-    val fs_ctxt_name = Name.variant (add_term_names (raw_induct', [])) "z";
+    val fs_ctxt_tyname = Name.variant (map fst (OldTerm.term_tfrees raw_induct')) "'n";
+    val fs_ctxt_name = Name.variant (OldTerm.add_term_names (raw_induct', [])) "z";
     val fsT = TFree (fs_ctxt_tyname, ind_sort);
 
     val inductive_forall_def' = Drule.instantiate'
@@ -254,7 +253,7 @@
         val prem = Logic.list_implies
           (map mk_fresh sets @
            map (fn prem =>
-             if null (term_frees prem inter ps) then prem
+             if null (OldTerm.term_frees prem inter ps) then prem
              else lift_prem prem) prems,
            HOLogic.mk_Trueprop (lift_pred p ts));
       in abs_params params' prem end) prems);
@@ -277,7 +276,7 @@
     val (vc_compat, vc_compat') = map (fn (params, sets, prems, (p, ts)) =>
       map (fn q => abs_params params (incr_boundvars ~1 (Logic.list_implies
           (List.mapPartial (fn prem =>
-             if null (ps inter term_frees prem) then SOME prem
+             if null (ps inter OldTerm.term_frees prem) then SOME prem
              else map_term (split_conj (K o I) names) prem prem) prems, q))))
         (maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
            (NominalPackage.fresh_star_const U T $ u $ t)) sets)
@@ -364,7 +363,7 @@
                    fold_rev (NominalPackage.mk_perm []) pis t) sets';
                  val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
                  val gprems1 = List.mapPartial (fn (th, t) =>
-                   if null (term_frees t inter ps) then SOME th
+                   if null (OldTerm.term_frees t inter ps) then SOME th
                    else
                      map_thm ctxt' (split_conj (K o I) names)
                        (etac conjunct1 1) monos NONE th)
@@ -406,7 +405,7 @@
                        (fold_rev (mk_perm_bool o cterm_of thy)
                          (pis' @ pis) th));
                  val gprems2 = map (fn (th, t) =>
-                   if null (term_frees t inter ps) then mk_pi th
+                   if null (OldTerm.term_frees t inter ps) then mk_pi th
                    else
                      mk_pi (the (map_thm ctxt (inst_conj_all names ps (rev pis''))
                        (inst_conj_all_tac (length pis'')) monos (SOME t) th)))
--- a/src/HOL/Nominal/nominal_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Nominal/nominal_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Nominal/nominal_package.ML
-    ID:         $Id$
     Author:     Stefan Berghofer and Christian Urban, TU Muenchen
 
 Nominal datatype package for Isabelle/HOL.
@@ -1414,7 +1413,7 @@
 
     val _ = warning "defining recursion combinator ...";
 
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
 
     val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used;
 
--- a/src/HOL/Nominal/nominal_primrec.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Nominal/nominal_primrec.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,6 @@
 (*  Title:      HOL/Nominal/nominal_primrec.ML
-    Author:     Stefan Berghofer, TU Muenchen and Norbert Voelker, FernUni Hagen
+    Author:     Norbert Voelker, FernUni Hagen
+    Author:     Stefan Berghofer, TU Muenchen
 
 Package for defining functions on nominal datatypes by primitive recursion.
 Taken from HOL/Tools/primrec_package.ML
@@ -71,7 +72,7 @@
     else
      (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
       check_vars "extra variables on rhs: "
-        (map dest_Free (term_frees rhs) |> subtract (op =) lfrees
+        (map dest_Free (OldTerm.term_frees rhs) |> subtract (op =) lfrees
           |> filter_out (is_fixed o fst));
       case AList.lookup (op =) rec_fns fname of
         NONE =>
@@ -154,7 +155,7 @@
               val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
               val rargs = map fst recs;
               val subs = map (rpair dummyT o fst)
-                (rev (rename_wrt_term rhs rargs));
+                (rev (Term.rename_wrt_term rhs rargs));
               val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
                 (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss')
                   handle RecError s => primrec_eq_err lthy s eq
--- a/src/HOL/OrderedGroup.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/OrderedGroup.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:   HOL/OrderedGroup.thy
-    ID:      $Id$
-    Author:  Gertrud Bauer, Steven Obua, Lawrence C Paulson, and Markus Wenzel,
-             with contributions by Jeremy Avigad
+    Author:  Gertrud Bauer, Steven Obua, Lawrence C Paulson, Markus Wenzel, Jeremy Avigad
 *)
 
 header {* Ordered Groups *}
@@ -1376,7 +1374,7 @@
         @{const_name HOL.uminus}, @{const_name HOL.minus}]
   | agrp_ord _ = ~1;
 
-fun termless_agrp (a, b) = (Term.term_lpo agrp_ord (a, b) = LESS);
+fun termless_agrp (a, b) = (TermOrd.term_lpo agrp_ord (a, b) = LESS);
 
 local
   val ac1 = mk_meta_eq @{thm add_assoc};
--- a/src/HOL/PReal.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/PReal.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -9,7 +9,7 @@
 header {* Positive real numbers *}
 
 theory PReal
-imports Rational "~~/src/HOL/Library/Dense_Linear_Order"
+imports Rational Dense_Linear_Order
 begin
 
 text{*Could be generalized and moved to @{text Ring_and_Field}*}
--- a/src/HOL/Plain.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Plain.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,15 +1,14 @@
-(*  Title:      HOL/Plain.thy
-    ID:         $Id$
-
-Contains fundamental HOL tools and packages.  Does not include Hilbert_Choice.
-*)
-
 header {* Plain HOL *}
 
 theory Plain
 imports Datatype FunDef Record SAT Extraction
 begin
 
+text {*
+  Plain bootstrap of fundamental HOL tools and packages; does not
+  include @{text Hilbert_Choice}.
+*}
+
 ML {* path_add "~~/src/HOL/Library" *}
 
 end
--- a/src/HOL/ROOT.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ROOT.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,5 @@
-(*  Title:      HOL/ROOT.ML
- 
-Classical Higher-order Logic -- batteries included.
-*)
+(* Classical Higher-order Logic -- batteries included *)
 
 use_thy "Complex_Main";
 
 val HOL_proofs = ! Proofterm.proofs;
-
--- a/src/HOL/Rational.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Rational.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -851,7 +851,7 @@
 definition Fract_norm :: "int \<Rightarrow> int \<Rightarrow> rat" where
   [simp, code del]: "Fract_norm a b = Fract a b"
 
-lemma [code]: "Fract_norm a b = (if a = 0 \<or> b = 0 then 0 else let c = zgcd a b in
+lemma Fract_norm_code [code]: "Fract_norm a b = (if a = 0 \<or> b = 0 then 0 else let c = zgcd a b in
   if b > 0 then Fract (a div c) (b div c) else Fract (- (a div c)) (- (b div c)))"
   by (simp add: eq_rat Zero_rat_def Let_def Fract_norm)
 
@@ -871,7 +871,7 @@
        then c = 0 \<or> d = 0
      else if d = 0
        then a = 0 \<or> b = 0
-     else a * d =  b * c)"
+     else a * d = b * c)"
   by (auto simp add: eq eq_rat)
 
 lemma rat_eq_refl [code nbe]:
--- a/src/HOL/Real.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Real.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,5 @@
 theory Real
-imports RComplete "~~/src/HOL/Real/RealVector"
+imports RComplete RealVector
 begin
 
 end
--- a/src/HOL/Real/HahnBanach/Bounds.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,83 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/Bounds.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Bounds *}
-
-theory Bounds
-imports Main ContNotDenum
-begin
-
-locale lub =
-  fixes A and x
-  assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
-    and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
-
-lemmas [elim?] = lub.least lub.upper
-
-definition
-  the_lub :: "'a::order set \<Rightarrow> 'a" where
-  "the_lub A = The (lub A)"
-
-notation (xsymbols)
-  the_lub  ("\<Squnion>_" [90] 90)
-
-lemma the_lub_equality [elim?]:
-  assumes "lub A x"
-  shows "\<Squnion>A = (x::'a::order)"
-proof -
-  interpret lub [A x] by fact
-  show ?thesis
-  proof (unfold the_lub_def)
-    from `lub A x` show "The (lub A) = x"
-    proof
-      fix x' assume lub': "lub A x'"
-      show "x' = x"
-      proof (rule order_antisym)
-	from lub' show "x' \<le> x"
-	proof
-          fix a assume "a \<in> A"
-          then show "a \<le> x" ..
-	qed
-	show "x \<le> x'"
-	proof
-          fix a assume "a \<in> A"
-          with lub' show "a \<le> x'" ..
-	qed
-      qed
-    qed
-  qed
-qed
-
-lemma the_lubI_ex:
-  assumes ex: "\<exists>x. lub A x"
-  shows "lub A (\<Squnion>A)"
-proof -
-  from ex obtain x where x: "lub A x" ..
-  also from x have [symmetric]: "\<Squnion>A = x" ..
-  finally show ?thesis .
-qed
-
-lemma lub_compat: "lub A x = isLub UNIV A x"
-proof -
-  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
-    by (rule ext) (simp only: isUb_def)
-  then show ?thesis
-    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
-qed
-
-lemma real_complete:
-  fixes A :: "real set"
-  assumes nonempty: "\<exists>a. a \<in> A"
-    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
-  shows "\<exists>x. lub A x"
-proof -
-  from ex_upper have "\<exists>y. isUb UNIV A y"
-    unfolding isUb_def setle_def by blast
-  with nonempty have "\<exists>x. isLub UNIV A x"
-    by (rule reals_complete)
-  then show ?thesis by (simp only: lub_compat)
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/FunctionNorm.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,279 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/FunctionNorm.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* The norm of a function *}
-
-theory FunctionNorm
-imports NormedSpace FunctionOrder
-begin
-
-subsection {* Continuous linear forms*}
-
-text {*
-  A linear form @{text f} on a normed vector space @{text "(V, \<parallel>\<cdot>\<parallel>)"}
-  is \emph{continuous}, iff it is bounded, i.e.
-  \begin{center}
-  @{text "\<exists>c \<in> R. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
-  \end{center}
-  In our application no other functions than linear forms are
-  considered, so we can define continuous linear forms as bounded
-  linear forms:
-*}
-
-locale continuous = var V + norm_syntax + linearform +
-  assumes bounded: "\<exists>c. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
-
-declare continuous.intro [intro?] continuous_axioms.intro [intro?]
-
-lemma continuousI [intro]:
-  fixes norm :: "_ \<Rightarrow> real"  ("\<parallel>_\<parallel>")
-  assumes "linearform V f"
-  assumes r: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
-  shows "continuous V norm f"
-proof
-  show "linearform V f" by fact
-  from r have "\<exists>c. \<forall>x\<in>V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" by blast
-  then show "continuous_axioms V norm f" ..
-qed
-
-
-subsection {* The norm of a linear form *}
-
-text {*
-  The least real number @{text c} for which holds
-  \begin{center}
-  @{text "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
-  \end{center}
-  is called the \emph{norm} of @{text f}.
-
-  For non-trivial vector spaces @{text "V \<noteq> {0}"} the norm can be
-  defined as
-  \begin{center}
-  @{text "\<parallel>f\<parallel> = \<sup>x \<noteq> 0. \<bar>f x\<bar> / \<parallel>x\<parallel>"}
-  \end{center}
-
-  For the case @{text "V = {0}"} the supremum would be taken from an
-  empty set. Since @{text \<real>} is unbounded, there would be no supremum.
-  To avoid this situation it must be guaranteed that there is an
-  element in this set. This element must be @{text "{} \<ge> 0"} so that
-  @{text fn_norm} has the norm properties. Furthermore it does not
-  have to change the norm in all other cases, so it must be @{text 0},
-  as all other elements are @{text "{} \<ge> 0"}.
-
-  Thus we define the set @{text B} where the supremum is taken from as
-  follows:
-  \begin{center}
-  @{text "{0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel>. x \<noteq> 0 \<and> x \<in> F}"}
-  \end{center}
-
-  @{text fn_norm} is equal to the supremum of @{text B}, if the
-  supremum exists (otherwise it is undefined).
-*}
-
-locale fn_norm = norm_syntax +
-  fixes B defines "B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
-  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
-  defines "\<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
-
-locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm
-
-lemma (in fn_norm) B_not_empty [intro]: "0 \<in> B V f"
-  by (simp add: B_def)
-
-text {*
-  The following lemma states that every continuous linear form on a
-  normed space @{text "(V, \<parallel>\<cdot>\<parallel>)"} has a function norm.
-*}
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_works:
-  assumes "continuous V norm f"
-  shows "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
-proof -
-  interpret continuous [V norm f] by fact
-  txt {* The existence of the supremum is shown using the
-    completeness of the reals. Completeness means, that every
-    non-empty bounded set of reals has a supremum. *}
-  have "\<exists>a. lub (B V f) a"
-  proof (rule real_complete)
-    txt {* First we have to show that @{text B} is non-empty: *}
-    have "0 \<in> B V f" ..
-    then show "\<exists>x. x \<in> B V f" ..
-
-    txt {* Then we have to show that @{text B} is bounded: *}
-    show "\<exists>c. \<forall>y \<in> B V f. y \<le> c"
-    proof -
-      txt {* We know that @{text f} is bounded by some value @{text c}. *}
-      from bounded obtain c where c: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
-
-      txt {* To prove the thesis, we have to show that there is some
-        @{text b}, such that @{text "y \<le> b"} for all @{text "y \<in>
-        B"}. Due to the definition of @{text B} there are two cases. *}
-
-      def b \<equiv> "max c 0"
-      have "\<forall>y \<in> B V f. y \<le> b"
-      proof
-        fix y assume y: "y \<in> B V f"
-        show "y \<le> b"
-        proof cases
-          assume "y = 0"
-          then show ?thesis unfolding b_def by arith
-        next
-          txt {* The second case is @{text "y = \<bar>f x\<bar> / \<parallel>x\<parallel>"} for some
-            @{text "x \<in> V"} with @{text "x \<noteq> 0"}. *}
-          assume "y \<noteq> 0"
-          with y obtain x where y_rep: "y = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
-              and x: "x \<in> V" and neq: "x \<noteq> 0"
-            by (auto simp add: B_def real_divide_def)
-          from x neq have gt: "0 < \<parallel>x\<parallel>" ..
-
-          txt {* The thesis follows by a short calculation using the
-            fact that @{text f} is bounded. *}
-
-          note y_rep
-          also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
-          proof (rule mult_right_mono)
-            from c x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
-            from gt have "0 < inverse \<parallel>x\<parallel>" 
-              by (rule positive_imp_inverse_positive)
-            then show "0 \<le> inverse \<parallel>x\<parallel>" by (rule order_less_imp_le)
-          qed
-          also have "\<dots> = c * (\<parallel>x\<parallel> * inverse \<parallel>x\<parallel>)"
-            by (rule real_mult_assoc)
-          also
-          from gt have "\<parallel>x\<parallel> \<noteq> 0" by simp
-          then have "\<parallel>x\<parallel> * inverse \<parallel>x\<parallel> = 1" by simp 
-          also have "c * 1 \<le> b" by (simp add: b_def le_maxI1)
-          finally show "y \<le> b" .
-        qed
-      qed
-      then show ?thesis ..
-    qed
-  qed
-  then show ?thesis unfolding fn_norm_def by (rule the_lubI_ex)
-qed
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_ub [iff?]:
-  assumes "continuous V norm f"
-  assumes b: "b \<in> B V f"
-  shows "b \<le> \<parallel>f\<parallel>\<hyphen>V"
-proof -
-  interpret continuous [V norm f] by fact
-  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
-    using `continuous V norm f` by (rule fn_norm_works)
-  from this and b show ?thesis ..
-qed
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_leastB:
-  assumes "continuous V norm f"
-  assumes b: "\<And>b. b \<in> B V f \<Longrightarrow> b \<le> y"
-  shows "\<parallel>f\<parallel>\<hyphen>V \<le> y"
-proof -
-  interpret continuous [V norm f] by fact
-  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
-    using `continuous V norm f` by (rule fn_norm_works)
-  from this and b show ?thesis ..
-qed
-
-text {* The norm of a continuous function is always @{text "\<ge> 0"}. *}
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_ge_zero [iff]:
-  assumes "continuous V norm f"
-  shows "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
-proof -
-  interpret continuous [V norm f] by fact
-  txt {* The function norm is defined as the supremum of @{text B}.
-    So it is @{text "\<ge> 0"} if all elements in @{text B} are @{text "\<ge>
-    0"}, provided the supremum exists and @{text B} is not empty. *}
-  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
-    using `continuous V norm f` by (rule fn_norm_works)
-  moreover have "0 \<in> B V f" ..
-  ultimately show ?thesis ..
-qed
-
-text {*
-  \medskip The fundamental property of function norms is:
-  \begin{center}
-  @{text "\<bar>f x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
-  \end{center}
-*}
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_le_cong:
-  assumes "continuous V norm f" "linearform V f"
-  assumes x: "x \<in> V"
-  shows "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
-proof -
-  interpret continuous [V norm f] by fact
-  interpret linearform [V f] .
-  show ?thesis
-  proof cases
-    assume "x = 0"
-    then have "\<bar>f x\<bar> = \<bar>f 0\<bar>" by simp
-    also have "f 0 = 0" by rule unfold_locales
-    also have "\<bar>\<dots>\<bar> = 0" by simp
-    also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
-      using `continuous V norm f` by (rule fn_norm_ge_zero)
-    from x have "0 \<le> norm x" ..
-    with a have "0 \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" by (simp add: zero_le_mult_iff)
-    finally show "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" .
-  next
-    assume "x \<noteq> 0"
-    with x have neq: "\<parallel>x\<parallel> \<noteq> 0" by simp
-    then have "\<bar>f x\<bar> = (\<bar>f x\<bar> * inverse \<parallel>x\<parallel>) * \<parallel>x\<parallel>" by simp
-    also have "\<dots> \<le>  \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
-    proof (rule mult_right_mono)
-      from x show "0 \<le> \<parallel>x\<parallel>" ..
-      from x and neq have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<in> B V f"
-	by (auto simp add: B_def real_divide_def)
-      with `continuous V norm f` show "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>V"
-	by (rule fn_norm_ub)
-    qed
-    finally show ?thesis .
-  qed
-qed
-
-text {*
-  \medskip The function norm is the least positive real number for
-  which the following inequation holds:
-  \begin{center}
-    @{text "\<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
-  \end{center}
-*}
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_least [intro?]:
-  assumes "continuous V norm f"
-  assumes ineq: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" and ge: "0 \<le> c"
-  shows "\<parallel>f\<parallel>\<hyphen>V \<le> c"
-proof -
-  interpret continuous [V norm f] by fact
-  show ?thesis
-  proof (rule fn_norm_leastB [folded B_def fn_norm_def])
-    fix b assume b: "b \<in> B V f"
-    show "b \<le> c"
-    proof cases
-      assume "b = 0"
-      with ge show ?thesis by simp
-    next
-      assume "b \<noteq> 0"
-      with b obtain x where b_rep: "b = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
-        and x_neq: "x \<noteq> 0" and x: "x \<in> V"
-	by (auto simp add: B_def real_divide_def)
-      note b_rep
-      also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
-      proof (rule mult_right_mono)
-	have "0 < \<parallel>x\<parallel>" using x x_neq ..
-	then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
-	from ineq and x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
-      qed
-      also have "\<dots> = c"
-      proof -
-	from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
-	then show ?thesis by simp
-      qed
-      finally show ?thesis .
-    qed
-  qed (insert `continuous V norm f`, simp_all add: continuous_def)
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/FunctionOrder.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,142 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/FunctionOrder.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* An order on functions *}
-
-theory FunctionOrder
-imports Subspace Linearform
-begin
-
-subsection {* The graph of a function *}
-
-text {*
-  We define the \emph{graph} of a (real) function @{text f} with
-  domain @{text F} as the set
-  \begin{center}
-  @{text "{(x, f x). x \<in> F}"}
-  \end{center}
-  So we are modeling partial functions by specifying the domain and
-  the mapping function. We use the term ``function'' also for its
-  graph.
-*}
-
-types 'a graph = "('a \<times> real) set"
-
-definition
-  graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph" where
-  "graph F f = {(x, f x) | x. x \<in> F}"
-
-lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
-  unfolding graph_def by blast
-
-lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"
-  unfolding graph_def by blast
-
-lemma graphE [elim?]:
-    "(x, y) \<in> graph F f \<Longrightarrow> (x \<in> F \<Longrightarrow> y = f x \<Longrightarrow> C) \<Longrightarrow> C"
-  unfolding graph_def by blast
-
-
-subsection {* Functions ordered by domain extension *}
-
-text {*
-  A function @{text h'} is an extension of @{text h}, iff the graph of
-  @{text h} is a subset of the graph of @{text h'}.
-*}
-
-lemma graph_extI:
-  "(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'
-    \<Longrightarrow> graph H h \<subseteq> graph H' h'"
-  unfolding graph_def by blast
-
-lemma graph_extD1 [dest?]:
-  "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"
-  unfolding graph_def by blast
-
-lemma graph_extD2 [dest?]:
-  "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"
-  unfolding graph_def by blast
-
-
-subsection {* Domain and function of a graph *}
-
-text {*
-  The inverse functions to @{text graph} are @{text domain} and @{text
-  funct}.
-*}
-
-definition
-  "domain" :: "'a graph \<Rightarrow> 'a set" where
-  "domain g = {x. \<exists>y. (x, y) \<in> g}"
-
-definition
-  funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)" where
-  "funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"
-
-text {*
-  The following lemma states that @{text g} is the graph of a function
-  if the relation induced by @{text g} is unique.
-*}
-
-lemma graph_domain_funct:
-  assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"
-  shows "graph (domain g) (funct g) = g"
-  unfolding domain_def funct_def graph_def
-proof auto  (* FIXME !? *)
-  fix a b assume g: "(a, b) \<in> g"
-  from g show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)
-  from g show "\<exists>y. (a, y) \<in> g" ..
-  from g show "b = (SOME y. (a, y) \<in> g)"
-  proof (rule some_equality [symmetric])
-    fix y assume "(a, y) \<in> g"
-    with g show "y = b" by (rule uniq)
-  qed
-qed
-
-
-subsection {* Norm-preserving extensions of a function *}
-
-text {*
-  Given a linear form @{text f} on the space @{text F} and a seminorm
-  @{text p} on @{text E}. The set of all linear extensions of @{text
-  f}, to superspaces @{text H} of @{text F}, which are bounded by
-  @{text p}, is defined as follows.
-*}
-
-definition
-  norm_pres_extensions ::
-    "'a::{plus, minus, uminus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
-      \<Rightarrow> 'a graph set" where
-    "norm_pres_extensions E p F f
-      = {g. \<exists>H h. g = graph H h
-          \<and> linearform H h
-          \<and> H \<unlhd> E
-          \<and> F \<unlhd> H
-          \<and> graph F f \<subseteq> graph H h
-          \<and> (\<forall>x \<in> H. h x \<le> p x)}"
-
-lemma norm_pres_extensionE [elim]:
-  "g \<in> norm_pres_extensions E p F f
-  \<Longrightarrow> (\<And>H h. g = graph H h \<Longrightarrow> linearform H h
-        \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H \<Longrightarrow> graph F f \<subseteq> graph H h
-        \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x \<Longrightarrow> C) \<Longrightarrow> C"
-  unfolding norm_pres_extensions_def by blast
-
-lemma norm_pres_extensionI2 [intro]:
-  "linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H
-    \<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x
-    \<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"
-  unfolding norm_pres_extensions_def by blast
-
-lemma norm_pres_extensionI:  (* FIXME ? *)
-  "\<exists>H h. g = graph H h
-    \<and> linearform H h
-    \<and> H \<unlhd> E
-    \<and> F \<unlhd> H
-    \<and> graph F f \<subseteq> graph H h
-    \<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"
-  unfolding norm_pres_extensions_def by blast
-
-end
--- a/src/HOL/Real/HahnBanach/HahnBanach.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,510 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* The Hahn-Banach Theorem *}
-
-theory HahnBanach
-imports HahnBanachLemmas
-begin
-
-text {*
-  We present the proof of two different versions of the Hahn-Banach
-  Theorem, closely following \cite[\S36]{Heuser:1986}.
-*}
-
-subsection {* The Hahn-Banach Theorem for vector spaces *}
-
-text {*
-  \textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
-  vector space @{text E}, let @{text p} be a semi-norm on @{text E},
-  and @{text f} be a linear form defined on @{text F} such that @{text
-  f} is bounded by @{text p}, i.e.  @{text "\<forall>x \<in> F. f x \<le> p x"}.  Then
-  @{text f} can be extended to a linear form @{text h} on @{text E}
-  such that @{text h} is norm-preserving, i.e. @{text h} is also
-  bounded by @{text p}.
-
-  \bigskip
-  \textbf{Proof Sketch.}
-  \begin{enumerate}
-
-  \item Define @{text M} as the set of norm-preserving extensions of
-  @{text f} to subspaces of @{text E}. The linear forms in @{text M}
-  are ordered by domain extension.
-
-  \item We show that every non-empty chain in @{text M} has an upper
-  bound in @{text M}.
-
-  \item With Zorn's Lemma we conclude that there is a maximal function
-  @{text g} in @{text M}.
-
-  \item The domain @{text H} of @{text g} is the whole space @{text
-  E}, as shown by classical contradiction:
-
-  \begin{itemize}
-
-  \item Assuming @{text g} is not defined on whole @{text E}, it can
-  still be extended in a norm-preserving way to a super-space @{text
-  H'} of @{text H}.
-
-  \item Thus @{text g} can not be maximal. Contradiction!
-
-  \end{itemize}
-  \end{enumerate}
-*}
-
-theorem HahnBanach:
-  assumes E: "vectorspace E" and "subspace F E"
-    and "seminorm E p" and "linearform F f"
-  assumes fp: "\<forall>x \<in> F. f x \<le> p x"
-  shows "\<exists>h. linearform E h \<and> (\<forall>x \<in> F. h x = f x) \<and> (\<forall>x \<in> E. h x \<le> p x)"
-    -- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
-    -- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
-    -- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [F E] by fact
-  interpret seminorm [E p] by fact
-  interpret linearform [F f] by fact
-  def M \<equiv> "norm_pres_extensions E p F f"
-  then have M: "M = \<dots>" by (simp only:)
-  from E have F: "vectorspace F" ..
-  note FE = `F \<unlhd> E`
-  {
-    fix c assume cM: "c \<in> chain M" and ex: "\<exists>x. x \<in> c"
-    have "\<Union>c \<in> M"
-      -- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
-      -- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
-      unfolding M_def
-    proof (rule norm_pres_extensionI)
-      let ?H = "domain (\<Union>c)"
-      let ?h = "funct (\<Union>c)"
-
-      have a: "graph ?H ?h = \<Union>c"
-      proof (rule graph_domain_funct)
-        fix x y z assume "(x, y) \<in> \<Union>c" and "(x, z) \<in> \<Union>c"
-        with M_def cM show "z = y" by (rule sup_definite)
-      qed
-      moreover from M cM a have "linearform ?H ?h"
-        by (rule sup_lf)
-      moreover from a M cM ex FE E have "?H \<unlhd> E"
-        by (rule sup_subE)
-      moreover from a M cM ex FE have "F \<unlhd> ?H"
-        by (rule sup_supF)
-      moreover from a M cM ex have "graph F f \<subseteq> graph ?H ?h"
-        by (rule sup_ext)
-      moreover from a M cM have "\<forall>x \<in> ?H. ?h x \<le> p x"
-        by (rule sup_norm_pres)
-      ultimately show "\<exists>H h. \<Union>c = graph H h
-          \<and> linearform H h
-          \<and> H \<unlhd> E
-          \<and> F \<unlhd> H
-          \<and> graph F f \<subseteq> graph H h
-          \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
-    qed
-  }
-  then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
-  -- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
-  proof (rule Zorn's_Lemma)
-      -- {* We show that @{text M} is non-empty: *}
-    show "graph F f \<in> M"
-      unfolding M_def
-    proof (rule norm_pres_extensionI2)
-      show "linearform F f" by fact
-      show "F \<unlhd> E" by fact
-      from F show "F \<unlhd> F" by (rule vectorspace.subspace_refl)
-      show "graph F f \<subseteq> graph F f" ..
-      show "\<forall>x\<in>F. f x \<le> p x" by fact
-    qed
-  qed
-  then obtain g where gM: "g \<in> M" and gx: "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
-    by blast
-  from gM obtain H h where
-      g_rep: "g = graph H h"
-    and linearform: "linearform H h"
-    and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
-    and graphs: "graph F f \<subseteq> graph H h"
-    and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
-      -- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
-      -- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
-      -- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
-  from HE E have H: "vectorspace H"
-    by (rule subspace.vectorspace)
-
-  have HE_eq: "H = E"
-    -- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
-  proof (rule classical)
-    assume neq: "H \<noteq> E"
-      -- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
-      -- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
-    have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
-    proof -
-      from HE have "H \<subseteq> E" ..
-      with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
-      obtain x': "x' \<noteq> 0"
-      proof
-        show "x' \<noteq> 0"
-        proof
-          assume "x' = 0"
-          with H have "x' \<in> H" by (simp only: vectorspace.zero)
-          with `x' \<notin> H` show False by contradiction
-        qed
-      qed
-
-      def H' \<equiv> "H + lin x'"
-        -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
-      have HH': "H \<unlhd> H'"
-      proof (unfold H'_def)
-        from x'E have "vectorspace (lin x')" ..
-        with H show "H \<unlhd> H + lin x'" ..
-      qed
-
-      obtain xi where
-        xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
-          \<and> xi \<le> p (y + x') - h y"
-        -- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
-        -- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
-           \label{ex-xi-use}\skp *}
-      proof -
-        from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
-            \<and> xi \<le> p (y + x') - h y"
-        proof (rule ex_xi)
-          fix u v assume u: "u \<in> H" and v: "v \<in> H"
-          with HE have uE: "u \<in> E" and vE: "v \<in> E" by auto
-          from H u v linearform have "h v - h u = h (v - u)"
-            by (simp add: linearform.diff)
-          also from hp and H u v have "\<dots> \<le> p (v - u)"
-            by (simp only: vectorspace.diff_closed)
-          also from x'E uE vE have "v - u = x' + - x' + v + - u"
-            by (simp add: diff_eq1)
-          also from x'E uE vE have "\<dots> = v + x' + - (u + x')"
-            by (simp add: add_ac)
-          also from x'E uE vE have "\<dots> = (v + x') - (u + x')"
-            by (simp add: diff_eq1)
-          also from x'E uE vE E have "p \<dots> \<le> p (v + x') + p (u + x')"
-            by (simp add: diff_subadditive)
-          finally have "h v - h u \<le> p (v + x') + p (u + x')" .
-          then show "- p (u + x') - h u \<le> p (v + x') - h v" by simp
-        qed
-        then show thesis by (blast intro: that)
-      qed
-
-      def h' \<equiv> "\<lambda>x. let (y, a) =
-          SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
-        -- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
-
-      have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
-        -- {* @{text h'} is an extension of @{text h} \dots \skp *}
-      proof
-        show "g \<subseteq> graph H' h'"
-        proof -
-          have  "graph H h \<subseteq> graph H' h'"
-          proof (rule graph_extI)
-            fix t assume t: "t \<in> H"
-            from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
-	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
-            with h'_def show "h t = h' t" by (simp add: Let_def)
-          next
-            from HH' show "H \<subseteq> H'" ..
-          qed
-          with g_rep show ?thesis by (simp only:)
-        qed
-
-        show "g \<noteq> graph H' h'"
-        proof -
-          have "graph H h \<noteq> graph H' h'"
-          proof
-            assume eq: "graph H h = graph H' h'"
-            have "x' \<in> H'"
-	      unfolding H'_def
-            proof
-              from H show "0 \<in> H" by (rule vectorspace.zero)
-              from x'E show "x' \<in> lin x'" by (rule x_lin_x)
-              from x'E show "x' = 0 + x'" by simp
-            qed
-            then have "(x', h' x') \<in> graph H' h'" ..
-            with eq have "(x', h' x') \<in> graph H h" by (simp only:)
-            then have "x' \<in> H" ..
-            with `x' \<notin> H` show False by contradiction
-          qed
-          with g_rep show ?thesis by simp
-        qed
-      qed
-      moreover have "graph H' h' \<in> M"
-        -- {* and @{text h'} is norm-preserving. \skp *}
-      proof (unfold M_def)
-        show "graph H' h' \<in> norm_pres_extensions E p F f"
-        proof (rule norm_pres_extensionI2)
-          show "linearform H' h'"
-	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
-	    by (rule h'_lf)
-          show "H' \<unlhd> E"
-	  unfolding H'_def
-          proof
-            show "H \<unlhd> E" by fact
-            show "vectorspace E" by fact
-            from x'E show "lin x' \<unlhd> E" ..
-          qed
-          from H `F \<unlhd> H` HH' show FH': "F \<unlhd> H'"
-            by (rule vectorspace.subspace_trans)
-          show "graph F f \<subseteq> graph H' h'"
-          proof (rule graph_extI)
-            fix x assume x: "x \<in> F"
-            with graphs have "f x = h x" ..
-            also have "\<dots> = h x + 0 * xi" by simp
-            also have "\<dots> = (let (y, a) = (x, 0) in h y + a * xi)"
-              by (simp add: Let_def)
-            also have "(x, 0) =
-                (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
-	      using E HE
-            proof (rule decomp_H'_H [symmetric])
-              from FH x show "x \<in> H" ..
-              from x' show "x' \<noteq> 0" .
-	      show "x' \<notin> H" by fact
-	      show "x' \<in> E" by fact
-            qed
-            also have
-              "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
-              in h y + a * xi) = h' x" by (simp only: h'_def)
-            finally show "f x = h' x" .
-          next
-            from FH' show "F \<subseteq> H'" ..
-          qed
-          show "\<forall>x \<in> H'. h' x \<le> p x"
-	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
-	      `seminorm E p` linearform and hp xi
-	    by (rule h'_norm_pres)
-        qed
-      qed
-      ultimately show ?thesis ..
-    qed
-    then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
-      -- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
-    with gx show "H = E" by contradiction
-  qed
-
-  from HE_eq and linearform have "linearform E h"
-    by (simp only:)
-  moreover have "\<forall>x \<in> F. h x = f x"
-  proof
-    fix x assume "x \<in> F"
-    with graphs have "f x = h x" ..
-    then show "h x = f x" ..
-  qed
-  moreover from HE_eq and hp have "\<forall>x \<in> E. h x \<le> p x"
-    by (simp only:)
-  ultimately show ?thesis by blast
-qed
-
-
-subsection  {* Alternative formulation *}
-
-text {*
-  The following alternative formulation of the Hahn-Banach
-  Theorem\label{abs-HahnBanach} uses the fact that for a real linear
-  form @{text f} and a seminorm @{text p} the following inequations
-  are equivalent:\footnote{This was shown in lemma @{thm [source]
-  abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
-  \begin{center}
-  \begin{tabular}{lll}
-  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
-  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
-  \end{tabular}
-  \end{center}
-*}
-
-theorem abs_HahnBanach:
-  assumes E: "vectorspace E" and FE: "subspace F E"
-    and lf: "linearform F f" and sn: "seminorm E p"
-  assumes fp: "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
-  shows "\<exists>g. linearform E g
-    \<and> (\<forall>x \<in> F. g x = f x)
-    \<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [F E] by fact
-  interpret linearform [F f] by fact
-  interpret seminorm [E p] by fact
-  have "\<exists>g. linearform E g \<and> (\<forall>x \<in> F. g x = f x) \<and> (\<forall>x \<in> E. g x \<le> p x)"
-    using E FE sn lf
-  proof (rule HahnBanach)
-    show "\<forall>x \<in> F. f x \<le> p x"
-      using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
-  qed
-  then obtain g where lg: "linearform E g" and *: "\<forall>x \<in> F. g x = f x"
-      and **: "\<forall>x \<in> E. g x \<le> p x" by blast
-  have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
-    using _ E sn lg **
-  proof (rule abs_ineq_iff [THEN iffD2])
-    show "E \<unlhd> E" ..
-  qed
-  with lg * show ?thesis by blast
-qed
-
-
-subsection {* The Hahn-Banach Theorem for normed spaces *}
-
-text {*
-  Every continuous linear form @{text f} on a subspace @{text F} of a
-  norm space @{text E}, can be extended to a continuous linear form
-  @{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
-*}
-
-theorem norm_HahnBanach:
-  fixes V and norm ("\<parallel>_\<parallel>")
-  fixes B defines "\<And>V f. B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
-  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
-  defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
-  assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
-    and linearform: "linearform F f" and "continuous F norm f"
-  shows "\<exists>g. linearform E g
-     \<and> continuous E norm g
-     \<and> (\<forall>x \<in> F. g x = f x)
-     \<and> \<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
-proof -
-  interpret normed_vectorspace [E norm] by fact
-  interpret normed_vectorspace_with_fn_norm [E norm B fn_norm]
-    by (auto simp: B_def fn_norm_def) intro_locales
-  interpret subspace [F E] by fact
-  interpret linearform [F f] by fact
-  interpret continuous [F norm f] by fact
-  have E: "vectorspace E" by intro_locales
-  have F: "vectorspace F" by rule intro_locales
-  have F_norm: "normed_vectorspace F norm"
-    using FE E_norm by (rule subspace_normed_vs)
-  have ge_zero: "0 \<le> \<parallel>f\<parallel>\<hyphen>F"
-    by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
-      [OF normed_vectorspace_with_fn_norm.intro,
-       OF F_norm `continuous F norm f` , folded B_def fn_norm_def])
-  txt {* We define a function @{text p} on @{text E} as follows:
-    @{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"} *}
-  def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
-
-  txt {* @{text p} is a seminorm on @{text E}: *}
-  have q: "seminorm E p"
-  proof
-    fix x y a assume x: "x \<in> E" and y: "y \<in> E"
-    
-    txt {* @{text p} is positive definite: *}
-    have "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
-    moreover from x have "0 \<le> \<parallel>x\<parallel>" ..
-    ultimately show "0 \<le> p x"  
-      by (simp add: p_def zero_le_mult_iff)
-
-    txt {* @{text p} is absolutely homogenous: *}
-
-    show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
-    proof -
-      have "p (a \<cdot> x) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>a \<cdot> x\<parallel>" by (simp only: p_def)
-      also from x have "\<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>" by (rule abs_homogenous)
-      also have "\<parallel>f\<parallel>\<hyphen>F * (\<bar>a\<bar> * \<parallel>x\<parallel>) = \<bar>a\<bar> * (\<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>)" by simp
-      also have "\<dots> = \<bar>a\<bar> * p x" by (simp only: p_def)
-      finally show ?thesis .
-    qed
-
-    txt {* Furthermore, @{text p} is subadditive: *}
-
-    show "p (x + y) \<le> p x + p y"
-    proof -
-      have "p (x + y) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel>" by (simp only: p_def)
-      also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
-      from x y have "\<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>" ..
-      with a have " \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>F * (\<parallel>x\<parallel> + \<parallel>y\<parallel>)"
-        by (simp add: mult_left_mono)
-      also have "\<dots> = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel> + \<parallel>f\<parallel>\<hyphen>F * \<parallel>y\<parallel>" by (simp only: right_distrib)
-      also have "\<dots> = p x + p y" by (simp only: p_def)
-      finally show ?thesis .
-    qed
-  qed
-
-  txt {* @{text f} is bounded by @{text p}. *}
-
-  have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
-  proof
-    fix x assume "x \<in> F"
-    with `continuous F norm f` and linearform
-    show "\<bar>f x\<bar> \<le> p x"
-      unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
-        [OF normed_vectorspace_with_fn_norm.intro,
-         OF F_norm, folded B_def fn_norm_def])
-  qed
-
-  txt {* Using the fact that @{text p} is a seminorm and @{text f} is bounded
-    by @{text p} we can apply the Hahn-Banach Theorem for real vector
-    spaces. So @{text f} can be extended in a norm-preserving way to
-    some function @{text g} on the whole vector space @{text E}. *}
-
-  with E FE linearform q obtain g where
-      linearformE: "linearform E g"
-    and a: "\<forall>x \<in> F. g x = f x"
-    and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
-    by (rule abs_HahnBanach [elim_format]) iprover
-
-  txt {* We furthermore have to show that @{text g} is also continuous: *}
-
-  have g_cont: "continuous E norm g" using linearformE
-  proof
-    fix x assume "x \<in> E"
-    with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
-      by (simp only: p_def)
-  qed
-
-  txt {* To complete the proof, we show that @{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. *}
-
-  have "\<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
-  proof (rule order_antisym)
-    txt {*
-      First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}.  The function norm @{text
-      "\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
-      \begin{center}
-      \begin{tabular}{l}
-      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
-      \end{tabular}
-      \end{center}
-      \noindent Furthermore holds
-      \begin{center}
-      \begin{tabular}{l}
-      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
-      \end{tabular}
-      \end{center}
-    *}
-
-    have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
-    proof
-      fix x assume "x \<in> E"
-      with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
-        by (simp only: p_def)
-    qed
-    from g_cont this ge_zero
-    show "\<parallel>g\<parallel>\<hyphen>E \<le> \<parallel>f\<parallel>\<hyphen>F"
-      by (rule fn_norm_least [of g, folded B_def fn_norm_def])
-
-    txt {* The other direction is achieved by a similar argument. *}
-
-    show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
-    proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
-	[OF normed_vectorspace_with_fn_norm.intro,
-	 OF F_norm, folded B_def fn_norm_def])
-      show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
-      proof
-	fix x assume x: "x \<in> F"
-	from a x have "g x = f x" ..
-	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
-	also from g_cont
-	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
-	proof (rule fn_norm_le_cong [of g, folded B_def fn_norm_def])
-	  from FE x show "x \<in> E" ..
-	qed
-	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
-      qed
-      show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
-	using g_cont
-	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
-      show "continuous F norm f" by fact
-    qed
-  qed
-  with linearformE a g_cont show ?thesis by blast
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/HahnBanachExtLemmas.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,281 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Extending non-maximal functions *}
-
-theory HahnBanachExtLemmas
-imports FunctionNorm
-begin
-
-text {*
-  In this section the following context is presumed.  Let @{text E} be
-  a real vector space with a seminorm @{text q} on @{text E}. @{text
-  F} is a subspace of @{text E} and @{text f} a linear function on
-  @{text F}. We consider a subspace @{text H} of @{text E} that is a
-  superspace of @{text F} and a linear form @{text h} on @{text
-  H}. @{text H} is a not equal to @{text E} and @{text "x\<^sub>0"} is
-  an element in @{text "E - H"}.  @{text H} is extended to the direct
-  sum @{text "H' = H + lin x\<^sub>0"}, so for any @{text "x \<in> H'"}
-  the decomposition of @{text "x = y + a \<cdot> x"} with @{text "y \<in> H"} is
-  unique. @{text h'} is defined on @{text H'} by @{text "h' x = h y +
-  a \<cdot> \<xi>"} for a certain @{text \<xi>}.
-
-  Subsequently we show some properties of this extension @{text h'} of
-  @{text h}.
-
-  \medskip This lemma will be used to show the existence of a linear
-  extension of @{text f} (see page \pageref{ex-xi-use}). It is a
-  consequence of the completeness of @{text \<real>}. To show
-  \begin{center}
-  \begin{tabular}{l}
-  @{text "\<exists>\<xi>. \<forall>y \<in> F. a y \<le> \<xi> \<and> \<xi> \<le> b y"}
-  \end{tabular}
-  \end{center}
-  \noindent it suffices to show that
-  \begin{center}
-  \begin{tabular}{l}
-  @{text "\<forall>u \<in> F. \<forall>v \<in> F. a u \<le> b v"}
-  \end{tabular}
-  \end{center}
-*}
-
-lemma ex_xi:
-  assumes "vectorspace F"
-  assumes r: "\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> b v"
-  shows "\<exists>xi::real. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
-proof -
-  interpret vectorspace [F] by fact
-  txt {* From the completeness of the reals follows:
-    The set @{text "S = {a u. u \<in> F}"} has a supremum, if it is
-    non-empty and has an upper bound. *}
-
-  let ?S = "{a u | u. u \<in> F}"
-  have "\<exists>xi. lub ?S xi"
-  proof (rule real_complete)
-    have "a 0 \<in> ?S" by blast
-    then show "\<exists>X. X \<in> ?S" ..
-    have "\<forall>y \<in> ?S. y \<le> b 0"
-    proof
-      fix y assume y: "y \<in> ?S"
-      then obtain u where u: "u \<in> F" and y: "y = a u" by blast
-      from u and zero have "a u \<le> b 0" by (rule r)
-      with y show "y \<le> b 0" by (simp only:)
-    qed
-    then show "\<exists>u. \<forall>y \<in> ?S. y \<le> u" ..
-  qed
-  then obtain xi where xi: "lub ?S xi" ..
-  {
-    fix y assume "y \<in> F"
-    then have "a y \<in> ?S" by blast
-    with xi have "a y \<le> xi" by (rule lub.upper)
-  } moreover {
-    fix y assume y: "y \<in> F"
-    from xi have "xi \<le> b y"
-    proof (rule lub.least)
-      fix au assume "au \<in> ?S"
-      then obtain u where u: "u \<in> F" and au: "au = a u" by blast
-      from u y have "a u \<le> b y" by (rule r)
-      with au show "au \<le> b y" by (simp only:)
-    qed
-  } ultimately show "\<exists>xi. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y" by blast
-qed
-
-text {*
-  \medskip The function @{text h'} is defined as a @{text "h' x = h y
-  + a \<cdot> \<xi>"} where @{text "x = y + a \<cdot> \<xi>"} is a linear extension of
-  @{text h} to @{text H'}.
-*}
-
-lemma h'_lf:
-  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
-      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
-    and H'_def: "H' \<equiv> H + lin x0"
-    and HE: "H \<unlhd> E"
-  assumes "linearform H h"
-  assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
-  assumes E: "vectorspace E"
-  shows "linearform H' h'"
-proof -
-  interpret linearform [H h] by fact
-  interpret vectorspace [E] by fact
-  show ?thesis
-  proof
-    note E = `vectorspace E`
-    have H': "vectorspace H'"
-    proof (unfold H'_def)
-      from `x0 \<in> E`
-      have "lin x0 \<unlhd> E" ..
-      with HE show "vectorspace (H + lin x0)" using E ..
-    qed
-    {
-      fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
-      show "h' (x1 + x2) = h' x1 + h' x2"
-      proof -
-	from H' x1 x2 have "x1 + x2 \<in> H'"
-          by (rule vectorspace.add_closed)
-	with x1 x2 obtain y y1 y2 a a1 a2 where
-          x1x2: "x1 + x2 = y + a \<cdot> x0" and y: "y \<in> H"
-          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
-          and x2_rep: "x2 = y2 + a2 \<cdot> x0" and y2: "y2 \<in> H"
-          unfolding H'_def sum_def lin_def by blast
-	
-	have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
-	proof (rule decomp_H') txt_raw {* \label{decomp-H-use} *}
-          from HE y1 y2 show "y1 + y2 \<in> H"
-            by (rule subspace.add_closed)
-          from x0 and HE y y1 y2
-          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E"  "y2 \<in> E" by auto
-          with x1_rep x2_rep have "(y1 + y2) + (a1 + a2) \<cdot> x0 = x1 + x2"
-            by (simp add: add_ac add_mult_distrib2)
-          also note x1x2
-          finally show "(y1 + y2) + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0" .
-	qed
-	
-	from h'_def x1x2 E HE y x0
-	have "h' (x1 + x2) = h y + a * xi"
-          by (rule h'_definite)
-	also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
-          by (simp only: ya)
-	also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
-          by simp
-	also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
-          by (simp add: left_distrib)
-	also from h'_def x1_rep E HE y1 x0
-	have "h y1 + a1 * xi = h' x1"
-          by (rule h'_definite [symmetric])
-	also from h'_def x2_rep E HE y2 x0
-	have "h y2 + a2 * xi = h' x2"
-          by (rule h'_definite [symmetric])
-	finally show ?thesis .
-      qed
-    next
-      fix x1 c assume x1: "x1 \<in> H'"
-      show "h' (c \<cdot> x1) = c * (h' x1)"
-      proof -
-	from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
-          by (rule vectorspace.mult_closed)
-	with x1 obtain y a y1 a1 where
-            cx1_rep: "c \<cdot> x1 = y + a \<cdot> x0" and y: "y \<in> H"
-          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
-          unfolding H'_def sum_def lin_def by blast
-	
-	have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
-	proof (rule decomp_H')
-          from HE y1 show "c \<cdot> y1 \<in> H"
-            by (rule subspace.mult_closed)
-          from x0 and HE y y1
-          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E" by auto
-          with x1_rep have "c \<cdot> y1 + (c * a1) \<cdot> x0 = c \<cdot> x1"
-            by (simp add: mult_assoc add_mult_distrib1)
-          also note cx1_rep
-          finally show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0" .
-	qed
-	
-	from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
-          by (rule h'_definite)
-	also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
-          by (simp only: ya)
-	also from y1 have "h (c \<cdot> y1) = c * h y1"
-          by simp
-	also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
-          by (simp only: right_distrib)
-	also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
-          by (rule h'_definite [symmetric])
-	finally show ?thesis .
-      qed
-    }
-  qed
-qed
-
-text {* \medskip The linear extension @{text h'} of @{text h}
-  is bounded by the seminorm @{text p}. *}
-
-lemma h'_norm_pres:
-  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
-      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
-    and H'_def: "H' \<equiv> H + lin x0"
-    and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
-  assumes E: "vectorspace E" and HE: "subspace H E"
-    and "seminorm E p" and "linearform H h"
-  assumes a: "\<forall>y \<in> H. h y \<le> p y"
-    and a': "\<forall>y \<in> H. - p (y + x0) - h y \<le> xi \<and> xi \<le> p (y + x0) - h y"
-  shows "\<forall>x \<in> H'. h' x \<le> p x"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [H E] by fact
-  interpret seminorm [E p] by fact
-  interpret linearform [H h] by fact
-  show ?thesis
-  proof
-    fix x assume x': "x \<in> H'"
-    show "h' x \<le> p x"
-    proof -
-      from a' have a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
-	and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
-      from x' obtain y a where
-          x_rep: "x = y + a \<cdot> x0" and y: "y \<in> H"
-	unfolding H'_def sum_def lin_def by blast
-      from y have y': "y \<in> E" ..
-      from y have ay: "inverse a \<cdot> y \<in> H" by simp
-      
-      from h'_def x_rep E HE y x0 have "h' x = h y + a * xi"
-	by (rule h'_definite)
-      also have "\<dots> \<le> p (y + a \<cdot> x0)"
-      proof (rule linorder_cases)
-	assume z: "a = 0"
-	then have "h y + a * xi = h y" by simp
-	also from a y have "\<dots> \<le> p y" ..
-	also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
-	finally show ?thesis .
-      next
-	txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
-          with @{text ya} taken as @{text "y / a"}: *}
-	assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
-	from a1 ay
-	have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
-	with lz have "a * xi \<le>
-          a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
-          by (simp add: mult_left_mono_neg order_less_imp_le)
-	
-	also have "\<dots> =
-          - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
-	  by (simp add: right_diff_distrib)
-	also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
-          p (a \<cdot> (inverse a \<cdot> y + x0))"
-          by (simp add: abs_homogenous)
-	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
-          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
-	also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
-          by simp
-	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
-	then show ?thesis by simp
-      next
-	txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
-          with @{text ya} taken as @{text "y / a"}: *}
-	assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
-	from a2 ay
-	have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
-	with gz have "a * xi \<le>
-          a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
-          by simp
-	also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
-	  by (simp add: right_diff_distrib)
-	also from gz x0 y'
-	have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
-          by (simp add: abs_homogenous)
-	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
-          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
-	also from nz y have "a * h (inverse a \<cdot> y) = h y"
-          by simp
-	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
-	then show ?thesis by simp
-      qed
-      also from x_rep have "\<dots> = p x" by (simp only:)
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/HahnBanachLemmas.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(*<*)
-theory HahnBanachLemmas imports HahnBanachSupLemmas HahnBanachExtLemmas begin
-end
-(*>*)
\ No newline at end of file
--- a/src/HOL/Real/HahnBanach/HahnBanachSupLemmas.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,446 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* The supremum w.r.t.~the function order *}
-
-theory HahnBanachSupLemmas
-imports FunctionNorm ZornLemma
-begin
-
-text {*
-  This section contains some lemmas that will be used in the proof of
-  the Hahn-Banach Theorem.  In this section the following context is
-  presumed.  Let @{text E} be a real vector space with a seminorm
-  @{text p} on @{text E}.  @{text F} is a subspace of @{text E} and
-  @{text f} a linear form on @{text F}. We consider a chain @{text c}
-  of norm-preserving extensions of @{text f}, such that @{text "\<Union>c =
-  graph H h"}.  We will show some properties about the limit function
-  @{text h}, i.e.\ the supremum of the chain @{text c}.
-
-  \medskip Let @{text c} be a chain of norm-preserving extensions of
-  the function @{text f} and let @{text "graph H h"} be the supremum
-  of @{text c}.  Every element in @{text H} is member of one of the
-  elements of the chain.
-*}
-lemmas [dest?] = chainD
-lemmas chainE2 [elim?] = chainD2 [elim_format, standard]
-
-lemma some_H'h't:
-  assumes M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and u: "graph H h = \<Union>c"
-    and x: "x \<in> H"
-  shows "\<exists>H' h'. graph H' h' \<in> c
-    \<and> (x, h x) \<in> graph H' h'
-    \<and> linearform H' h' \<and> H' \<unlhd> E
-    \<and> F \<unlhd> H' \<and> graph F f \<subseteq> graph H' h'
-    \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
-proof -
-  from x have "(x, h x) \<in> graph H h" ..
-  also from u have "\<dots> = \<Union>c" .
-  finally obtain g where gc: "g \<in> c" and gh: "(x, h x) \<in> g" by blast
-
-  from cM have "c \<subseteq> M" ..
-  with gc have "g \<in> M" ..
-  also from M have "\<dots> = norm_pres_extensions E p F f" .
-  finally obtain H' and h' where g: "g = graph H' h'"
-    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
-      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x" ..
-
-  from gc and g have "graph H' h' \<in> c" by (simp only:)
-  moreover from gh and g have "(x, h x) \<in> graph H' h'" by (simp only:)
-  ultimately show ?thesis using * by blast
-qed
-
-text {*
-  \medskip Let @{text c} be a chain of norm-preserving extensions of
-  the function @{text f} and let @{text "graph H h"} be the supremum
-  of @{text c}.  Every element in the domain @{text H} of the supremum
-  function is member of the domain @{text H'} of some function @{text
-  h'}, such that @{text h} extends @{text h'}.
-*}
-
-lemma some_H'h':
-  assumes M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and u: "graph H h = \<Union>c"
-    and x: "x \<in> H"
-  shows "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
-    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
-    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
-proof -
-  from M cM u x obtain H' h' where
-      x_hx: "(x, h x) \<in> graph H' h'"
-    and c: "graph H' h' \<in> c"
-    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
-      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
-    by (rule some_H'h't [elim_format]) blast
-  from x_hx have "x \<in> H'" ..
-  moreover from cM u c have "graph H' h' \<subseteq> graph H h"
-    by (simp only: chain_ball_Union_upper)
-  ultimately show ?thesis using * by blast
-qed
-
-text {*
-  \medskip Any two elements @{text x} and @{text y} in the domain
-  @{text H} of the supremum function @{text h} are both in the domain
-  @{text H'} of some function @{text h'}, such that @{text h} extends
-  @{text h'}.
-*}
-
-lemma some_H'h'2:
-  assumes M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and u: "graph H h = \<Union>c"
-    and x: "x \<in> H"
-    and y: "y \<in> H"
-  shows "\<exists>H' h'. x \<in> H' \<and> y \<in> H'
-    \<and> graph H' h' \<subseteq> graph H h
-    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
-    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
-proof -
-  txt {* @{text y} is in the domain @{text H''} of some function @{text h''},
-  such that @{text h} extends @{text h''}. *}
-
-  from M cM u and y obtain H' h' where
-      y_hy: "(y, h y) \<in> graph H' h'"
-    and c': "graph H' h' \<in> c"
-    and * :
-      "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
-      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
-    by (rule some_H'h't [elim_format]) blast
-
-  txt {* @{text x} is in the domain @{text H'} of some function @{text h'},
-    such that @{text h} extends @{text h'}. *}
-
-  from M cM u and x obtain H'' h'' where
-      x_hx: "(x, h x) \<in> graph H'' h''"
-    and c'': "graph H'' h'' \<in> c"
-    and ** :
-      "linearform H'' h''"  "H'' \<unlhd> E"  "F \<unlhd> H''"
-      "graph F f \<subseteq> graph H'' h''"  "\<forall>x \<in> H''. h'' x \<le> p x"
-    by (rule some_H'h't [elim_format]) blast
-
-  txt {* Since both @{text h'} and @{text h''} are elements of the chain,
-    @{text h''} is an extension of @{text h'} or vice versa. Thus both
-    @{text x} and @{text y} are contained in the greater
-    one. \label{cases1}*}
-
-  from cM c'' c' have "graph H'' h'' \<subseteq> graph H' h' \<or> graph H' h' \<subseteq> graph H'' h''"
-    (is "?case1 \<or> ?case2") ..
-  then show ?thesis
-  proof
-    assume ?case1
-    have "(x, h x) \<in> graph H'' h''" by fact
-    also have "\<dots> \<subseteq> graph H' h'" by fact
-    finally have xh:"(x, h x) \<in> graph H' h'" .
-    then have "x \<in> H'" ..
-    moreover from y_hy have "y \<in> H'" ..
-    moreover from cM u and c' have "graph H' h' \<subseteq> graph H h"
-      by (simp only: chain_ball_Union_upper)
-    ultimately show ?thesis using * by blast
-  next
-    assume ?case2
-    from x_hx have "x \<in> H''" ..
-    moreover {
-      have "(y, h y) \<in> graph H' h'" by (rule y_hy)
-      also have "\<dots> \<subseteq> graph H'' h''" by fact
-      finally have "(y, h y) \<in> graph H'' h''" .
-    } then have "y \<in> H''" ..
-    moreover from cM u and c'' have "graph H'' h'' \<subseteq> graph H h"
-      by (simp only: chain_ball_Union_upper)
-    ultimately show ?thesis using ** by blast
-  qed
-qed
-
-text {*
-  \medskip The relation induced by the graph of the supremum of a
-  chain @{text c} is definite, i.~e.~t is the graph of a function.
-*}
-
-lemma sup_definite:
-  assumes M_def: "M \<equiv> norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and xy: "(x, y) \<in> \<Union>c"
-    and xz: "(x, z) \<in> \<Union>c"
-  shows "z = y"
-proof -
-  from cM have c: "c \<subseteq> M" ..
-  from xy obtain G1 where xy': "(x, y) \<in> G1" and G1: "G1 \<in> c" ..
-  from xz obtain G2 where xz': "(x, z) \<in> G2" and G2: "G2 \<in> c" ..
-
-  from G1 c have "G1 \<in> M" ..
-  then obtain H1 h1 where G1_rep: "G1 = graph H1 h1"
-    unfolding M_def by blast
-
-  from G2 c have "G2 \<in> M" ..
-  then obtain H2 h2 where G2_rep: "G2 = graph H2 h2"
-    unfolding M_def by blast
-
-  txt {* @{text "G\<^sub>1"} is contained in @{text "G\<^sub>2"}
-    or vice versa, since both @{text "G\<^sub>1"} and @{text
-    "G\<^sub>2"} are members of @{text c}. \label{cases2}*}
-
-  from cM G1 G2 have "G1 \<subseteq> G2 \<or> G2 \<subseteq> G1" (is "?case1 \<or> ?case2") ..
-  then show ?thesis
-  proof
-    assume ?case1
-    with xy' G2_rep have "(x, y) \<in> graph H2 h2" by blast
-    then have "y = h2 x" ..
-    also
-    from xz' G2_rep have "(x, z) \<in> graph H2 h2" by (simp only:)
-    then have "z = h2 x" ..
-    finally show ?thesis .
-  next
-    assume ?case2
-    with xz' G1_rep have "(x, z) \<in> graph H1 h1" by blast
-    then have "z = h1 x" ..
-    also
-    from xy' G1_rep have "(x, y) \<in> graph H1 h1" by (simp only:)
-    then have "y = h1 x" ..
-    finally show ?thesis ..
-  qed
-qed
-
-text {*
-  \medskip The limit function @{text h} is linear. Every element
-  @{text x} in the domain of @{text h} is in the domain of a function
-  @{text h'} in the chain of norm preserving extensions.  Furthermore,
-  @{text h} is an extension of @{text h'} so the function values of
-  @{text x} are identical for @{text h'} and @{text h}.  Finally, the
-  function @{text h'} is linear by construction of @{text M}.
-*}
-
-lemma sup_lf:
-  assumes M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and u: "graph H h = \<Union>c"
-  shows "linearform H h"
-proof
-  fix x y assume x: "x \<in> H" and y: "y \<in> H"
-  with M cM u obtain H' h' where
-        x': "x \<in> H'" and y': "y \<in> H'"
-      and b: "graph H' h' \<subseteq> graph H h"
-      and linearform: "linearform H' h'"
-      and subspace: "H' \<unlhd> E"
-    by (rule some_H'h'2 [elim_format]) blast
-
-  show "h (x + y) = h x + h y"
-  proof -
-    from linearform x' y' have "h' (x + y) = h' x + h' y"
-      by (rule linearform.add)
-    also from b x' have "h' x = h x" ..
-    also from b y' have "h' y = h y" ..
-    also from subspace x' y' have "x + y \<in> H'"
-      by (rule subspace.add_closed)
-    with b have "h' (x + y) = h (x + y)" ..
-    finally show ?thesis .
-  qed
-next
-  fix x a assume x: "x \<in> H"
-  with M cM u obtain H' h' where
-        x': "x \<in> H'"
-      and b: "graph H' h' \<subseteq> graph H h"
-      and linearform: "linearform H' h'"
-      and subspace: "H' \<unlhd> E"
-    by (rule some_H'h' [elim_format]) blast
-
-  show "h (a \<cdot> x) = a * h x"
-  proof -
-    from linearform x' have "h' (a \<cdot> x) = a * h' x"
-      by (rule linearform.mult)
-    also from b x' have "h' x = h x" ..
-    also from subspace x' have "a \<cdot> x \<in> H'"
-      by (rule subspace.mult_closed)
-    with b have "h' (a \<cdot> x) = h (a \<cdot> x)" ..
-    finally show ?thesis .
-  qed
-qed
-
-text {*
-  \medskip The limit of a non-empty chain of norm preserving
-  extensions of @{text f} is an extension of @{text f}, since every
-  element of the chain is an extension of @{text f} and the supremum
-  is an extension for every element of the chain.
-*}
-
-lemma sup_ext:
-  assumes graph: "graph H h = \<Union>c"
-    and M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and ex: "\<exists>x. x \<in> c"
-  shows "graph F f \<subseteq> graph H h"
-proof -
-  from ex obtain x where xc: "x \<in> c" ..
-  from cM have "c \<subseteq> M" ..
-  with xc have "x \<in> M" ..
-  with M have "x \<in> norm_pres_extensions E p F f"
-    by (simp only:)
-  then obtain G g where "x = graph G g" and "graph F f \<subseteq> graph G g" ..
-  then have "graph F f \<subseteq> x" by (simp only:)
-  also from xc have "\<dots> \<subseteq> \<Union>c" by blast
-  also from graph have "\<dots> = graph H h" ..
-  finally show ?thesis .
-qed
-
-text {*
-  \medskip The domain @{text H} of the limit function is a superspace
-  of @{text F}, since @{text F} is a subset of @{text H}. The
-  existence of the @{text 0} element in @{text F} and the closure
-  properties follow from the fact that @{text F} is a vector space.
-*}
-
-lemma sup_supF:
-  assumes graph: "graph H h = \<Union>c"
-    and M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and ex: "\<exists>x. x \<in> c"
-    and FE: "F \<unlhd> E"
-  shows "F \<unlhd> H"
-proof
-  from FE show "F \<noteq> {}" by (rule subspace.non_empty)
-  from graph M cM ex have "graph F f \<subseteq> graph H h" by (rule sup_ext)
-  then show "F \<subseteq> H" ..
-  fix x y assume "x \<in> F" and "y \<in> F"
-  with FE show "x + y \<in> F" by (rule subspace.add_closed)
-next
-  fix x a assume "x \<in> F"
-  with FE show "a \<cdot> x \<in> F" by (rule subspace.mult_closed)
-qed
-
-text {*
-  \medskip The domain @{text H} of the limit function is a subspace of
-  @{text E}.
-*}
-
-lemma sup_subE:
-  assumes graph: "graph H h = \<Union>c"
-    and M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and ex: "\<exists>x. x \<in> c"
-    and FE: "F \<unlhd> E"
-    and E: "vectorspace E"
-  shows "H \<unlhd> E"
-proof
-  show "H \<noteq> {}"
-  proof -
-    from FE E have "0 \<in> F" by (rule subspace.zero)
-    also from graph M cM ex FE have "F \<unlhd> H" by (rule sup_supF)
-    then have "F \<subseteq> H" ..
-    finally show ?thesis by blast
-  qed
-  show "H \<subseteq> E"
-  proof
-    fix x assume "x \<in> H"
-    with M cM graph
-    obtain H' h' where x: "x \<in> H'" and H'E: "H' \<unlhd> E"
-      by (rule some_H'h' [elim_format]) blast
-    from H'E have "H' \<subseteq> E" ..
-    with x show "x \<in> E" ..
-  qed
-  fix x y assume x: "x \<in> H" and y: "y \<in> H"
-  show "x + y \<in> H"
-  proof -
-    from M cM graph x y obtain H' h' where
-          x': "x \<in> H'" and y': "y \<in> H'" and H'E: "H' \<unlhd> E"
-        and graphs: "graph H' h' \<subseteq> graph H h"
-      by (rule some_H'h'2 [elim_format]) blast
-    from H'E x' y' have "x + y \<in> H'"
-      by (rule subspace.add_closed)
-    also from graphs have "H' \<subseteq> H" ..
-    finally show ?thesis .
-  qed
-next
-  fix x a assume x: "x \<in> H"
-  show "a \<cdot> x \<in> H"
-  proof -
-    from M cM graph x
-    obtain H' h' where x': "x \<in> H'" and H'E: "H' \<unlhd> E"
-        and graphs: "graph H' h' \<subseteq> graph H h"
-      by (rule some_H'h' [elim_format]) blast
-    from H'E x' have "a \<cdot> x \<in> H'" by (rule subspace.mult_closed)
-    also from graphs have "H' \<subseteq> H" ..
-    finally show ?thesis .
-  qed
-qed
-
-text {*
-  \medskip The limit function is bounded by the norm @{text p} as
-  well, since all elements in the chain are bounded by @{text p}.
-*}
-
-lemma sup_norm_pres:
-  assumes graph: "graph H h = \<Union>c"
-    and M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-  shows "\<forall>x \<in> H. h x \<le> p x"
-proof
-  fix x assume "x \<in> H"
-  with M cM graph obtain H' h' where x': "x \<in> H'"
-      and graphs: "graph H' h' \<subseteq> graph H h"
-      and a: "\<forall>x \<in> H'. h' x \<le> p x"
-    by (rule some_H'h' [elim_format]) blast
-  from graphs x' have [symmetric]: "h' x = h x" ..
-  also from a x' have "h' x \<le> p x " ..
-  finally show "h x \<le> p x" .
-qed
-
-text {*
-  \medskip The following lemma is a property of linear forms on real
-  vector spaces. It will be used for the lemma @{text abs_HahnBanach}
-  (see page \pageref{abs-HahnBanach}). \label{abs-ineq-iff} For real
-  vector spaces the following inequations are equivalent:
-  \begin{center}
-  \begin{tabular}{lll}
-  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
-  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
-  \end{tabular}
-  \end{center}
-*}
-
-lemma abs_ineq_iff:
-  assumes "subspace H E" and "vectorspace E" and "seminorm E p"
-    and "linearform H h"
-  shows "(\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" (is "?L = ?R")
-proof
-  interpret subspace [H E] by fact
-  interpret vectorspace [E] by fact
-  interpret seminorm [E p] by fact
-  interpret linearform [H h] by fact
-  have H: "vectorspace H" using `vectorspace E` ..
-  {
-    assume l: ?L
-    show ?R
-    proof
-      fix x assume x: "x \<in> H"
-      have "h x \<le> \<bar>h x\<bar>" by arith
-      also from l x have "\<dots> \<le> p x" ..
-      finally show "h x \<le> p x" .
-    qed
-  next
-    assume r: ?R
-    show ?L
-    proof
-      fix x assume x: "x \<in> H"
-      show "\<And>a b :: real. - a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> \<bar>b\<bar> \<le> a"
-        by arith
-      from `linearform H h` and H x
-      have "- h x = h (- x)" by (rule linearform.neg [symmetric])
-      also
-      from H x have "- x \<in> H" by (rule vectorspace.neg_closed)
-      with r have "h (- x) \<le> p (- x)" ..
-      also have "\<dots> = p x"
-	using `seminorm E p` `vectorspace E`
-      proof (rule seminorm.minus)
-        from x show "x \<in> E" ..
-      qed
-      finally have "- h x \<le> p x" .
-      then show "- p x \<le> h x" by simp
-      from r x show "h x \<le> p x" ..
-    qed
-  }
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/Linearform.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,61 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/Linearform.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Linearforms *}
-
-theory Linearform
-imports VectorSpace
-begin
-
-text {*
-  A \emph{linear form} is a function on a vector space into the reals
-  that is additive and multiplicative.
-*}
-
-locale linearform = var V + var f +
-  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
-  assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
-    and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x"
-
-declare linearform.intro [intro?]
-
-lemma (in linearform) neg [iff]:
-  assumes "vectorspace V"
-  shows "x \<in> V \<Longrightarrow> f (- x) = - f x"
-proof -
-  interpret vectorspace [V] by fact
-  assume x: "x \<in> V"
-  then have "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1)
-  also from x have "\<dots> = (- 1) * (f x)" by (rule mult)
-  also from x have "\<dots> = - (f x)" by simp
-  finally show ?thesis .
-qed
-
-lemma (in linearform) diff [iff]:
-  assumes "vectorspace V"
-  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y"
-proof -
-  interpret vectorspace [V] by fact
-  assume x: "x \<in> V" and y: "y \<in> V"
-  then have "x - y = x + - y" by (rule diff_eq1)
-  also have "f \<dots> = f x + f (- y)" by (rule add) (simp_all add: x y)
-  also have "f (- y) = - f y" using `vectorspace V` y by (rule neg)
-  finally show ?thesis by simp
-qed
-
-text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
-
-lemma (in linearform) zero [iff]:
-  assumes "vectorspace V"
-  shows "f 0 = 0"
-proof -
-  interpret vectorspace [V] by fact
-  have "f 0 = f (0 - 0)" by simp
-  also have "\<dots> = f 0 - f 0" using `vectorspace V` by (rule diff) simp_all
-  also have "\<dots> = 0" by simp
-  finally show ?thesis .
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/NormedSpace.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,118 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/NormedSpace.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Normed vector spaces *}
-
-theory NormedSpace
-imports Subspace
-begin
-
-subsection {* Quasinorms *}
-
-text {*
-  A \emph{seminorm} @{text "\<parallel>\<cdot>\<parallel>"} is a function on a real vector space
-  into the reals that has the following properties: it is positive
-  definite, absolute homogenous and subadditive.
-*}
-
-locale norm_syntax =
-  fixes norm :: "'a \<Rightarrow> real"    ("\<parallel>_\<parallel>")
-
-locale seminorm = var V + norm_syntax +
-  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
-  assumes ge_zero [iff?]: "x \<in> V \<Longrightarrow> 0 \<le> \<parallel>x\<parallel>"
-    and abs_homogenous [iff?]: "x \<in> V \<Longrightarrow> \<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>"
-    and subadditive [iff?]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
-
-declare seminorm.intro [intro?]
-
-lemma (in seminorm) diff_subadditive:
-  assumes "vectorspace V"
-  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x - y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
-proof -
-  interpret vectorspace [V] by fact
-  assume x: "x \<in> V" and y: "y \<in> V"
-  then have "x - y = x + - 1 \<cdot> y"
-    by (simp add: diff_eq2 negate_eq2a)
-  also from x y have "\<parallel>\<dots>\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>- 1 \<cdot> y\<parallel>"
-    by (simp add: subadditive)
-  also from y have "\<parallel>- 1 \<cdot> y\<parallel> = \<bar>- 1\<bar> * \<parallel>y\<parallel>"
-    by (rule abs_homogenous)
-  also have "\<dots> = \<parallel>y\<parallel>" by simp
-  finally show ?thesis .
-qed
-
-lemma (in seminorm) minus:
-  assumes "vectorspace V"
-  shows "x \<in> V \<Longrightarrow> \<parallel>- x\<parallel> = \<parallel>x\<parallel>"
-proof -
-  interpret vectorspace [V] by fact
-  assume x: "x \<in> V"
-  then have "- x = - 1 \<cdot> x" by (simp only: negate_eq1)
-  also from x have "\<parallel>\<dots>\<parallel> = \<bar>- 1\<bar> * \<parallel>x\<parallel>"
-    by (rule abs_homogenous)
-  also have "\<dots> = \<parallel>x\<parallel>" by simp
-  finally show ?thesis .
-qed
-
-
-subsection {* Norms *}
-
-text {*
-  A \emph{norm} @{text "\<parallel>\<cdot>\<parallel>"} is a seminorm that maps only the
-  @{text 0} vector to @{text 0}.
-*}
-
-locale norm = seminorm +
-  assumes zero_iff [iff]: "x \<in> V \<Longrightarrow> (\<parallel>x\<parallel> = 0) = (x = 0)"
-
-
-subsection {* Normed vector spaces *}
-
-text {*
-  A vector space together with a norm is called a \emph{normed
-  space}.
-*}
-
-locale normed_vectorspace = vectorspace + norm
-
-declare normed_vectorspace.intro [intro?]
-
-lemma (in normed_vectorspace) gt_zero [intro?]:
-  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> 0 < \<parallel>x\<parallel>"
-proof -
-  assume x: "x \<in> V" and neq: "x \<noteq> 0"
-  from x have "0 \<le> \<parallel>x\<parallel>" ..
-  also have [symmetric]: "\<dots> \<noteq> 0"
-  proof
-    assume "\<parallel>x\<parallel> = 0"
-    with x have "x = 0" by simp
-    with neq show False by contradiction
-  qed
-  finally show ?thesis .
-qed
-
-text {*
-  Any subspace of a normed vector space is again a normed vectorspace.
-*}
-
-lemma subspace_normed_vs [intro?]:
-  fixes F E norm
-  assumes "subspace F E" "normed_vectorspace E norm"
-  shows "normed_vectorspace F norm"
-proof -
-  interpret subspace [F E] by fact
-  interpret normed_vectorspace [E norm] by fact
-  show ?thesis
-  proof
-    show "vectorspace F" by (rule vectorspace) unfold_locales
-  next
-    have "NormedSpace.norm E norm" ..
-    with subset show "NormedSpace.norm F norm"
-      by (simp add: norm_def seminorm_def norm_axioms_def)
-  qed
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/README.html	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,38 +0,0 @@
-<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
-
-<!-- $Id$ -->
-
-<HTML>
-
-<HEAD>
-  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
-  <TITLE>HOL/Real/HahnBanach/README</TITLE>
-</HEAD>
-
-<BODY>
-
-<H3>The Hahn-Banach Theorem for Real Vector Spaces (Isabelle/Isar)</H3>
-
-Author: Gertrud Bauer, Technische Universit&auml;t M&uuml;nchen<P>
-
-This directory contains the proof of the Hahn-Banach theorem for real vectorspaces,
-following H. Heuser, Funktionalanalysis, p. 228 -232.
-The Hahn-Banach theorem is one of the fundamental theorems of functioal analysis.
-It is a conclusion of Zorn's lemma.<P>
-
-Two different formaulations of the theorem are presented, one for general real vectorspaces
-and its application to normed vectorspaces. <P>
-
-The theorem says, that every continous linearform, defined on arbitrary subspaces
-(not only one-dimensional subspaces), can be extended to a continous linearform on
-the whole vectorspace.
-
-
-<HR>
-
-<ADDRESS>
-<A NAME="bauerg@in.tum.de" HREF="mailto:bauerg@in.tum.de">bauerg@in.tum.de</A>
-</ADDRESS>
-
-</BODY>
-</HTML>
--- a/src/HOL/Real/HahnBanach/ROOT.ML	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/ROOT.ML
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-
-The Hahn-Banach theorem for real vector spaces (Isabelle/Isar).
-*)
-
-time_use_thy "HahnBanach";
--- a/src/HOL/Real/HahnBanach/Subspace.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,514 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/Subspace.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Subspaces *}
-
-theory Subspace
-imports VectorSpace
-begin
-
-subsection {* Definition *}
-
-text {*
-  A non-empty subset @{text U} of a vector space @{text V} is a
-  \emph{subspace} of @{text V}, iff @{text U} is closed under addition
-  and scalar multiplication.
-*}
-
-locale subspace = var U + var V +
-  constrains U :: "'a\<Colon>{minus, plus, zero, uminus} set"
-  assumes non_empty [iff, intro]: "U \<noteq> {}"
-    and subset [iff]: "U \<subseteq> V"
-    and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U"
-    and mult_closed [iff]: "x \<in> U \<Longrightarrow> a \<cdot> x \<in> U"
-
-notation (symbols)
-  subspace  (infix "\<unlhd>" 50)
-
-declare vectorspace.intro [intro?] subspace.intro [intro?]
-
-lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V"
-  by (rule subspace.subset)
-
-lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V"
-  using subset by blast
-
-lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V"
-  by (rule subspace.subsetD)
-
-lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V"
-  by (rule subspace.subsetD)
-
-lemma (in subspace) diff_closed [iff]:
-  assumes "vectorspace V"
-  assumes x: "x \<in> U" and y: "y \<in> U"
-  shows "x - y \<in> U"
-proof -
-  interpret vectorspace [V] by fact
-  from x y show ?thesis by (simp add: diff_eq1 negate_eq1)
-qed
-
-text {*
-  \medskip Similar as for linear spaces, the existence of the zero
-  element in every subspace follows from the non-emptiness of the
-  carrier set and by vector space laws.
-*}
-
-lemma (in subspace) zero [intro]:
-  assumes "vectorspace V"
-  shows "0 \<in> U"
-proof -
-  interpret vectorspace [V] by fact
-  have "U \<noteq> {}" by (rule U_V.non_empty)
-  then obtain x where x: "x \<in> U" by blast
-  then have "x \<in> V" .. then have "0 = x - x" by simp
-  also from `vectorspace V` x x have "\<dots> \<in> U" by (rule U_V.diff_closed)
-  finally show ?thesis .
-qed
-
-lemma (in subspace) neg_closed [iff]:
-  assumes "vectorspace V"
-  assumes x: "x \<in> U"
-  shows "- x \<in> U"
-proof -
-  interpret vectorspace [V] by fact
-  from x show ?thesis by (simp add: negate_eq1)
-qed
-
-text {* \medskip Further derived laws: every subspace is a vector space. *}
-
-lemma (in subspace) vectorspace [iff]:
-  assumes "vectorspace V"
-  shows "vectorspace U"
-proof -
-  interpret vectorspace [V] by fact
-  show ?thesis
-  proof
-    show "U \<noteq> {}" ..
-    fix x y z assume x: "x \<in> U" and y: "y \<in> U" and z: "z \<in> U"
-    fix a b :: real
-    from x y show "x + y \<in> U" by simp
-    from x show "a \<cdot> x \<in> U" by simp
-    from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac)
-    from x y show "x + y = y + x" by (simp add: add_ac)
-    from x show "x - x = 0" by simp
-    from x show "0 + x = x" by simp
-    from x y show "a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" by (simp add: distrib)
-    from x show "(a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" by (simp add: distrib)
-    from x show "(a * b) \<cdot> x = a \<cdot> b \<cdot> x" by (simp add: mult_assoc)
-    from x show "1 \<cdot> x = x" by simp
-    from x show "- x = - 1 \<cdot> x" by (simp add: negate_eq1)
-    from x y show "x - y = x + - y" by (simp add: diff_eq1)
-  qed
-qed
-
-
-text {* The subspace relation is reflexive. *}
-
-lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
-proof
-  show "V \<noteq> {}" ..
-  show "V \<subseteq> V" ..
-  fix x y assume x: "x \<in> V" and y: "y \<in> V"
-  fix a :: real
-  from x y show "x + y \<in> V" by simp
-  from x show "a \<cdot> x \<in> V" by simp
-qed
-
-text {* The subspace relation is transitive. *}
-
-lemma (in vectorspace) subspace_trans [trans]:
-  "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W"
-proof
-  assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W"
-  from uv show "U \<noteq> {}" by (rule subspace.non_empty)
-  show "U \<subseteq> W"
-  proof -
-    from uv have "U \<subseteq> V" by (rule subspace.subset)
-    also from vw have "V \<subseteq> W" by (rule subspace.subset)
-    finally show ?thesis .
-  qed
-  fix x y assume x: "x \<in> U" and y: "y \<in> U"
-  from uv and x y show "x + y \<in> U" by (rule subspace.add_closed)
-  from uv and x show "\<And>a. a \<cdot> x \<in> U" by (rule subspace.mult_closed)
-qed
-
-
-subsection {* Linear closure *}
-
-text {*
-  The \emph{linear closure} of a vector @{text x} is the set of all
-  scalar multiples of @{text x}.
-*}
-
-definition
-  lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set" where
-  "lin x = {a \<cdot> x | a. True}"
-
-lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
-  unfolding lin_def by blast
-
-lemma linI' [iff]: "a \<cdot> x \<in> lin x"
-  unfolding lin_def by blast
-
-lemma linE [elim]: "x \<in> lin v \<Longrightarrow> (\<And>a::real. x = a \<cdot> v \<Longrightarrow> C) \<Longrightarrow> C"
-  unfolding lin_def by blast
-
-
-text {* Every vector is contained in its linear closure. *}
-
-lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x"
-proof -
-  assume "x \<in> V"
-  then have "x = 1 \<cdot> x" by simp
-  also have "\<dots> \<in> lin x" ..
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x"
-proof
-  assume "x \<in> V"
-  then show "0 = 0 \<cdot> x" by simp
-qed
-
-text {* Any linear closure is a subspace. *}
-
-lemma (in vectorspace) lin_subspace [intro]:
-  "x \<in> V \<Longrightarrow> lin x \<unlhd> V"
-proof
-  assume x: "x \<in> V"
-  then show "lin x \<noteq> {}" by (auto simp add: x_lin_x)
-  show "lin x \<subseteq> V"
-  proof
-    fix x' assume "x' \<in> lin x"
-    then obtain a where "x' = a \<cdot> x" ..
-    with x show "x' \<in> V" by simp
-  qed
-  fix x' x'' assume x': "x' \<in> lin x" and x'': "x'' \<in> lin x"
-  show "x' + x'' \<in> lin x"
-  proof -
-    from x' obtain a' where "x' = a' \<cdot> x" ..
-    moreover from x'' obtain a'' where "x'' = a'' \<cdot> x" ..
-    ultimately have "x' + x'' = (a' + a'') \<cdot> x"
-      using x by (simp add: distrib)
-    also have "\<dots> \<in> lin x" ..
-    finally show ?thesis .
-  qed
-  fix a :: real
-  show "a \<cdot> x' \<in> lin x"
-  proof -
-    from x' obtain a' where "x' = a' \<cdot> x" ..
-    with x have "a \<cdot> x' = (a * a') \<cdot> x" by (simp add: mult_assoc)
-    also have "\<dots> \<in> lin x" ..
-    finally show ?thesis .
-  qed
-qed
-
-
-text {* Any linear closure is a vector space. *}
-
-lemma (in vectorspace) lin_vectorspace [intro]:
-  assumes "x \<in> V"
-  shows "vectorspace (lin x)"
-proof -
-  from `x \<in> V` have "subspace (lin x) V"
-    by (rule lin_subspace)
-  from this and vectorspace_axioms show ?thesis
-    by (rule subspace.vectorspace)
-qed
-
-
-subsection {* Sum of two vectorspaces *}
-
-text {*
-  The \emph{sum} of two vectorspaces @{text U} and @{text V} is the
-  set of all sums of elements from @{text U} and @{text V}.
-*}
-
-instantiation "fun" :: (type, type) plus
-begin
-
-definition
-  sum_def: "plus_fun U V = {u + v | u v. u \<in> U \<and> v \<in> V}"  (* FIXME not fully general!? *)
-
-instance ..
-
-end
-
-lemma sumE [elim]:
-    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
-  unfolding sum_def by blast
-
-lemma sumI [intro]:
-    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
-  unfolding sum_def by blast
-
-lemma sumI' [intro]:
-    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
-  unfolding sum_def by blast
-
-text {* @{text U} is a subspace of @{text "U + V"}. *}
-
-lemma subspace_sum1 [iff]:
-  assumes "vectorspace U" "vectorspace V"
-  shows "U \<unlhd> U + V"
-proof -
-  interpret vectorspace [U] by fact
-  interpret vectorspace [V] by fact
-  show ?thesis
-  proof
-    show "U \<noteq> {}" ..
-    show "U \<subseteq> U + V"
-    proof
-      fix x assume x: "x \<in> U"
-      moreover have "0 \<in> V" ..
-      ultimately have "x + 0 \<in> U + V" ..
-      with x show "x \<in> U + V" by simp
-    qed
-    fix x y assume x: "x \<in> U" and "y \<in> U"
-    then show "x + y \<in> U" by simp
-    from x show "\<And>a. a \<cdot> x \<in> U" by simp
-  qed
-qed
-
-text {* The sum of two subspaces is again a subspace. *}
-
-lemma sum_subspace [intro?]:
-  assumes "subspace U E" "vectorspace E" "subspace V E"
-  shows "U + V \<unlhd> E"
-proof -
-  interpret subspace [U E] by fact
-  interpret vectorspace [E] by fact
-  interpret subspace [V E] by fact
-  show ?thesis
-  proof
-    have "0 \<in> U + V"
-    proof
-      show "0 \<in> U" using `vectorspace E` ..
-      show "0 \<in> V" using `vectorspace E` ..
-      show "(0::'a) = 0 + 0" by simp
-    qed
-    then show "U + V \<noteq> {}" by blast
-    show "U + V \<subseteq> E"
-    proof
-      fix x assume "x \<in> U + V"
-      then obtain u v where "x = u + v" and
-	"u \<in> U" and "v \<in> V" ..
-      then show "x \<in> E" by simp
-    qed
-    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
-    show "x + y \<in> U + V"
-    proof -
-      from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
-      moreover
-      from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" ..
-      ultimately
-      have "ux + uy \<in> U"
-	and "vx + vy \<in> V"
-	and "x + y = (ux + uy) + (vx + vy)"
-	using x y by (simp_all add: add_ac)
-      then show ?thesis ..
-    qed
-    fix a show "a \<cdot> x \<in> U + V"
-    proof -
-      from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
-      then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
-	and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
-      then show ?thesis ..
-    qed
-  qed
-qed
-
-text{* The sum of two subspaces is a vectorspace. *}
-
-lemma sum_vs [intro?]:
-    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
-  by (rule subspace.vectorspace) (rule sum_subspace)
-
-
-subsection {* Direct sums *}
-
-text {*
-  The sum of @{text U} and @{text V} is called \emph{direct}, iff the
-  zero element is the only common element of @{text U} and @{text
-  V}. For every element @{text x} of the direct sum of @{text U} and
-  @{text V} the decomposition in @{text "x = u + v"} with
-  @{text "u \<in> U"} and @{text "v \<in> V"} is unique.
-*}
-
-lemma decomp:
-  assumes "vectorspace E" "subspace U E" "subspace V E"
-  assumes direct: "U \<inter> V = {0}"
-    and u1: "u1 \<in> U" and u2: "u2 \<in> U"
-    and v1: "v1 \<in> V" and v2: "v2 \<in> V"
-    and sum: "u1 + v1 = u2 + v2"
-  shows "u1 = u2 \<and> v1 = v2"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [U E] by fact
-  interpret subspace [V E] by fact
-  show ?thesis
-  proof
-    have U: "vectorspace U"  (* FIXME: use interpret *)
-      using `subspace U E` `vectorspace E` by (rule subspace.vectorspace)
-    have V: "vectorspace V"
-      using `subspace V E` `vectorspace E` by (rule subspace.vectorspace)
-    from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1"
-      by (simp add: add_diff_swap)
-    from u1 u2 have u: "u1 - u2 \<in> U"
-      by (rule vectorspace.diff_closed [OF U])
-    with eq have v': "v2 - v1 \<in> U" by (simp only:)
-    from v2 v1 have v: "v2 - v1 \<in> V"
-      by (rule vectorspace.diff_closed [OF V])
-    with eq have u': " u1 - u2 \<in> V" by (simp only:)
-    
-    show "u1 = u2"
-    proof (rule add_minus_eq)
-      from u1 show "u1 \<in> E" ..
-      from u2 show "u2 \<in> E" ..
-      from u u' and direct show "u1 - u2 = 0" by blast
-    qed
-    show "v1 = v2"
-    proof (rule add_minus_eq [symmetric])
-      from v1 show "v1 \<in> E" ..
-      from v2 show "v2 \<in> E" ..
-      from v v' and direct show "v2 - v1 = 0" by blast
-    qed
-  qed
-qed
-
-text {*
-  An application of the previous lemma will be used in the proof of
-  the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
-  element @{text "y + a \<cdot> x\<^sub>0"} of the direct sum of a
-  vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"}
-  the components @{text "y \<in> H"} and @{text a} are uniquely
-  determined.
-*}
-
-lemma decomp_H':
-  assumes "vectorspace E" "subspace H E"
-  assumes y1: "y1 \<in> H" and y2: "y2 \<in> H"
-    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
-    and eq: "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'"
-  shows "y1 = y2 \<and> a1 = a2"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [H E] by fact
-  show ?thesis
-  proof
-    have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'"
-    proof (rule decomp)
-      show "a1 \<cdot> x' \<in> lin x'" ..
-      show "a2 \<cdot> x' \<in> lin x'" ..
-      show "H \<inter> lin x' = {0}"
-      proof
-	show "H \<inter> lin x' \<subseteq> {0}"
-	proof
-          fix x assume x: "x \<in> H \<inter> lin x'"
-          then obtain a where xx': "x = a \<cdot> x'"
-            by blast
-          have "x = 0"
-          proof cases
-            assume "a = 0"
-            with xx' and x' show ?thesis by simp
-          next
-            assume a: "a \<noteq> 0"
-            from x have "x \<in> H" ..
-            with xx' have "inverse a \<cdot> a \<cdot> x' \<in> H" by simp
-            with a and x' have "x' \<in> H" by (simp add: mult_assoc2)
-            with `x' \<notin> H` show ?thesis by contradiction
-          qed
-          then show "x \<in> {0}" ..
-	qed
-	show "{0} \<subseteq> H \<inter> lin x'"
-	proof -
-          have "0 \<in> H" using `vectorspace E` ..
-          moreover have "0 \<in> lin x'" using `x' \<in> E` ..
-          ultimately show ?thesis by blast
-	qed
-      qed
-      show "lin x' \<unlhd> E" using `x' \<in> E` ..
-    qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq)
-    then show "y1 = y2" ..
-    from c have "a1 \<cdot> x' = a2 \<cdot> x'" ..
-    with x' show "a1 = a2" by (simp add: mult_right_cancel)
-  qed
-qed
-
-text {*
-  Since for any element @{text "y + a \<cdot> x'"} of the direct sum of a
-  vectorspace @{text H} and the linear closure of @{text x'} the
-  components @{text "y \<in> H"} and @{text a} are unique, it follows from
-  @{text "y \<in> H"} that @{text "a = 0"}.
-*}
-
-lemma decomp_H'_H:
-  assumes "vectorspace E" "subspace H E"
-  assumes t: "t \<in> H"
-    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
-  shows "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [H E] by fact
-  show ?thesis
-  proof (rule, simp_all only: split_paired_all split_conv)
-    from t x' show "t = t + 0 \<cdot> x' \<and> t \<in> H" by simp
-    fix y and a assume ya: "t = y + a \<cdot> x' \<and> y \<in> H"
-    have "y = t \<and> a = 0"
-    proof (rule decomp_H')
-      from ya x' show "y + a \<cdot> x' = t + 0 \<cdot> x'" by simp
-      from ya show "y \<in> H" ..
-    qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+)
-    with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp
-  qed
-qed
-
-text {*
-  The components @{text "y \<in> H"} and @{text a} in @{text "y + a \<cdot> x'"}
-  are unique, so the function @{text h'} defined by
-  @{text "h' (y + a \<cdot> x') = h y + a \<cdot> \<xi>"} is definite.
-*}
-
-lemma h'_definite:
-  fixes H
-  assumes h'_def:
-    "h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
-                in (h y) + a * xi)"
-    and x: "x = y + a \<cdot> x'"
-  assumes "vectorspace E" "subspace H E"
-  assumes y: "y \<in> H"
-    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
-  shows "h' x = h y + a * xi"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [H E] by fact
-  from x y x' have "x \<in> H + lin x'" by auto
-  have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p")
-  proof (rule ex_ex1I)
-    from x y show "\<exists>p. ?P p" by blast
-    fix p q assume p: "?P p" and q: "?P q"
-    show "p = q"
-    proof -
-      from p have xp: "x = fst p + snd p \<cdot> x' \<and> fst p \<in> H"
-        by (cases p) simp
-      from q have xq: "x = fst q + snd q \<cdot> x' \<and> fst q \<in> H"
-        by (cases q) simp
-      have "fst p = fst q \<and> snd p = snd q"
-      proof (rule decomp_H')
-        from xp show "fst p \<in> H" ..
-        from xq show "fst q \<in> H" ..
-        from xp and xq show "fst p + snd p \<cdot> x' = fst q + snd q \<cdot> x'"
-          by simp
-      qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+)
-      then show ?thesis by (cases p, cases q) simp
-    qed
-  qed
-  then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)"
-    by (rule some1_equality) (simp add: x y)
-  with h'_def show "h' x = h y + a * xi" by (simp add: Let_def)
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/VectorSpace.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,417 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/VectorSpace.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Vector spaces *}
-
-theory VectorSpace
-imports Real Bounds Zorn
-begin
-
-subsection {* Signature *}
-
-text {*
-  For the definition of real vector spaces a type @{typ 'a} of the
-  sort @{text "{plus, minus, zero}"} is considered, on which a real
-  scalar multiplication @{text \<cdot>} is declared.
-*}
-
-consts
-  prod  :: "real \<Rightarrow> 'a::{plus, minus, zero} \<Rightarrow> 'a"     (infixr "'(*')" 70)
-
-notation (xsymbols)
-  prod  (infixr "\<cdot>" 70)
-notation (HTML output)
-  prod  (infixr "\<cdot>" 70)
-
-
-subsection {* Vector space laws *}
-
-text {*
-  A \emph{vector space} is a non-empty set @{text V} of elements from
-  @{typ 'a} with the following vector space laws: The set @{text V} is
-  closed under addition and scalar multiplication, addition is
-  associative and commutative; @{text "- x"} is the inverse of @{text
-  x} w.~r.~t.~addition and @{text 0} is the neutral element of
-  addition.  Addition and multiplication are distributive; scalar
-  multiplication is associative and the real number @{text "1"} is
-  the neutral element of scalar multiplication.
-*}
-
-locale vectorspace = var V +
-  assumes non_empty [iff, intro?]: "V \<noteq> {}"
-    and add_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y \<in> V"
-    and mult_closed [iff]: "x \<in> V \<Longrightarrow> a \<cdot> x \<in> V"
-    and add_assoc: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y) + z = x + (y + z)"
-    and add_commute: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = y + x"
-    and diff_self [simp]: "x \<in> V \<Longrightarrow> x - x = 0"
-    and add_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 + x = x"
-    and add_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y"
-    and add_mult_distrib2: "x \<in> V \<Longrightarrow> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x"
-    and mult_assoc: "x \<in> V \<Longrightarrow> (a * b) \<cdot> x = a \<cdot> (b \<cdot> x)"
-    and mult_1 [simp]: "x \<in> V \<Longrightarrow> 1 \<cdot> x = x"
-    and negate_eq1: "x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x"
-    and diff_eq1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = x + - y"
-
-lemma (in vectorspace) negate_eq2: "x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x"
-  by (rule negate_eq1 [symmetric])
-
-lemma (in vectorspace) negate_eq2a: "x \<in> V \<Longrightarrow> -1 \<cdot> x = - x"
-  by (simp add: negate_eq1)
-
-lemma (in vectorspace) diff_eq2: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + - y = x - y"
-  by (rule diff_eq1 [symmetric])
-
-lemma (in vectorspace) diff_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y \<in> V"
-  by (simp add: diff_eq1 negate_eq1)
-
-lemma (in vectorspace) neg_closed [iff]: "x \<in> V \<Longrightarrow> - x \<in> V"
-  by (simp add: negate_eq1)
-
-lemma (in vectorspace) add_left_commute:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> x + (y + z) = y + (x + z)"
-proof -
-  assume xyz: "x \<in> V"  "y \<in> V"  "z \<in> V"
-  then have "x + (y + z) = (x + y) + z"
-    by (simp only: add_assoc)
-  also from xyz have "\<dots> = (y + x) + z" by (simp only: add_commute)
-  also from xyz have "\<dots> = y + (x + z)" by (simp only: add_assoc)
-  finally show ?thesis .
-qed
-
-theorems (in vectorspace) add_ac =
-  add_assoc add_commute add_left_commute
-
-
-text {* The existence of the zero element of a vector space
-  follows from the non-emptiness of carrier set. *}
-
-lemma (in vectorspace) zero [iff]: "0 \<in> V"
-proof -
-  from non_empty obtain x where x: "x \<in> V" by blast
-  then have "0 = x - x" by (rule diff_self [symmetric])
-  also from x x have "\<dots> \<in> V" by (rule diff_closed)
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) add_zero_right [simp]:
-  "x \<in> V \<Longrightarrow>  x + 0 = x"
-proof -
-  assume x: "x \<in> V"
-  from this and zero have "x + 0 = 0 + x" by (rule add_commute)
-  also from x have "\<dots> = x" by (rule add_zero_left)
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) mult_assoc2:
-    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = (a * b) \<cdot> x"
-  by (simp only: mult_assoc)
-
-lemma (in vectorspace) diff_mult_distrib1:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x - y) = a \<cdot> x - a \<cdot> y"
-  by (simp add: diff_eq1 negate_eq1 add_mult_distrib1 mult_assoc2)
-
-lemma (in vectorspace) diff_mult_distrib2:
-  "x \<in> V \<Longrightarrow> (a - b) \<cdot> x = a \<cdot> x - (b \<cdot> x)"
-proof -
-  assume x: "x \<in> V"
-  have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
-    by (simp add: real_diff_def)
-  also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
-    by (rule add_mult_distrib2)
-  also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
-    by (simp add: negate_eq1 mult_assoc2)
-  also from x have "\<dots> = a \<cdot> x - (b \<cdot> x)"
-    by (simp add: diff_eq1)
-  finally show ?thesis .
-qed
-
-lemmas (in vectorspace) distrib =
-  add_mult_distrib1 add_mult_distrib2
-  diff_mult_distrib1 diff_mult_distrib2
-
-
-text {* \medskip Further derived laws: *}
-
-lemma (in vectorspace) mult_zero_left [simp]:
-  "x \<in> V \<Longrightarrow> 0 \<cdot> x = 0"
-proof -
-  assume x: "x \<in> V"
-  have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp
-  also have "\<dots> = (1 + - 1) \<cdot> x" by simp
-  also from x have "\<dots> =  1 \<cdot> x + (- 1) \<cdot> x"
-    by (rule add_mult_distrib2)
-  also from x have "\<dots> = x + (- 1) \<cdot> x" by simp
-  also from x have "\<dots> = x + - x" by (simp add: negate_eq2a)
-  also from x have "\<dots> = x - x" by (simp add: diff_eq2)
-  also from x have "\<dots> = 0" by simp
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) mult_zero_right [simp]:
-  "a \<cdot> 0 = (0::'a)"
-proof -
-  have "a \<cdot> 0 = a \<cdot> (0 - (0::'a))" by simp
-  also have "\<dots> =  a \<cdot> 0 - a \<cdot> 0"
-    by (rule diff_mult_distrib1) simp_all
-  also have "\<dots> = 0" by simp
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) minus_mult_cancel [simp]:
-    "x \<in> V \<Longrightarrow> (- a) \<cdot> - x = a \<cdot> x"
-  by (simp add: negate_eq1 mult_assoc2)
-
-lemma (in vectorspace) add_minus_left_eq_diff:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + y = y - x"
-proof -
-  assume xy: "x \<in> V"  "y \<in> V"
-  then have "- x + y = y + - x" by (simp add: add_commute)
-  also from xy have "\<dots> = y - x" by (simp add: diff_eq1)
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) add_minus [simp]:
-    "x \<in> V \<Longrightarrow> x + - x = 0"
-  by (simp add: diff_eq2)
-
-lemma (in vectorspace) add_minus_left [simp]:
-    "x \<in> V \<Longrightarrow> - x + x = 0"
-  by (simp add: diff_eq2 add_commute)
-
-lemma (in vectorspace) minus_minus [simp]:
-    "x \<in> V \<Longrightarrow> - (- x) = x"
-  by (simp add: negate_eq1 mult_assoc2)
-
-lemma (in vectorspace) minus_zero [simp]:
-    "- (0::'a) = 0"
-  by (simp add: negate_eq1)
-
-lemma (in vectorspace) minus_zero_iff [simp]:
-  "x \<in> V \<Longrightarrow> (- x = 0) = (x = 0)"
-proof
-  assume x: "x \<in> V"
-  {
-    from x have "x = - (- x)" by (simp add: minus_minus)
-    also assume "- x = 0"
-    also have "- \<dots> = 0" by (rule minus_zero)
-    finally show "x = 0" .
-  next
-    assume "x = 0"
-    then show "- x = 0" by simp
-  }
-qed
-
-lemma (in vectorspace) add_minus_cancel [simp]:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + (- x + y) = y"
-  by (simp add: add_assoc [symmetric] del: add_commute)
-
-lemma (in vectorspace) minus_add_cancel [simp]:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + (x + y) = y"
-  by (simp add: add_assoc [symmetric] del: add_commute)
-
-lemma (in vectorspace) minus_add_distrib [simp]:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - (x + y) = - x + - y"
-  by (simp add: negate_eq1 add_mult_distrib1)
-
-lemma (in vectorspace) diff_zero [simp]:
-    "x \<in> V \<Longrightarrow> x - 0 = x"
-  by (simp add: diff_eq1)
-
-lemma (in vectorspace) diff_zero_right [simp]:
-    "x \<in> V \<Longrightarrow> 0 - x = - x"
-  by (simp add: diff_eq1)
-
-lemma (in vectorspace) add_left_cancel:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y = x + z) = (y = z)"
-proof
-  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
-  {
-    from y have "y = 0 + y" by simp
-    also from x y have "\<dots> = (- x + x) + y" by simp
-    also from x y have "\<dots> = - x + (x + y)"
-      by (simp add: add_assoc neg_closed)
-    also assume "x + y = x + z"
-    also from x z have "- x + (x + z) = - x + x + z"
-      by (simp add: add_assoc [symmetric] neg_closed)
-    also from x z have "\<dots> = z" by simp
-    finally show "y = z" .
-  next
-    assume "y = z"
-    then show "x + y = x + z" by (simp only:)
-  }
-qed
-
-lemma (in vectorspace) add_right_cancel:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (y + x = z + x) = (y = z)"
-  by (simp only: add_commute add_left_cancel)
-
-lemma (in vectorspace) add_assoc_cong:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x' \<in> V \<Longrightarrow> y' \<in> V \<Longrightarrow> z \<in> V
-    \<Longrightarrow> x + y = x' + y' \<Longrightarrow> x + (y + z) = x' + (y' + z)"
-  by (simp only: add_assoc [symmetric])
-
-lemma (in vectorspace) mult_left_commute:
-    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = b \<cdot> a \<cdot> x"
-  by (simp add: real_mult_commute mult_assoc2)
-
-lemma (in vectorspace) mult_zero_uniq:
-  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> a \<cdot> x = 0 \<Longrightarrow> a = 0"
-proof (rule classical)
-  assume a: "a \<noteq> 0"
-  assume x: "x \<in> V"  "x \<noteq> 0" and ax: "a \<cdot> x = 0"
-  from x a have "x = (inverse a * a) \<cdot> x" by simp
-  also from `x \<in> V` have "\<dots> = inverse a \<cdot> (a \<cdot> x)" by (rule mult_assoc)
-  also from ax have "\<dots> = inverse a \<cdot> 0" by simp
-  also have "\<dots> = 0" by simp
-  finally have "x = 0" .
-  with `x \<noteq> 0` show "a = 0" by contradiction
-qed
-
-lemma (in vectorspace) mult_left_cancel:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (a \<cdot> x = a \<cdot> y) = (x = y)"
-proof
-  assume x: "x \<in> V" and y: "y \<in> V" and a: "a \<noteq> 0"
-  from x have "x = 1 \<cdot> x" by simp
-  also from a have "\<dots> = (inverse a * a) \<cdot> x" by simp
-  also from x have "\<dots> = inverse a \<cdot> (a \<cdot> x)"
-    by (simp only: mult_assoc)
-  also assume "a \<cdot> x = a \<cdot> y"
-  also from a y have "inverse a \<cdot> \<dots> = y"
-    by (simp add: mult_assoc2)
-  finally show "x = y" .
-next
-  assume "x = y"
-  then show "a \<cdot> x = a \<cdot> y" by (simp only:)
-qed
-
-lemma (in vectorspace) mult_right_cancel:
-  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (a \<cdot> x = b \<cdot> x) = (a = b)"
-proof
-  assume x: "x \<in> V" and neq: "x \<noteq> 0"
-  {
-    from x have "(a - b) \<cdot> x = a \<cdot> x - b \<cdot> x"
-      by (simp add: diff_mult_distrib2)
-    also assume "a \<cdot> x = b \<cdot> x"
-    with x have "a \<cdot> x - b \<cdot> x = 0" by simp
-    finally have "(a - b) \<cdot> x = 0" .
-    with x neq have "a - b = 0" by (rule mult_zero_uniq)
-    then show "a = b" by simp
-  next
-    assume "a = b"
-    then show "a \<cdot> x = b \<cdot> x" by (simp only:)
-  }
-qed
-
-lemma (in vectorspace) eq_diff_eq:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x = z - y) = (x + y = z)"
-proof
-  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
-  {
-    assume "x = z - y"
-    then have "x + y = z - y + y" by simp
-    also from y z have "\<dots> = z + - y + y"
-      by (simp add: diff_eq1)
-    also have "\<dots> = z + (- y + y)"
-      by (rule add_assoc) (simp_all add: y z)
-    also from y z have "\<dots> = z + 0"
-      by (simp only: add_minus_left)
-    also from z have "\<dots> = z"
-      by (simp only: add_zero_right)
-    finally show "x + y = z" .
-  next
-    assume "x + y = z"
-    then have "z - y = (x + y) - y" by simp
-    also from x y have "\<dots> = x + y + - y"
-      by (simp add: diff_eq1)
-    also have "\<dots> = x + (y + - y)"
-      by (rule add_assoc) (simp_all add: x y)
-    also from x y have "\<dots> = x" by simp
-    finally show "x = z - y" ..
-  }
-qed
-
-lemma (in vectorspace) add_minus_eq_minus:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = 0 \<Longrightarrow> x = - y"
-proof -
-  assume x: "x \<in> V" and y: "y \<in> V"
-  from x y have "x = (- y + y) + x" by simp
-  also from x y have "\<dots> = - y + (x + y)" by (simp add: add_ac)
-  also assume "x + y = 0"
-  also from y have "- y + 0 = - y" by simp
-  finally show "x = - y" .
-qed
-
-lemma (in vectorspace) add_minus_eq:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = 0 \<Longrightarrow> x = y"
-proof -
-  assume x: "x \<in> V" and y: "y \<in> V"
-  assume "x - y = 0"
-  with x y have eq: "x + - y = 0" by (simp add: diff_eq1)
-  with _ _ have "x = - (- y)"
-    by (rule add_minus_eq_minus) (simp_all add: x y)
-  with x y show "x = y" by simp
-qed
-
-lemma (in vectorspace) add_diff_swap:
-  "a \<in> V \<Longrightarrow> b \<in> V \<Longrightarrow> c \<in> V \<Longrightarrow> d \<in> V \<Longrightarrow> a + b = c + d
-    \<Longrightarrow> a - c = d - b"
-proof -
-  assume vs: "a \<in> V"  "b \<in> V"  "c \<in> V"  "d \<in> V"
-    and eq: "a + b = c + d"
-  then have "- c + (a + b) = - c + (c + d)"
-    by (simp add: add_left_cancel)
-  also have "\<dots> = d" using `c \<in> V` `d \<in> V` by (rule minus_add_cancel)
-  finally have eq: "- c + (a + b) = d" .
-  from vs have "a - c = (- c + (a + b)) + - b"
-    by (simp add: add_ac diff_eq1)
-  also from vs eq have "\<dots>  = d + - b"
-    by (simp add: add_right_cancel)
-  also from vs have "\<dots> = d - b" by (simp add: diff_eq2)
-  finally show "a - c = d - b" .
-qed
-
-lemma (in vectorspace) vs_add_cancel_21:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> u \<in> V
-    \<Longrightarrow> (x + (y + z) = y + u) = (x + z = u)"
-proof
-  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"  "u \<in> V"
-  {
-    from vs have "x + z = - y + y + (x + z)" by simp
-    also have "\<dots> = - y + (y + (x + z))"
-      by (rule add_assoc) (simp_all add: vs)
-    also from vs have "y + (x + z) = x + (y + z)"
-      by (simp add: add_ac)
-    also assume "x + (y + z) = y + u"
-    also from vs have "- y + (y + u) = u" by simp
-    finally show "x + z = u" .
-  next
-    assume "x + z = u"
-    with vs show "x + (y + z) = y + u"
-      by (simp only: add_left_commute [of x])
-  }
-qed
-
-lemma (in vectorspace) add_cancel_end:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + (y + z) = y) = (x = - z)"
-proof
-  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"
-  {
-    assume "x + (y + z) = y"
-    with vs have "(x + z) + y = 0 + y"
-      by (simp add: add_ac)
-    with vs have "x + z = 0"
-      by (simp only: add_right_cancel add_closed zero)
-    with vs show "x = - z" by (simp add: add_minus_eq_minus)
-  next
-    assume eq: "x = - z"
-    then have "x + (y + z) = - z + (y + z)" by simp
-    also have "\<dots> = y + (- z + z)"
-      by (rule add_left_commute) (simp_all add: vs)
-    also from vs have "\<dots> = y"  by simp
-    finally show "x + (y + z) = y" .
-  }
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/ZornLemma.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,58 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/ZornLemma.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Zorn's Lemma *}
-
-theory ZornLemma
-imports Zorn
-begin
-
-text {*
-  Zorn's Lemmas states: if every linear ordered subset of an ordered
-  set @{text S} has an upper bound in @{text S}, then there exists a
-  maximal element in @{text S}.  In our application, @{text S} is a
-  set of sets ordered by set inclusion. Since the union of a chain of
-  sets is an upper bound for all elements of the chain, the conditions
-  of Zorn's lemma can be modified: if @{text S} is non-empty, it
-  suffices to show that for every non-empty chain @{text c} in @{text
-  S} the union of @{text c} also lies in @{text S}.
-*}
-
-theorem Zorn's_Lemma:
-  assumes r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
-    and aS: "a \<in> S"
-  shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
-proof (rule Zorn_Lemma2)
-  show "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
-  proof
-    fix c assume "c \<in> chain S"
-    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
-    proof cases
-
-      txt {* If @{text c} is an empty chain, then every element in
-	@{text S} is an upper bound of @{text c}. *}
-
-      assume "c = {}"
-      with aS show ?thesis by fast
-
-      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
-	bound of @{text c}, lying in @{text S}. *}
-
-    next
-      assume "c \<noteq> {}"
-      show ?thesis
-      proof
-        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
-        show "\<Union>c \<in> S"
-        proof (rule r)
-          from `c \<noteq> {}` show "\<exists>x. x \<in> c" by fast
-	  show "c \<in> chain S" by fact
-        qed
-      qed
-    qed
-  qed
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/document/root.bib	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,27 +0,0 @@
-
-@Book{Heuser:1986,
-  author = 	 {H. Heuser},
-  title = 	 {Funktionalanalysis: Theorie und Anwendung},
-  publisher = 	 {Teubner},
-  year = 	 1986
-}
-
-@InCollection{Narici:1996,
-  author = 	 {L. Narici and E. Beckenstein},
-  title = 	 {The {Hahn-Banach Theorem}: The Life and Times},
-  booktitle = 	 {Topology Atlas},
-  publisher =	 {York University, Toronto, Ontario, Canada},
-  year =	 1996,
-  note =	 {\url{http://at.yorku.ca/topology/preprint.htm} and
-                  \url{http://at.yorku.ca/p/a/a/a/16.htm}}
-}
-
-@Article{Nowak:1993,
-  author =       {B. Nowak and A. Trybulec},
-  title =	 {{Hahn-Banach} Theorem},
-  journal =      {Journal of Formalized Mathematics},
-  year =         {1993},
-  volume =       {5},
-  institution =  {University of Bialystok},
-  note =         {\url{http://mizar.uwb.edu.pl/JFM/Vol5/hahnban.html}}
-}
--- a/src/HOL/Real/HahnBanach/document/root.tex	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,83 +0,0 @@
-\documentclass[10pt,a4paper,twoside]{article}
-\usepackage{graphicx}
-\usepackage{latexsym,theorem}
-\usepackage{isabelle,isabellesym}
-\usepackage{pdfsetup} %last one!
-
-\isabellestyle{it}
-\urlstyle{rm}
-
-\newcommand{\isasymsup}{\isamath{\sup\,}}
-\newcommand{\skp}{\smallskip}
-
-
-\begin{document}
-
-\pagestyle{headings}
-\pagenumbering{arabic}
-
-\title{The Hahn-Banach Theorem \\ for Real Vector Spaces}
-\author{Gertrud Bauer \\ \url{http://www.in.tum.de/~bauerg/}}
-\maketitle
-
-\begin{abstract}
-  The Hahn-Banach Theorem is one of the most fundamental results in functional
-  analysis. We present a fully formal proof of two versions of the theorem,
-  one for general linear spaces and another for normed spaces.  This
-  development is based on simply-typed classical set-theory, as provided by
-  Isabelle/HOL.
-\end{abstract}
-
-
-\tableofcontents
-\parindent 0pt \parskip 0.5ex
-
-\clearpage
-\section{Preface}
-
-This is a fully formal proof of the Hahn-Banach Theorem. It closely follows
-the informal presentation given in Heuser's textbook \cite[{\S} 36]{Heuser:1986}.
-Another formal proof of the same theorem has been done in Mizar
-\cite{Nowak:1993}.  A general overview of the relevance and history of the
-Hahn-Banach Theorem is given by Narici and Beckenstein \cite{Narici:1996}.
-
-\medskip The document is structured as follows.  The first part contains
-definitions of basic notions of linear algebra: vector spaces, subspaces,
-normed spaces, continuous linear-forms, norm of functions and an order on
-functions by domain extension.  The second part contains some lemmas about the
-supremum (w.r.t.\ the function order) and extension of non-maximal functions.
-With these preliminaries, the main proof of the theorem (in its two versions)
-is conducted in the third part.  The dependencies of individual theories are
-as follows.
-
-\begin{center}
-  \includegraphics[scale=0.5]{session_graph}  
-\end{center}
-
-\clearpage
-\part {Basic Notions}
-
-\input{Bounds}
-\input{VectorSpace}
-\input{Subspace}
-\input{NormedSpace}
-\input{Linearform}
-\input{FunctionOrder}
-\input{FunctionNorm}
-\input{ZornLemma}
-
-\clearpage
-\part {Lemmas for the Proof}
-
-\input{HahnBanachSupLemmas}
-\input{HahnBanachExtLemmas}
-\input{HahnBanachLemmas}
-
-\clearpage
-\part {The Main Proof}
-
-\input{HahnBanach}
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
--- a/src/HOL/Real/RealVector.thy	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,841 +0,0 @@
-(*  Title:      HOL/Real/RealVector.thy
-    Author:     Brian Huffman
-*)
-
-header {* Vector Spaces and Algebras over the Reals *}
-
-theory RealVector
-imports "~~/src/HOL/RealPow"
-begin
-
-subsection {* Locale for additive functions *}
-
-locale additive =
-  fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
-  assumes add: "f (x + y) = f x + f y"
-begin
-
-lemma zero: "f 0 = 0"
-proof -
-  have "f 0 = f (0 + 0)" by simp
-  also have "\<dots> = f 0 + f 0" by (rule add)
-  finally show "f 0 = 0" by simp
-qed
-
-lemma minus: "f (- x) = - f x"
-proof -
-  have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
-  also have "\<dots> = - f x + f x" by (simp add: zero)
-  finally show "f (- x) = - f x" by (rule add_right_imp_eq)
-qed
-
-lemma diff: "f (x - y) = f x - f y"
-by (simp add: diff_def add minus)
-
-lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
-apply (cases "finite A")
-apply (induct set: finite)
-apply (simp add: zero)
-apply (simp add: add)
-apply (simp add: zero)
-done
-
-end
-
-subsection {* Vector spaces *}
-
-locale vector_space =
-  fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
-  assumes scale_right_distrib: "scale a (x + y) = scale a x + scale a y"
-  and scale_left_distrib: "scale (a + b) x = scale a x + scale b x"
-  and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
-  and scale_one [simp]: "scale 1 x = x"
-begin
-
-lemma scale_left_commute:
-  "scale a (scale b x) = scale b (scale a x)"
-by (simp add: mult_commute)
-
-lemma scale_zero_left [simp]: "scale 0 x = 0"
-  and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
-  and scale_left_diff_distrib: "scale (a - b) x = scale a x - scale b x"
-proof -
-  interpret s: additive ["\<lambda>a. scale a x"]
-    proof qed (rule scale_left_distrib)
-  show "scale 0 x = 0" by (rule s.zero)
-  show "scale (- a) x = - (scale a x)" by (rule s.minus)
-  show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
-qed
-
-lemma scale_zero_right [simp]: "scale a 0 = 0"
-  and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
-  and scale_right_diff_distrib: "scale a (x - y) = scale a x - scale a y"
-proof -
-  interpret s: additive ["\<lambda>x. scale a x"]
-    proof qed (rule scale_right_distrib)
-  show "scale a 0 = 0" by (rule s.zero)
-  show "scale a (- x) = - (scale a x)" by (rule s.minus)
-  show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
-qed
-
-lemma scale_eq_0_iff [simp]:
-  "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
-proof cases
-  assume "a = 0" thus ?thesis by simp
-next
-  assume anz [simp]: "a \<noteq> 0"
-  { assume "scale a x = 0"
-    hence "scale (inverse a) (scale a x) = 0" by simp
-    hence "x = 0" by simp }
-  thus ?thesis by force
-qed
-
-lemma scale_left_imp_eq:
-  "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
-proof -
-  assume nonzero: "a \<noteq> 0"
-  assume "scale a x = scale a y"
-  hence "scale a (x - y) = 0"
-     by (simp add: scale_right_diff_distrib)
-  hence "x - y = 0" by (simp add: nonzero)
-  thus "x = y" by (simp only: right_minus_eq)
-qed
-
-lemma scale_right_imp_eq:
-  "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
-proof -
-  assume nonzero: "x \<noteq> 0"
-  assume "scale a x = scale b x"
-  hence "scale (a - b) x = 0"
-     by (simp add: scale_left_diff_distrib)
-  hence "a - b = 0" by (simp add: nonzero)
-  thus "a = b" by (simp only: right_minus_eq)
-qed
-
-lemma scale_cancel_left:
-  "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
-by (auto intro: scale_left_imp_eq)
-
-lemma scale_cancel_right:
-  "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
-by (auto intro: scale_right_imp_eq)
-
-end
-
-subsection {* Real vector spaces *}
-
-class scaleR = type +
-  fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
-begin
-
-abbreviation
-  divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
-where
-  "x /\<^sub>R r == scaleR (inverse r) x"
-
-end
-
-instantiation real :: scaleR
-begin
-
-definition
-  real_scaleR_def [simp]: "scaleR a x = a * x"
-
-instance ..
-
-end
-
-class real_vector = scaleR + ab_group_add +
-  assumes scaleR_right_distrib: "scaleR a (x + y) = scaleR a x + scaleR a y"
-  and scaleR_left_distrib: "scaleR (a + b) x = scaleR a x + scaleR b x"
-  and scaleR_scaleR [simp]: "scaleR a (scaleR b x) = scaleR (a * b) x"
-  and scaleR_one [simp]: "scaleR 1 x = x"
-
-interpretation real_vector:
-  vector_space ["scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"]
-apply unfold_locales
-apply (rule scaleR_right_distrib)
-apply (rule scaleR_left_distrib)
-apply (rule scaleR_scaleR)
-apply (rule scaleR_one)
-done
-
-text {* Recover original theorem names *}
-
-lemmas scaleR_left_commute = real_vector.scale_left_commute
-lemmas scaleR_zero_left = real_vector.scale_zero_left
-lemmas scaleR_minus_left = real_vector.scale_minus_left
-lemmas scaleR_left_diff_distrib = real_vector.scale_left_diff_distrib
-lemmas scaleR_zero_right = real_vector.scale_zero_right
-lemmas scaleR_minus_right = real_vector.scale_minus_right
-lemmas scaleR_right_diff_distrib = real_vector.scale_right_diff_distrib
-lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
-lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
-lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
-lemmas scaleR_cancel_left = real_vector.scale_cancel_left
-lemmas scaleR_cancel_right = real_vector.scale_cancel_right
-
-class real_algebra = real_vector + ring +
-  assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
-  and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
-
-class real_algebra_1 = real_algebra + ring_1
-
-class real_div_algebra = real_algebra_1 + division_ring
-
-class real_field = real_div_algebra + field
-
-instance real :: real_field
-apply (intro_classes, unfold real_scaleR_def)
-apply (rule right_distrib)
-apply (rule left_distrib)
-apply (rule mult_assoc [symmetric])
-apply (rule mult_1_left)
-apply (rule mult_assoc)
-apply (rule mult_left_commute)
-done
-
-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)"]
-proof qed (rule scaleR_right_distrib)
-
-lemma nonzero_inverse_scaleR_distrib:
-  fixes x :: "'a::real_div_algebra" shows
-  "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
-by (rule inverse_unique, simp)
-
-lemma inverse_scaleR_distrib:
-  fixes x :: "'a::{real_div_algebra,division_by_zero}"
-  shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
-apply (case_tac "a = 0", simp)
-apply (case_tac "x = 0", simp)
-apply (erule (1) nonzero_inverse_scaleR_distrib)
-done
-
-
-subsection {* Embedding of the Reals into any @{text real_algebra_1}:
-@{term of_real} *}
-
-definition
-  of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
-  "of_real r = scaleR r 1"
-
-lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
-by (simp add: of_real_def)
-
-lemma of_real_0 [simp]: "of_real 0 = 0"
-by (simp add: of_real_def)
-
-lemma of_real_1 [simp]: "of_real 1 = 1"
-by (simp add: of_real_def)
-
-lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
-by (simp add: of_real_def scaleR_left_distrib)
-
-lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
-by (simp add: of_real_def)
-
-lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
-by (simp add: of_real_def scaleR_left_diff_distrib)
-
-lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
-by (simp add: of_real_def mult_commute)
-
-lemma nonzero_of_real_inverse:
-  "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
-   inverse (of_real x :: 'a::real_div_algebra)"
-by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
-
-lemma of_real_inverse [simp]:
-  "of_real (inverse x) =
-   inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
-by (simp add: of_real_def inverse_scaleR_distrib)
-
-lemma nonzero_of_real_divide:
-  "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
-   (of_real x / of_real y :: 'a::real_field)"
-by (simp add: divide_inverse nonzero_of_real_inverse)
-
-lemma of_real_divide [simp]:
-  "of_real (x / y) =
-   (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
-by (simp add: divide_inverse)
-
-lemma of_real_power [simp]:
-  "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1,recpower}) ^ n"
-by (induct n) (simp_all add: power_Suc)
-
-lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
-by (simp add: of_real_def scaleR_cancel_right)
-
-lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
-
-lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
-proof
-  fix r
-  show "of_real r = id r"
-    by (simp add: of_real_def)
-qed
-
-text{*Collapse nested embeddings*}
-lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
-by (induct n) auto
-
-lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
-by (cases z rule: int_diff_cases, simp)
-
-lemma of_real_number_of_eq:
-  "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})"
-by (simp add: number_of_eq)
-
-text{*Every real algebra has characteristic zero*}
-instance real_algebra_1 < ring_char_0
-proof
-  fix m n :: nat
-  have "(of_real (of_nat m) = (of_real (of_nat n)::'a)) = (m = n)"
-    by (simp only: of_real_eq_iff of_nat_eq_iff)
-  thus "(of_nat m = (of_nat n::'a)) = (m = n)"
-    by (simp only: of_real_of_nat_eq)
-qed
-
-instance real_field < field_char_0 ..
-
-
-subsection {* The Set of Real Numbers *}
-
-definition
-  Reals :: "'a::real_algebra_1 set" where
-  [code del]: "Reals \<equiv> range of_real"
-
-notation (xsymbols)
-  Reals  ("\<real>")
-
-lemma Reals_of_real [simp]: "of_real r \<in> Reals"
-by (simp add: Reals_def)
-
-lemma Reals_of_int [simp]: "of_int z \<in> Reals"
-by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
-
-lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
-by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
-
-lemma Reals_number_of [simp]:
-  "(number_of w::'a::{number_ring,real_algebra_1}) \<in> Reals"
-by (subst of_real_number_of_eq [symmetric], rule Reals_of_real)
-
-lemma Reals_0 [simp]: "0 \<in> Reals"
-apply (unfold Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_0 [symmetric])
-done
-
-lemma Reals_1 [simp]: "1 \<in> Reals"
-apply (unfold Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_1 [symmetric])
-done
-
-lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_add [symmetric])
-done
-
-lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_minus [symmetric])
-done
-
-lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_diff [symmetric])
-done
-
-lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_mult [symmetric])
-done
-
-lemma nonzero_Reals_inverse:
-  fixes a :: "'a::real_div_algebra"
-  shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_real_inverse [symmetric])
-done
-
-lemma Reals_inverse [simp]:
-  fixes a :: "'a::{real_div_algebra,division_by_zero}"
-  shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_inverse [symmetric])
-done
-
-lemma nonzero_Reals_divide:
-  fixes a b :: "'a::real_field"
-  shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_real_divide [symmetric])
-done
-
-lemma Reals_divide [simp]:
-  fixes a b :: "'a::{real_field,division_by_zero}"
-  shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_divide [symmetric])
-done
-
-lemma Reals_power [simp]:
-  fixes a :: "'a::{real_algebra_1,recpower}"
-  shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_power [symmetric])
-done
-
-lemma Reals_cases [cases set: Reals]:
-  assumes "q \<in> \<real>"
-  obtains (of_real) r where "q = of_real r"
-  unfolding Reals_def
-proof -
-  from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
-  then obtain r where "q = of_real r" ..
-  then show thesis ..
-qed
-
-lemma Reals_induct [case_names of_real, induct set: Reals]:
-  "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
-  by (rule Reals_cases) auto
-
-
-subsection {* Real normed vector spaces *}
-
-class norm = type +
-  fixes norm :: "'a \<Rightarrow> real"
-
-instantiation real :: norm
-begin
-
-definition
-  real_norm_def [simp]: "norm r \<equiv> \<bar>r\<bar>"
-
-instance ..
-
-end
-
-class sgn_div_norm = scaleR + norm + sgn +
-  assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
-
-class real_normed_vector = real_vector + sgn_div_norm +
-  assumes norm_ge_zero [simp]: "0 \<le> norm x"
-  and norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
-  and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
-  and norm_scaleR: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
-
-class real_normed_algebra = real_algebra + real_normed_vector +
-  assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
-
-class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
-  assumes norm_one [simp]: "norm 1 = 1"
-
-class real_normed_div_algebra = real_div_algebra + real_normed_vector +
-  assumes norm_mult: "norm (x * y) = norm x * norm y"
-
-class real_normed_field = real_field + real_normed_div_algebra
-
-instance real_normed_div_algebra < real_normed_algebra_1
-proof
-  fix x y :: 'a
-  show "norm (x * y) \<le> norm x * norm y"
-    by (simp add: norm_mult)
-next
-  have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
-    by (rule norm_mult)
-  thus "norm (1::'a) = 1" by simp
-qed
-
-instance real :: real_normed_field
-apply (intro_classes, unfold real_norm_def real_scaleR_def)
-apply (simp add: real_sgn_def)
-apply (rule abs_ge_zero)
-apply (rule abs_eq_0)
-apply (rule abs_triangle_ineq)
-apply (rule abs_mult)
-apply (rule abs_mult)
-done
-
-lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
-by simp
-
-lemma zero_less_norm_iff [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "(0 < norm x) = (x \<noteq> 0)"
-by (simp add: order_less_le)
-
-lemma norm_not_less_zero [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "\<not> norm x < 0"
-by (simp add: linorder_not_less)
-
-lemma norm_le_zero_iff [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "(norm x \<le> 0) = (x = 0)"
-by (simp add: order_le_less)
-
-lemma norm_minus_cancel [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "norm (- x) = norm x"
-proof -
-  have "norm (- x) = norm (scaleR (- 1) x)"
-    by (simp only: scaleR_minus_left scaleR_one)
-  also have "\<dots> = \<bar>- 1\<bar> * norm x"
-    by (rule norm_scaleR)
-  finally show ?thesis by simp
-qed
-
-lemma norm_minus_commute:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm (a - b) = norm (b - a)"
-proof -
-  have "norm (- (b - a)) = norm (b - a)"
-    by (rule norm_minus_cancel)
-  thus ?thesis by simp
-qed
-
-lemma norm_triangle_ineq2:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm a - norm b \<le> norm (a - b)"
-proof -
-  have "norm (a - b + b) \<le> norm (a - b) + norm b"
-    by (rule norm_triangle_ineq)
-  thus ?thesis by simp
-qed
-
-lemma norm_triangle_ineq3:
-  fixes a b :: "'a::real_normed_vector"
-  shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
-apply (subst abs_le_iff)
-apply auto
-apply (rule norm_triangle_ineq2)
-apply (subst norm_minus_commute)
-apply (rule norm_triangle_ineq2)
-done
-
-lemma norm_triangle_ineq4:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm (a - b) \<le> norm a + norm b"
-proof -
-  have "norm (a + - b) \<le> norm a + norm (- b)"
-    by (rule norm_triangle_ineq)
-  thus ?thesis
-    by (simp only: diff_minus norm_minus_cancel)
-qed
-
-lemma norm_diff_ineq:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm a - norm b \<le> norm (a + b)"
-proof -
-  have "norm a - norm (- b) \<le> norm (a - - b)"
-    by (rule norm_triangle_ineq2)
-  thus ?thesis by simp
-qed
-
-lemma norm_diff_triangle_ineq:
-  fixes a b c d :: "'a::real_normed_vector"
-  shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
-proof -
-  have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
-    by (simp add: diff_minus add_ac)
-  also have "\<dots> \<le> norm (a - c) + norm (b - d)"
-    by (rule norm_triangle_ineq)
-  finally show ?thesis .
-qed
-
-lemma abs_norm_cancel [simp]:
-  fixes a :: "'a::real_normed_vector"
-  shows "\<bar>norm a\<bar> = norm a"
-by (rule abs_of_nonneg [OF norm_ge_zero])
-
-lemma norm_add_less:
-  fixes x y :: "'a::real_normed_vector"
-  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
-by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
-
-lemma norm_mult_less:
-  fixes x y :: "'a::real_normed_algebra"
-  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
-apply (rule order_le_less_trans [OF norm_mult_ineq])
-apply (simp add: mult_strict_mono')
-done
-
-lemma norm_of_real [simp]:
-  "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
-unfolding of_real_def by (simp add: norm_scaleR)
-
-lemma norm_number_of [simp]:
-  "norm (number_of w::'a::{number_ring,real_normed_algebra_1})
-    = \<bar>number_of w\<bar>"
-by (subst of_real_number_of_eq [symmetric], rule norm_of_real)
-
-lemma norm_of_int [simp]:
-  "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
-by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
-
-lemma norm_of_nat [simp]:
-  "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
-apply (subst of_real_of_nat_eq [symmetric])
-apply (subst norm_of_real, simp)
-done
-
-lemma nonzero_norm_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
-apply (rule inverse_unique [symmetric])
-apply (simp add: norm_mult [symmetric])
-done
-
-lemma norm_inverse:
-  fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
-  shows "norm (inverse a) = inverse (norm a)"
-apply (case_tac "a = 0", simp)
-apply (erule nonzero_norm_inverse)
-done
-
-lemma nonzero_norm_divide:
-  fixes a b :: "'a::real_normed_field"
-  shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
-by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
-
-lemma norm_divide:
-  fixes a b :: "'a::{real_normed_field,division_by_zero}"
-  shows "norm (a / b) = norm a / norm b"
-by (simp add: divide_inverse norm_mult norm_inverse)
-
-lemma norm_power_ineq:
-  fixes x :: "'a::{real_normed_algebra_1,recpower}"
-  shows "norm (x ^ n) \<le> norm x ^ n"
-proof (induct n)
-  case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
-next
-  case (Suc n)
-  have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
-    by (rule norm_mult_ineq)
-  also from Suc have "\<dots> \<le> norm x * norm x ^ n"
-    using norm_ge_zero by (rule mult_left_mono)
-  finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
-    by (simp add: power_Suc)
-qed
-
-lemma norm_power:
-  fixes x :: "'a::{real_normed_div_algebra,recpower}"
-  shows "norm (x ^ n) = norm x ^ n"
-by (induct n) (simp_all add: power_Suc norm_mult)
-
-
-subsection {* Sign function *}
-
-lemma norm_sgn:
-  "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
-by (simp add: sgn_div_norm norm_scaleR)
-
-lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
-by (simp add: sgn_div_norm)
-
-lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
-by (simp add: sgn_div_norm)
-
-lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
-by (simp add: sgn_div_norm)
-
-lemma sgn_scaleR:
-  "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
-by (simp add: sgn_div_norm norm_scaleR mult_ac)
-
-lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
-by (simp add: sgn_div_norm)
-
-lemma sgn_of_real:
-  "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
-unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
-
-lemma sgn_mult:
-  fixes x y :: "'a::real_normed_div_algebra"
-  shows "sgn (x * y) = sgn x * sgn y"
-by (simp add: sgn_div_norm norm_mult mult_commute)
-
-lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
-by (simp add: sgn_div_norm divide_inverse)
-
-lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
-unfolding real_sgn_eq by simp
-
-lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
-unfolding real_sgn_eq by simp
-
-
-subsection {* Bounded Linear and Bilinear Operators *}
-
-locale bounded_linear = additive +
-  constrains f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
-  assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
-begin
-
-lemma pos_bounded:
-  "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
-proof -
-  obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
-    using bounded by fast
-  show ?thesis
-  proof (intro exI impI conjI allI)
-    show "0 < max 1 K"
-      by (rule order_less_le_trans [OF zero_less_one le_maxI1])
-  next
-    fix x
-    have "norm (f x) \<le> norm x * K" using K .
-    also have "\<dots> \<le> norm x * max 1 K"
-      by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
-    finally show "norm (f x) \<le> norm x * max 1 K" .
-  qed
-qed
-
-lemma nonneg_bounded:
-  "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
-proof -
-  from pos_bounded
-  show ?thesis by (auto intro: order_less_imp_le)
-qed
-
-end
-
-locale bounded_bilinear =
-  fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
-                 \<Rightarrow> 'c::real_normed_vector"
-    (infixl "**" 70)
-  assumes add_left: "prod (a + a') b = prod a b + prod a' b"
-  assumes add_right: "prod a (b + b') = prod a b + prod a b'"
-  assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
-  assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
-  assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
-begin
-
-lemma pos_bounded:
-  "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
-apply (cut_tac bounded, erule exE)
-apply (rule_tac x="max 1 K" in exI, safe)
-apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
-apply (drule spec, drule spec, erule order_trans)
-apply (rule mult_left_mono [OF le_maxI2])
-apply (intro mult_nonneg_nonneg norm_ge_zero)
-done
-
-lemma nonneg_bounded:
-  "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
-proof -
-  from pos_bounded
-  show ?thesis by (auto intro: order_less_imp_le)
-qed
-
-lemma additive_right: "additive (\<lambda>b. prod a b)"
-by (rule additive.intro, rule add_right)
-
-lemma additive_left: "additive (\<lambda>a. prod a b)"
-by (rule additive.intro, rule add_left)
-
-lemma zero_left: "prod 0 b = 0"
-by (rule additive.zero [OF additive_left])
-
-lemma zero_right: "prod a 0 = 0"
-by (rule additive.zero [OF additive_right])
-
-lemma minus_left: "prod (- a) b = - prod a b"
-by (rule additive.minus [OF additive_left])
-
-lemma minus_right: "prod a (- b) = - prod a b"
-by (rule additive.minus [OF additive_right])
-
-lemma diff_left:
-  "prod (a - a') b = prod a b - prod a' b"
-by (rule additive.diff [OF additive_left])
-
-lemma diff_right:
-  "prod a (b - b') = prod a b - prod a b'"
-by (rule additive.diff [OF additive_right])
-
-lemma bounded_linear_left:
-  "bounded_linear (\<lambda>a. a ** b)"
-apply (unfold_locales)
-apply (rule add_left)
-apply (rule scaleR_left)
-apply (cut_tac bounded, safe)
-apply (rule_tac x="norm b * K" in exI)
-apply (simp add: mult_ac)
-done
-
-lemma bounded_linear_right:
-  "bounded_linear (\<lambda>b. a ** b)"
-apply (unfold_locales)
-apply (rule add_right)
-apply (rule scaleR_right)
-apply (cut_tac bounded, safe)
-apply (rule_tac x="norm a * K" in exI)
-apply (simp add: mult_ac)
-done
-
-lemma prod_diff_prod:
-  "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
-by (simp add: diff_left diff_right)
-
-end
-
-interpretation mult:
-  bounded_bilinear ["op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"]
-apply (rule bounded_bilinear.intro)
-apply (rule left_distrib)
-apply (rule right_distrib)
-apply (rule mult_scaleR_left)
-apply (rule mult_scaleR_right)
-apply (rule_tac x="1" in exI)
-apply (simp add: norm_mult_ineq)
-done
-
-interpretation mult_left:
-  bounded_linear ["(\<lambda>x::'a::real_normed_algebra. x * y)"]
-by (rule mult.bounded_linear_left)
-
-interpretation mult_right:
-  bounded_linear ["(\<lambda>y::'a::real_normed_algebra. x * y)"]
-by (rule mult.bounded_linear_right)
-
-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"]
-apply (rule bounded_bilinear.intro)
-apply (rule scaleR_left_distrib)
-apply (rule scaleR_right_distrib)
-apply simp
-apply (rule scaleR_left_commute)
-apply (rule_tac x="1" in exI)
-apply (simp add: norm_scaleR)
-done
-
-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"]
-by (rule scaleR.bounded_linear_right)
-
-interpretation of_real: bounded_linear ["\<lambda>r. of_real r"]
-unfolding of_real_def by (rule scaleR.bounded_linear_left)
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/RealVector.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,841 @@
+(*  Title:      HOL/RealVector.thy
+    Author:     Brian Huffman
+*)
+
+header {* Vector Spaces and Algebras over the Reals *}
+
+theory RealVector
+imports RealPow
+begin
+
+subsection {* Locale for additive functions *}
+
+locale additive =
+  fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
+  assumes add: "f (x + y) = f x + f y"
+begin
+
+lemma zero: "f 0 = 0"
+proof -
+  have "f 0 = f (0 + 0)" by simp
+  also have "\<dots> = f 0 + f 0" by (rule add)
+  finally show "f 0 = 0" by simp
+qed
+
+lemma minus: "f (- x) = - f x"
+proof -
+  have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
+  also have "\<dots> = - f x + f x" by (simp add: zero)
+  finally show "f (- x) = - f x" by (rule add_right_imp_eq)
+qed
+
+lemma diff: "f (x - y) = f x - f y"
+by (simp add: diff_def add minus)
+
+lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
+apply (cases "finite A")
+apply (induct set: finite)
+apply (simp add: zero)
+apply (simp add: add)
+apply (simp add: zero)
+done
+
+end
+
+subsection {* Vector spaces *}
+
+locale vector_space =
+  fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
+  assumes scale_right_distrib: "scale a (x + y) = scale a x + scale a y"
+  and scale_left_distrib: "scale (a + b) x = scale a x + scale b x"
+  and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
+  and scale_one [simp]: "scale 1 x = x"
+begin
+
+lemma scale_left_commute:
+  "scale a (scale b x) = scale b (scale a x)"
+by (simp add: mult_commute)
+
+lemma scale_zero_left [simp]: "scale 0 x = 0"
+  and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
+  and scale_left_diff_distrib: "scale (a - b) x = scale a x - scale b x"
+proof -
+  interpret s: additive "\<lambda>a. scale a x"
+    proof qed (rule scale_left_distrib)
+  show "scale 0 x = 0" by (rule s.zero)
+  show "scale (- a) x = - (scale a x)" by (rule s.minus)
+  show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
+qed
+
+lemma scale_zero_right [simp]: "scale a 0 = 0"
+  and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
+  and scale_right_diff_distrib: "scale a (x - y) = scale a x - scale a y"
+proof -
+  interpret s: additive "\<lambda>x. scale a x"
+    proof qed (rule scale_right_distrib)
+  show "scale a 0 = 0" by (rule s.zero)
+  show "scale a (- x) = - (scale a x)" by (rule s.minus)
+  show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
+qed
+
+lemma scale_eq_0_iff [simp]:
+  "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
+proof cases
+  assume "a = 0" thus ?thesis by simp
+next
+  assume anz [simp]: "a \<noteq> 0"
+  { assume "scale a x = 0"
+    hence "scale (inverse a) (scale a x) = 0" by simp
+    hence "x = 0" by simp }
+  thus ?thesis by force
+qed
+
+lemma scale_left_imp_eq:
+  "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
+proof -
+  assume nonzero: "a \<noteq> 0"
+  assume "scale a x = scale a y"
+  hence "scale a (x - y) = 0"
+     by (simp add: scale_right_diff_distrib)
+  hence "x - y = 0" by (simp add: nonzero)
+  thus "x = y" by (simp only: right_minus_eq)
+qed
+
+lemma scale_right_imp_eq:
+  "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
+proof -
+  assume nonzero: "x \<noteq> 0"
+  assume "scale a x = scale b x"
+  hence "scale (a - b) x = 0"
+     by (simp add: scale_left_diff_distrib)
+  hence "a - b = 0" by (simp add: nonzero)
+  thus "a = b" by (simp only: right_minus_eq)
+qed
+
+lemma scale_cancel_left:
+  "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
+by (auto intro: scale_left_imp_eq)
+
+lemma scale_cancel_right:
+  "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
+by (auto intro: scale_right_imp_eq)
+
+end
+
+subsection {* Real vector spaces *}
+
+class scaleR = type +
+  fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
+begin
+
+abbreviation
+  divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
+where
+  "x /\<^sub>R r == scaleR (inverse r) x"
+
+end
+
+instantiation real :: scaleR
+begin
+
+definition
+  real_scaleR_def [simp]: "scaleR a x = a * x"
+
+instance ..
+
+end
+
+class real_vector = scaleR + ab_group_add +
+  assumes scaleR_right_distrib: "scaleR a (x + y) = scaleR a x + scaleR a y"
+  and scaleR_left_distrib: "scaleR (a + b) x = scaleR a x + scaleR b x"
+  and scaleR_scaleR [simp]: "scaleR a (scaleR b x) = scaleR (a * b) x"
+  and scaleR_one [simp]: "scaleR 1 x = x"
+
+interpretation real_vector!:
+  vector_space "scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"
+apply unfold_locales
+apply (rule scaleR_right_distrib)
+apply (rule scaleR_left_distrib)
+apply (rule scaleR_scaleR)
+apply (rule scaleR_one)
+done
+
+text {* Recover original theorem names *}
+
+lemmas scaleR_left_commute = real_vector.scale_left_commute
+lemmas scaleR_zero_left = real_vector.scale_zero_left
+lemmas scaleR_minus_left = real_vector.scale_minus_left
+lemmas scaleR_left_diff_distrib = real_vector.scale_left_diff_distrib
+lemmas scaleR_zero_right = real_vector.scale_zero_right
+lemmas scaleR_minus_right = real_vector.scale_minus_right
+lemmas scaleR_right_diff_distrib = real_vector.scale_right_diff_distrib
+lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
+lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
+lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
+lemmas scaleR_cancel_left = real_vector.scale_cancel_left
+lemmas scaleR_cancel_right = real_vector.scale_cancel_right
+
+class real_algebra = real_vector + ring +
+  assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
+  and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
+
+class real_algebra_1 = real_algebra + ring_1
+
+class real_div_algebra = real_algebra_1 + division_ring
+
+class real_field = real_div_algebra + field
+
+instance real :: real_field
+apply (intro_classes, unfold real_scaleR_def)
+apply (rule right_distrib)
+apply (rule left_distrib)
+apply (rule mult_assoc [symmetric])
+apply (rule mult_1_left)
+apply (rule mult_assoc)
+apply (rule mult_left_commute)
+done
+
+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)"
+proof qed (rule scaleR_right_distrib)
+
+lemma nonzero_inverse_scaleR_distrib:
+  fixes x :: "'a::real_div_algebra" shows
+  "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
+by (rule inverse_unique, simp)
+
+lemma inverse_scaleR_distrib:
+  fixes x :: "'a::{real_div_algebra,division_by_zero}"
+  shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
+apply (case_tac "a = 0", simp)
+apply (case_tac "x = 0", simp)
+apply (erule (1) nonzero_inverse_scaleR_distrib)
+done
+
+
+subsection {* Embedding of the Reals into any @{text real_algebra_1}:
+@{term of_real} *}
+
+definition
+  of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
+  "of_real r = scaleR r 1"
+
+lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
+by (simp add: of_real_def)
+
+lemma of_real_0 [simp]: "of_real 0 = 0"
+by (simp add: of_real_def)
+
+lemma of_real_1 [simp]: "of_real 1 = 1"
+by (simp add: of_real_def)
+
+lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
+by (simp add: of_real_def scaleR_left_distrib)
+
+lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
+by (simp add: of_real_def)
+
+lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
+by (simp add: of_real_def scaleR_left_diff_distrib)
+
+lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
+by (simp add: of_real_def mult_commute)
+
+lemma nonzero_of_real_inverse:
+  "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
+   inverse (of_real x :: 'a::real_div_algebra)"
+by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
+
+lemma of_real_inverse [simp]:
+  "of_real (inverse x) =
+   inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
+by (simp add: of_real_def inverse_scaleR_distrib)
+
+lemma nonzero_of_real_divide:
+  "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
+   (of_real x / of_real y :: 'a::real_field)"
+by (simp add: divide_inverse nonzero_of_real_inverse)
+
+lemma of_real_divide [simp]:
+  "of_real (x / y) =
+   (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
+by (simp add: divide_inverse)
+
+lemma of_real_power [simp]:
+  "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1,recpower}) ^ n"
+by (induct n) (simp_all add: power_Suc)
+
+lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
+by (simp add: of_real_def scaleR_cancel_right)
+
+lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
+
+lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
+proof
+  fix r
+  show "of_real r = id r"
+    by (simp add: of_real_def)
+qed
+
+text{*Collapse nested embeddings*}
+lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
+by (induct n) auto
+
+lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
+by (cases z rule: int_diff_cases, simp)
+
+lemma of_real_number_of_eq:
+  "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})"
+by (simp add: number_of_eq)
+
+text{*Every real algebra has characteristic zero*}
+instance real_algebra_1 < ring_char_0
+proof
+  fix m n :: nat
+  have "(of_real (of_nat m) = (of_real (of_nat n)::'a)) = (m = n)"
+    by (simp only: of_real_eq_iff of_nat_eq_iff)
+  thus "(of_nat m = (of_nat n::'a)) = (m = n)"
+    by (simp only: of_real_of_nat_eq)
+qed
+
+instance real_field < field_char_0 ..
+
+
+subsection {* The Set of Real Numbers *}
+
+definition
+  Reals :: "'a::real_algebra_1 set" where
+  [code del]: "Reals \<equiv> range of_real"
+
+notation (xsymbols)
+  Reals  ("\<real>")
+
+lemma Reals_of_real [simp]: "of_real r \<in> Reals"
+by (simp add: Reals_def)
+
+lemma Reals_of_int [simp]: "of_int z \<in> Reals"
+by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
+
+lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
+by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
+
+lemma Reals_number_of [simp]:
+  "(number_of w::'a::{number_ring,real_algebra_1}) \<in> Reals"
+by (subst of_real_number_of_eq [symmetric], rule Reals_of_real)
+
+lemma Reals_0 [simp]: "0 \<in> Reals"
+apply (unfold Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_0 [symmetric])
+done
+
+lemma Reals_1 [simp]: "1 \<in> Reals"
+apply (unfold Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_1 [symmetric])
+done
+
+lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_add [symmetric])
+done
+
+lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_minus [symmetric])
+done
+
+lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_diff [symmetric])
+done
+
+lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_mult [symmetric])
+done
+
+lemma nonzero_Reals_inverse:
+  fixes a :: "'a::real_div_algebra"
+  shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (erule nonzero_of_real_inverse [symmetric])
+done
+
+lemma Reals_inverse [simp]:
+  fixes a :: "'a::{real_div_algebra,division_by_zero}"
+  shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_inverse [symmetric])
+done
+
+lemma nonzero_Reals_divide:
+  fixes a b :: "'a::real_field"
+  shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (erule nonzero_of_real_divide [symmetric])
+done
+
+lemma Reals_divide [simp]:
+  fixes a b :: "'a::{real_field,division_by_zero}"
+  shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_divide [symmetric])
+done
+
+lemma Reals_power [simp]:
+  fixes a :: "'a::{real_algebra_1,recpower}"
+  shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_power [symmetric])
+done
+
+lemma Reals_cases [cases set: Reals]:
+  assumes "q \<in> \<real>"
+  obtains (of_real) r where "q = of_real r"
+  unfolding Reals_def
+proof -
+  from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
+  then obtain r where "q = of_real r" ..
+  then show thesis ..
+qed
+
+lemma Reals_induct [case_names of_real, induct set: Reals]:
+  "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
+  by (rule Reals_cases) auto
+
+
+subsection {* Real normed vector spaces *}
+
+class norm = type +
+  fixes norm :: "'a \<Rightarrow> real"
+
+instantiation real :: norm
+begin
+
+definition
+  real_norm_def [simp]: "norm r \<equiv> \<bar>r\<bar>"
+
+instance ..
+
+end
+
+class sgn_div_norm = scaleR + norm + sgn +
+  assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
+
+class real_normed_vector = real_vector + sgn_div_norm +
+  assumes norm_ge_zero [simp]: "0 \<le> norm x"
+  and norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
+  and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
+  and norm_scaleR: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
+
+class real_normed_algebra = real_algebra + real_normed_vector +
+  assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
+
+class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
+  assumes norm_one [simp]: "norm 1 = 1"
+
+class real_normed_div_algebra = real_div_algebra + real_normed_vector +
+  assumes norm_mult: "norm (x * y) = norm x * norm y"
+
+class real_normed_field = real_field + real_normed_div_algebra
+
+instance real_normed_div_algebra < real_normed_algebra_1
+proof
+  fix x y :: 'a
+  show "norm (x * y) \<le> norm x * norm y"
+    by (simp add: norm_mult)
+next
+  have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
+    by (rule norm_mult)
+  thus "norm (1::'a) = 1" by simp
+qed
+
+instance real :: real_normed_field
+apply (intro_classes, unfold real_norm_def real_scaleR_def)
+apply (simp add: real_sgn_def)
+apply (rule abs_ge_zero)
+apply (rule abs_eq_0)
+apply (rule abs_triangle_ineq)
+apply (rule abs_mult)
+apply (rule abs_mult)
+done
+
+lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
+by simp
+
+lemma zero_less_norm_iff [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "(0 < norm x) = (x \<noteq> 0)"
+by (simp add: order_less_le)
+
+lemma norm_not_less_zero [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "\<not> norm x < 0"
+by (simp add: linorder_not_less)
+
+lemma norm_le_zero_iff [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "(norm x \<le> 0) = (x = 0)"
+by (simp add: order_le_less)
+
+lemma norm_minus_cancel [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "norm (- x) = norm x"
+proof -
+  have "norm (- x) = norm (scaleR (- 1) x)"
+    by (simp only: scaleR_minus_left scaleR_one)
+  also have "\<dots> = \<bar>- 1\<bar> * norm x"
+    by (rule norm_scaleR)
+  finally show ?thesis by simp
+qed
+
+lemma norm_minus_commute:
+  fixes a b :: "'a::real_normed_vector"
+  shows "norm (a - b) = norm (b - a)"
+proof -
+  have "norm (- (b - a)) = norm (b - a)"
+    by (rule norm_minus_cancel)
+  thus ?thesis by simp
+qed
+
+lemma norm_triangle_ineq2:
+  fixes a b :: "'a::real_normed_vector"
+  shows "norm a - norm b \<le> norm (a - b)"
+proof -
+  have "norm (a - b + b) \<le> norm (a - b) + norm b"
+    by (rule norm_triangle_ineq)
+  thus ?thesis by simp
+qed
+
+lemma norm_triangle_ineq3:
+  fixes a b :: "'a::real_normed_vector"
+  shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
+apply (subst abs_le_iff)
+apply auto
+apply (rule norm_triangle_ineq2)
+apply (subst norm_minus_commute)
+apply (rule norm_triangle_ineq2)
+done
+
+lemma norm_triangle_ineq4:
+  fixes a b :: "'a::real_normed_vector"
+  shows "norm (a - b) \<le> norm a + norm b"
+proof -
+  have "norm (a + - b) \<le> norm a + norm (- b)"
+    by (rule norm_triangle_ineq)
+  thus ?thesis
+    by (simp only: diff_minus norm_minus_cancel)
+qed
+
+lemma norm_diff_ineq:
+  fixes a b :: "'a::real_normed_vector"
+  shows "norm a - norm b \<le> norm (a + b)"
+proof -
+  have "norm a - norm (- b) \<le> norm (a - - b)"
+    by (rule norm_triangle_ineq2)
+  thus ?thesis by simp
+qed
+
+lemma norm_diff_triangle_ineq:
+  fixes a b c d :: "'a::real_normed_vector"
+  shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
+proof -
+  have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
+    by (simp add: diff_minus add_ac)
+  also have "\<dots> \<le> norm (a - c) + norm (b - d)"
+    by (rule norm_triangle_ineq)
+  finally show ?thesis .
+qed
+
+lemma abs_norm_cancel [simp]:
+  fixes a :: "'a::real_normed_vector"
+  shows "\<bar>norm a\<bar> = norm a"
+by (rule abs_of_nonneg [OF norm_ge_zero])
+
+lemma norm_add_less:
+  fixes x y :: "'a::real_normed_vector"
+  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
+by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
+
+lemma norm_mult_less:
+  fixes x y :: "'a::real_normed_algebra"
+  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
+apply (rule order_le_less_trans [OF norm_mult_ineq])
+apply (simp add: mult_strict_mono')
+done
+
+lemma norm_of_real [simp]:
+  "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
+unfolding of_real_def by (simp add: norm_scaleR)
+
+lemma norm_number_of [simp]:
+  "norm (number_of w::'a::{number_ring,real_normed_algebra_1})
+    = \<bar>number_of w\<bar>"
+by (subst of_real_number_of_eq [symmetric], rule norm_of_real)
+
+lemma norm_of_int [simp]:
+  "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
+by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
+
+lemma norm_of_nat [simp]:
+  "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
+apply (subst of_real_of_nat_eq [symmetric])
+apply (subst norm_of_real, simp)
+done
+
+lemma nonzero_norm_inverse:
+  fixes a :: "'a::real_normed_div_algebra"
+  shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
+apply (rule inverse_unique [symmetric])
+apply (simp add: norm_mult [symmetric])
+done
+
+lemma norm_inverse:
+  fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
+  shows "norm (inverse a) = inverse (norm a)"
+apply (case_tac "a = 0", simp)
+apply (erule nonzero_norm_inverse)
+done
+
+lemma nonzero_norm_divide:
+  fixes a b :: "'a::real_normed_field"
+  shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
+by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
+
+lemma norm_divide:
+  fixes a b :: "'a::{real_normed_field,division_by_zero}"
+  shows "norm (a / b) = norm a / norm b"
+by (simp add: divide_inverse norm_mult norm_inverse)
+
+lemma norm_power_ineq:
+  fixes x :: "'a::{real_normed_algebra_1,recpower}"
+  shows "norm (x ^ n) \<le> norm x ^ n"
+proof (induct n)
+  case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
+next
+  case (Suc n)
+  have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
+    by (rule norm_mult_ineq)
+  also from Suc have "\<dots> \<le> norm x * norm x ^ n"
+    using norm_ge_zero by (rule mult_left_mono)
+  finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
+    by (simp add: power_Suc)
+qed
+
+lemma norm_power:
+  fixes x :: "'a::{real_normed_div_algebra,recpower}"
+  shows "norm (x ^ n) = norm x ^ n"
+by (induct n) (simp_all add: power_Suc norm_mult)
+
+
+subsection {* Sign function *}
+
+lemma norm_sgn:
+  "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
+by (simp add: sgn_div_norm norm_scaleR)
+
+lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
+by (simp add: sgn_div_norm)
+
+lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
+by (simp add: sgn_div_norm)
+
+lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
+by (simp add: sgn_div_norm)
+
+lemma sgn_scaleR:
+  "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
+by (simp add: sgn_div_norm norm_scaleR mult_ac)
+
+lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
+by (simp add: sgn_div_norm)
+
+lemma sgn_of_real:
+  "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
+unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
+
+lemma sgn_mult:
+  fixes x y :: "'a::real_normed_div_algebra"
+  shows "sgn (x * y) = sgn x * sgn y"
+by (simp add: sgn_div_norm norm_mult mult_commute)
+
+lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
+by (simp add: sgn_div_norm divide_inverse)
+
+lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
+unfolding real_sgn_eq by simp
+
+lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
+unfolding real_sgn_eq by simp
+
+
+subsection {* Bounded Linear and Bilinear Operators *}
+
+locale bounded_linear = additive +
+  constrains f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
+  assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
+begin
+
+lemma pos_bounded:
+  "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
+proof -
+  obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
+    using bounded by fast
+  show ?thesis
+  proof (intro exI impI conjI allI)
+    show "0 < max 1 K"
+      by (rule order_less_le_trans [OF zero_less_one le_maxI1])
+  next
+    fix x
+    have "norm (f x) \<le> norm x * K" using K .
+    also have "\<dots> \<le> norm x * max 1 K"
+      by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
+    finally show "norm (f x) \<le> norm x * max 1 K" .
+  qed
+qed
+
+lemma nonneg_bounded:
+  "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
+proof -
+  from pos_bounded
+  show ?thesis by (auto intro: order_less_imp_le)
+qed
+
+end
+
+locale bounded_bilinear =
+  fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
+                 \<Rightarrow> 'c::real_normed_vector"
+    (infixl "**" 70)
+  assumes add_left: "prod (a + a') b = prod a b + prod a' b"
+  assumes add_right: "prod a (b + b') = prod a b + prod a b'"
+  assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
+  assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
+  assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
+begin
+
+lemma pos_bounded:
+  "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
+apply (cut_tac bounded, erule exE)
+apply (rule_tac x="max 1 K" in exI, safe)
+apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
+apply (drule spec, drule spec, erule order_trans)
+apply (rule mult_left_mono [OF le_maxI2])
+apply (intro mult_nonneg_nonneg norm_ge_zero)
+done
+
+lemma nonneg_bounded:
+  "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
+proof -
+  from pos_bounded
+  show ?thesis by (auto intro: order_less_imp_le)
+qed
+
+lemma additive_right: "additive (\<lambda>b. prod a b)"
+by (rule additive.intro, rule add_right)
+
+lemma additive_left: "additive (\<lambda>a. prod a b)"
+by (rule additive.intro, rule add_left)
+
+lemma zero_left: "prod 0 b = 0"
+by (rule additive.zero [OF additive_left])
+
+lemma zero_right: "prod a 0 = 0"
+by (rule additive.zero [OF additive_right])
+
+lemma minus_left: "prod (- a) b = - prod a b"
+by (rule additive.minus [OF additive_left])
+
+lemma minus_right: "prod a (- b) = - prod a b"
+by (rule additive.minus [OF additive_right])
+
+lemma diff_left:
+  "prod (a - a') b = prod a b - prod a' b"
+by (rule additive.diff [OF additive_left])
+
+lemma diff_right:
+  "prod a (b - b') = prod a b - prod a b'"
+by (rule additive.diff [OF additive_right])
+
+lemma bounded_linear_left:
+  "bounded_linear (\<lambda>a. a ** b)"
+apply (unfold_locales)
+apply (rule add_left)
+apply (rule scaleR_left)
+apply (cut_tac bounded, safe)
+apply (rule_tac x="norm b * K" in exI)
+apply (simp add: mult_ac)
+done
+
+lemma bounded_linear_right:
+  "bounded_linear (\<lambda>b. a ** b)"
+apply (unfold_locales)
+apply (rule add_right)
+apply (rule scaleR_right)
+apply (cut_tac bounded, safe)
+apply (rule_tac x="norm a * K" in exI)
+apply (simp add: mult_ac)
+done
+
+lemma prod_diff_prod:
+  "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
+by (simp add: diff_left diff_right)
+
+end
+
+interpretation mult!:
+  bounded_bilinear "op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"
+apply (rule bounded_bilinear.intro)
+apply (rule left_distrib)
+apply (rule right_distrib)
+apply (rule mult_scaleR_left)
+apply (rule mult_scaleR_right)
+apply (rule_tac x="1" in exI)
+apply (simp add: norm_mult_ineq)
+done
+
+interpretation mult_left!:
+  bounded_linear "(\<lambda>x::'a::real_normed_algebra. x * y)"
+by (rule mult.bounded_linear_left)
+
+interpretation mult_right!:
+  bounded_linear "(\<lambda>y::'a::real_normed_algebra. x * y)"
+by (rule mult.bounded_linear_right)
+
+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"
+apply (rule bounded_bilinear.intro)
+apply (rule scaleR_left_distrib)
+apply (rule scaleR_right_distrib)
+apply simp
+apply (rule scaleR_left_commute)
+apply (rule_tac x="1" in exI)
+apply (simp add: norm_scaleR)
+done
+
+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"
+by (rule scaleR.bounded_linear_right)
+
+interpretation of_real!: bounded_linear "\<lambda>r. of_real r"
+unfolding of_real_def by (rule scaleR.bounded_linear_left)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SEQ.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,1136 @@
+(*  Title       : SEQ.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Convergence of sequences and series
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+    Additional contributions by Jeremy Avigad and Brian Huffman
+*)
+
+header {* Sequences and Convergence *}
+
+theory SEQ
+imports RealVector RComplete
+begin
+
+definition
+  Zseq :: "[nat \<Rightarrow> 'a::real_normed_vector] \<Rightarrow> bool" where
+    --{*Standard definition of sequence converging to zero*}
+  [code del]: "Zseq X = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. norm (X n) < r)"
+
+definition
+  LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
+    ("((_)/ ----> (_))" [60, 60] 60) where
+    --{*Standard definition of convergence of sequence*}
+  [code del]: "X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))"
+
+definition
+  lim :: "(nat => 'a::real_normed_vector) => 'a" where
+    --{*Standard definition of limit using choice operator*}
+  "lim X = (THE L. X ----> L)"
+
+definition
+  convergent :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Standard definition of convergence*}
+  "convergent X = (\<exists>L. X ----> L)"
+
+definition
+  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Standard definition for bounded sequence*}
+  [code del]: "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
+
+definition
+  monoseq :: "(nat=>real)=>bool" where
+    --{*Definition for monotonicity*}
+  [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
+  subseq :: "(nat => nat) => bool" where
+    --{*Definition of subsequence*}
+  [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
+
+definition
+  Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Standard definition of the Cauchy condition*}
+  [code del]: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)"
+
+
+subsection {* Bounded Sequences *}
+
+lemma BseqI': assumes K: "\<And>n. norm (X n) \<le> K" shows "Bseq X"
+unfolding Bseq_def
+proof (intro exI conjI allI)
+  show "0 < max K 1" by simp
+next
+  fix n::nat
+  have "norm (X n) \<le> K" by (rule K)
+  thus "norm (X n) \<le> max K 1" by simp
+qed
+
+lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
+unfolding Bseq_def by auto
+
+lemma BseqI2': assumes K: "\<forall>n\<ge>N. norm (X n) \<le> K" shows "Bseq X"
+proof (rule BseqI')
+  let ?A = "norm ` X ` {..N}"
+  have 1: "finite ?A" by simp
+  fix n::nat
+  show "norm (X n) \<le> max K (Max ?A)"
+  proof (cases rule: linorder_le_cases)
+    assume "n \<ge> N"
+    hence "norm (X n) \<le> K" using K by simp
+    thus "norm (X n) \<le> max K (Max ?A)" by simp
+  next
+    assume "n \<le> N"
+    hence "norm (X n) \<in> ?A" by simp
+    with 1 have "norm (X n) \<le> Max ?A" by (rule Max_ge)
+    thus "norm (X n) \<le> max K (Max ?A)" by simp
+  qed
+qed
+
+lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
+unfolding Bseq_def by auto
+
+lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
+apply (erule BseqE)
+apply (rule_tac N="k" and K="K" in BseqI2')
+apply clarify
+apply (drule_tac x="n - k" in spec, simp)
+done
+
+
+subsection {* Sequences That Converge to Zero *}
+
+lemma ZseqI:
+  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r) \<Longrightarrow> Zseq X"
+unfolding Zseq_def by simp
+
+lemma ZseqD:
+  "\<lbrakk>Zseq X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r"
+unfolding Zseq_def by simp
+
+lemma Zseq_zero: "Zseq (\<lambda>n. 0)"
+unfolding Zseq_def by simp
+
+lemma Zseq_const_iff: "Zseq (\<lambda>n. k) = (k = 0)"
+unfolding Zseq_def by force
+
+lemma Zseq_norm_iff: "Zseq (\<lambda>n. norm (X n)) = Zseq (\<lambda>n. X n)"
+unfolding Zseq_def by simp
+
+lemma Zseq_imp_Zseq:
+  assumes X: "Zseq X"
+  assumes Y: "\<And>n. norm (Y n) \<le> norm (X n) * K"
+  shows "Zseq (\<lambda>n. Y n)"
+proof (cases)
+  assume K: "0 < K"
+  show ?thesis
+  proof (rule ZseqI)
+    fix r::real assume "0 < r"
+    hence "0 < r / K"
+      using K by (rule divide_pos_pos)
+    then obtain N where "\<forall>n\<ge>N. norm (X n) < r / K"
+      using ZseqD [OF X] by fast
+    hence "\<forall>n\<ge>N. norm (X n) * K < r"
+      by (simp add: pos_less_divide_eq K)
+    hence "\<forall>n\<ge>N. norm (Y n) < r"
+      by (simp add: order_le_less_trans [OF Y])
+    thus "\<exists>N. \<forall>n\<ge>N. norm (Y n) < r" ..
+  qed
+next
+  assume "\<not> 0 < K"
+  hence K: "K \<le> 0" by (simp only: linorder_not_less)
+  {
+    fix n::nat
+    have "norm (Y n) \<le> norm (X n) * K" by (rule Y)
+    also have "\<dots> \<le> norm (X n) * 0"
+      using K norm_ge_zero by (rule mult_left_mono)
+    finally have "norm (Y n) = 0" by simp
+  }
+  thus ?thesis by (simp add: Zseq_zero)
+qed
+
+lemma Zseq_le: "\<lbrakk>Zseq Y; \<forall>n. norm (X n) \<le> norm (Y n)\<rbrakk> \<Longrightarrow> Zseq X"
+by (erule_tac K="1" in Zseq_imp_Zseq, simp)
+
+lemma Zseq_add:
+  assumes X: "Zseq X"
+  assumes Y: "Zseq Y"
+  shows "Zseq (\<lambda>n. X n + Y n)"
+proof (rule ZseqI)
+  fix r::real assume "0 < r"
+  hence r: "0 < r / 2" by simp
+  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r/2"
+    using ZseqD [OF X r] by fast
+  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < r/2"
+    using ZseqD [OF Y r] by fast
+  show "\<exists>N. \<forall>n\<ge>N. norm (X n + Y n) < r"
+  proof (intro exI allI impI)
+    fix n assume n: "max M N \<le> n"
+    have "norm (X n + Y n) \<le> norm (X n) + norm (Y n)"
+      by (rule norm_triangle_ineq)
+    also have "\<dots> < r/2 + r/2"
+    proof (rule add_strict_mono)
+      from M n show "norm (X n) < r/2" by simp
+      from N n show "norm (Y n) < r/2" by simp
+    qed
+    finally show "norm (X n + Y n) < r" by simp
+  qed
+qed
+
+lemma Zseq_minus: "Zseq X \<Longrightarrow> Zseq (\<lambda>n. - X n)"
+unfolding Zseq_def by simp
+
+lemma Zseq_diff: "\<lbrakk>Zseq X; Zseq Y\<rbrakk> \<Longrightarrow> Zseq (\<lambda>n. X n - Y n)"
+by (simp only: diff_minus Zseq_add Zseq_minus)
+
+lemma (in bounded_linear) Zseq:
+  assumes X: "Zseq X"
+  shows "Zseq (\<lambda>n. f (X n))"
+proof -
+  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
+    using bounded by fast
+  with X show ?thesis
+    by (rule Zseq_imp_Zseq)
+qed
+
+lemma (in bounded_bilinear) Zseq:
+  assumes X: "Zseq X"
+  assumes Y: "Zseq Y"
+  shows "Zseq (\<lambda>n. X n ** Y n)"
+proof (rule ZseqI)
+  fix r::real assume r: "0 < r"
+  obtain K where K: "0 < K"
+    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
+    using pos_bounded by fast
+  from K have K': "0 < inverse K"
+    by (rule positive_imp_inverse_positive)
+  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r"
+    using ZseqD [OF X r] by fast
+  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < inverse K"
+    using ZseqD [OF Y K'] by fast
+  show "\<exists>N. \<forall>n\<ge>N. norm (X n ** Y n) < r"
+  proof (intro exI allI impI)
+    fix n assume n: "max M N \<le> n"
+    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
+      by (rule norm_le)
+    also have "norm (X n) * norm (Y n) * K < r * inverse K * K"
+    proof (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero K)
+      from M n show Xn: "norm (X n) < r" by simp
+      from N n show Yn: "norm (Y n) < inverse K" by simp
+    qed
+    also from K have "r * inverse K * K = r" by simp
+    finally show "norm (X n ** Y n) < r" .
+  qed
+qed
+
+lemma (in bounded_bilinear) Zseq_prod_Bseq:
+  assumes X: "Zseq X"
+  assumes Y: "Bseq Y"
+  shows "Zseq (\<lambda>n. X n ** Y n)"
+proof -
+  obtain K where K: "0 \<le> K"
+    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
+    using nonneg_bounded by fast
+  obtain B where B: "0 < B"
+    and norm_Y: "\<And>n. norm (Y n) \<le> B"
+    using Y [unfolded Bseq_def] by fast
+  from X show ?thesis
+  proof (rule Zseq_imp_Zseq)
+    fix n::nat
+    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
+      by (rule norm_le)
+    also have "\<dots> \<le> norm (X n) * B * K"
+      by (intro mult_mono' order_refl norm_Y norm_ge_zero
+                mult_nonneg_nonneg K)
+    also have "\<dots> = norm (X n) * (B * K)"
+      by (rule mult_assoc)
+    finally show "norm (X n ** Y n) \<le> norm (X n) * (B * K)" .
+  qed
+qed
+
+lemma (in bounded_bilinear) Bseq_prod_Zseq:
+  assumes X: "Bseq X"
+  assumes Y: "Zseq Y"
+  shows "Zseq (\<lambda>n. X n ** Y n)"
+proof -
+  obtain K where K: "0 \<le> K"
+    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
+    using nonneg_bounded by fast
+  obtain B where B: "0 < B"
+    and norm_X: "\<And>n. norm (X n) \<le> B"
+    using X [unfolded Bseq_def] by fast
+  from Y show ?thesis
+  proof (rule Zseq_imp_Zseq)
+    fix n::nat
+    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
+      by (rule norm_le)
+    also have "\<dots> \<le> B * norm (Y n) * K"
+      by (intro mult_mono' order_refl norm_X norm_ge_zero
+                mult_nonneg_nonneg K)
+    also have "\<dots> = norm (Y n) * (B * K)"
+      by (simp only: mult_ac)
+    finally show "norm (X n ** Y n) \<le> norm (Y n) * (B * K)" .
+  qed
+qed
+
+lemma (in bounded_bilinear) Zseq_left:
+  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. X n ** a)"
+by (rule bounded_linear_left [THEN bounded_linear.Zseq])
+
+lemma (in bounded_bilinear) Zseq_right:
+  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. a ** X n)"
+by (rule bounded_linear_right [THEN bounded_linear.Zseq])
+
+lemmas Zseq_mult = mult.Zseq
+lemmas Zseq_mult_right = mult.Zseq_right
+lemmas Zseq_mult_left = mult.Zseq_left
+
+
+subsection {* Limits of Sequences *}
+
+lemma LIMSEQ_iff:
+      "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
+by (rule LIMSEQ_def)
+
+lemma LIMSEQ_Zseq_iff: "((\<lambda>n. X n) ----> L) = Zseq (\<lambda>n. X n - L)"
+by (simp only: LIMSEQ_def Zseq_def)
+
+lemma LIMSEQ_I:
+  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_D:
+  "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_const_iff: "(\<lambda>n. k) ----> l = (k = l)"
+by (simp add: LIMSEQ_Zseq_iff Zseq_const_iff)
+
+lemma LIMSEQ_norm: "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
+apply (simp add: LIMSEQ_def, safe)
+apply (drule_tac x="r" in spec, safe)
+apply (rule_tac x="no" in exI, safe)
+apply (drule_tac x="n" in spec, safe)
+apply (erule order_le_less_trans [OF norm_triangle_ineq3])
+done
+
+lemma LIMSEQ_ignore_initial_segment:
+  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
+apply (rule LIMSEQ_I)
+apply (drule (1) LIMSEQ_D)
+apply (erule exE, rename_tac N)
+apply (rule_tac x=N in exI)
+apply simp
+done
+
+lemma LIMSEQ_offset:
+  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
+apply (rule LIMSEQ_I)
+apply (drule (1) LIMSEQ_D)
+apply (erule exE, rename_tac N)
+apply (rule_tac x="N + k" in exI)
+apply clarify
+apply (drule_tac x="n - k" in spec)
+apply (simp add: le_diff_conv2)
+done
+
+lemma LIMSEQ_Suc: "f ----> l \<Longrightarrow> (\<lambda>n. f (Suc n)) ----> l"
+by (drule_tac k="1" in LIMSEQ_ignore_initial_segment, simp)
+
+lemma LIMSEQ_imp_Suc: "(\<lambda>n. f (Suc n)) ----> l \<Longrightarrow> f ----> l"
+by (rule_tac k="1" in LIMSEQ_offset, simp)
+
+lemma LIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l"
+by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
+
+lemma add_diff_add:
+  fixes a b c d :: "'a::ab_group_add"
+  shows "(a + c) - (b + d) = (a - b) + (c - d)"
+by simp
+
+lemma minus_diff_minus:
+  fixes a b :: "'a::ab_group_add"
+  shows "(- a) - (- b) = - (a - b)"
+by simp
+
+lemma LIMSEQ_add: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
+by (simp only: LIMSEQ_Zseq_iff add_diff_add Zseq_add)
+
+lemma LIMSEQ_minus: "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
+by (simp only: LIMSEQ_Zseq_iff minus_diff_minus Zseq_minus)
+
+lemma LIMSEQ_minus_cancel: "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
+by (drule LIMSEQ_minus, simp)
+
+lemma LIMSEQ_diff: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
+by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus)
+
+lemma LIMSEQ_unique: "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
+by (drule (1) LIMSEQ_diff, simp add: LIMSEQ_const_iff)
+
+lemma (in bounded_linear) LIMSEQ:
+  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
+by (simp only: LIMSEQ_Zseq_iff diff [symmetric] Zseq)
+
+lemma (in bounded_bilinear) LIMSEQ:
+  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n ** Y n) ----> a ** b"
+by (simp only: LIMSEQ_Zseq_iff prod_diff_prod
+               Zseq_add Zseq Zseq_left Zseq_right)
+
+lemma LIMSEQ_mult:
+  fixes a b :: "'a::real_normed_algebra"
+  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
+by (rule mult.LIMSEQ)
+
+lemma inverse_diff_inverse:
+  "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
+   \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
+by (simp add: ring_simps)
+
+lemma Bseq_inverse_lemma:
+  fixes x :: "'a::real_normed_div_algebra"
+  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
+apply (subst nonzero_norm_inverse, clarsimp)
+apply (erule (1) le_imp_inverse_le)
+done
+
+lemma Bseq_inverse:
+  fixes a :: "'a::real_normed_div_algebra"
+  assumes X: "X ----> a"
+  assumes a: "a \<noteq> 0"
+  shows "Bseq (\<lambda>n. inverse (X n))"
+proof -
+  from a have "0 < norm a" by simp
+  hence "\<exists>r>0. r < norm a" by (rule dense)
+  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
+  obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (X n - a) < r"
+    using LIMSEQ_D [OF X r1] by fast
+  show ?thesis
+  proof (rule BseqI2' [rule_format])
+    fix n assume n: "N \<le> n"
+    hence 1: "norm (X n - a) < r" by (rule N)
+    hence 2: "X n \<noteq> 0" using r2 by auto
+    hence "norm (inverse (X n)) = inverse (norm (X n))"
+      by (rule nonzero_norm_inverse)
+    also have "\<dots> \<le> inverse (norm a - r)"
+    proof (rule le_imp_inverse_le)
+      show "0 < norm a - r" using r2 by simp
+    next
+      have "norm a - norm (X n) \<le> norm (a - X n)"
+        by (rule norm_triangle_ineq2)
+      also have "\<dots> = norm (X n - a)"
+        by (rule norm_minus_commute)
+      also have "\<dots> < r" using 1 .
+      finally show "norm a - r \<le> norm (X n)" by simp
+    qed
+    finally show "norm (inverse (X n)) \<le> inverse (norm a - r)" .
+  qed
+qed
+
+lemma LIMSEQ_inverse_lemma:
+  fixes a :: "'a::real_normed_div_algebra"
+  shows "\<lbrakk>X ----> a; a \<noteq> 0; \<forall>n. X n \<noteq> 0\<rbrakk>
+         \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
+apply (subst LIMSEQ_Zseq_iff)
+apply (simp add: inverse_diff_inverse nonzero_imp_inverse_nonzero)
+apply (rule Zseq_minus)
+apply (rule Zseq_mult_left)
+apply (rule mult.Bseq_prod_Zseq)
+apply (erule (1) Bseq_inverse)
+apply (simp add: LIMSEQ_Zseq_iff)
+done
+
+lemma LIMSEQ_inverse:
+  fixes a :: "'a::real_normed_div_algebra"
+  assumes X: "X ----> a"
+  assumes a: "a \<noteq> 0"
+  shows "(\<lambda>n. inverse (X n)) ----> inverse a"
+proof -
+  from a have "0 < norm a" by simp
+  then obtain k where "\<forall>n\<ge>k. norm (X n - a) < norm a"
+    using LIMSEQ_D [OF X] by fast
+  hence "\<forall>n\<ge>k. X n \<noteq> 0" by auto
+  hence k: "\<forall>n. X (n + k) \<noteq> 0" by simp
+
+  from X have "(\<lambda>n. X (n + k)) ----> a"
+    by (rule LIMSEQ_ignore_initial_segment)
+  hence "(\<lambda>n. inverse (X (n + k))) ----> inverse a"
+    using a k by (rule LIMSEQ_inverse_lemma)
+  thus "(\<lambda>n. inverse (X n)) ----> inverse a"
+    by (rule LIMSEQ_offset)
+qed
+
+lemma LIMSEQ_divide:
+  fixes a b :: "'a::real_normed_field"
+  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
+by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)
+
+lemma LIMSEQ_pow:
+  fixes a :: "'a::{real_normed_algebra,recpower}"
+  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
+by (induct m) (simp_all add: power_Suc LIMSEQ_const LIMSEQ_mult)
+
+lemma LIMSEQ_setsum:
+  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
+  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
+proof (cases "finite S")
+  case True
+  thus ?thesis using n
+  proof (induct)
+    case empty
+    show ?case
+      by (simp add: LIMSEQ_const)
+  next
+    case insert
+    thus ?case
+      by (simp add: LIMSEQ_add)
+  qed
+next
+  case False
+  thus ?thesis
+    by (simp add: LIMSEQ_const)
+qed
+
+lemma LIMSEQ_setprod:
+  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
+  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
+  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)"
+proof (cases "finite S")
+  case True
+  thus ?thesis using n
+  proof (induct)
+    case empty
+    show ?case
+      by (simp add: LIMSEQ_const)
+  next
+    case insert
+    thus ?case
+      by (simp add: LIMSEQ_mult)
+  qed
+next
+  case False
+  thus ?thesis
+    by (simp add: setprod_def LIMSEQ_const)
+qed
+
+lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b"
+by (simp add: LIMSEQ_add LIMSEQ_const)
+
+(* FIXME: delete *)
+lemma LIMSEQ_add_minus:
+     "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
+by (simp only: LIMSEQ_add LIMSEQ_minus)
+
+lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n  - b)) ----> a - b"
+by (simp add: LIMSEQ_diff LIMSEQ_const)
+
+lemma LIMSEQ_diff_approach_zero: 
+  "g ----> L ==> (%x. f x - g x) ----> 0  ==>
+     f ----> L"
+  apply (drule LIMSEQ_add)
+  apply assumption
+  apply simp
+done
+
+lemma LIMSEQ_diff_approach_zero2: 
+  "f ----> L ==> (%x. f x - g x) ----> 0  ==>
+     g ----> L";
+  apply (drule LIMSEQ_diff)
+  apply assumption
+  apply simp
+done
+
+text{*A sequence tends to zero iff its abs does*}
+lemma LIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----> 0) = (X ----> 0)"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
+by (drule LIMSEQ_norm, simp)
+
+text{*An unbounded sequence's inverse tends to 0*}
+
+lemma LIMSEQ_inverse_zero:
+  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
+apply (rule LIMSEQ_I)
+apply (drule_tac x="inverse r" in spec, safe)
+apply (rule_tac x="N" in exI, safe)
+apply (drule_tac x="n" in spec, safe)
+apply (frule positive_imp_inverse_positive)
+apply (frule (1) less_imp_inverse_less)
+apply (subgoal_tac "0 < X n", simp)
+apply (erule (1) order_less_trans)
+done
+
+text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
+
+lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
+apply (rule LIMSEQ_inverse_zero, safe)
+apply (cut_tac x = r in reals_Archimedean2)
+apply (safe, rule_tac x = n in exI)
+apply (auto simp add: real_of_nat_Suc)
+done
+
+text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
+infinity is now easily proved*}
+
+lemma LIMSEQ_inverse_real_of_nat_add:
+     "(%n. r + inverse(real(Suc n))) ----> r"
+by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
+
+lemma LIMSEQ_inverse_real_of_nat_add_minus:
+     "(%n. r + -inverse(real(Suc n))) ----> r"
+by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
+
+lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
+     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
+by (cut_tac b=1 in
+        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)
+
+lemma LIMSEQ_le_const:
+  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
+apply (rule ccontr, simp only: linorder_not_le)
+apply (drule_tac r="a - x" in LIMSEQ_D, simp)
+apply clarsimp
+apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
+apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
+apply simp
+done
+
+lemma LIMSEQ_le_const2:
+  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
+apply (subgoal_tac "- a \<le> - x", simp)
+apply (rule LIMSEQ_le_const)
+apply (erule LIMSEQ_minus)
+apply simp
+done
+
+lemma LIMSEQ_le:
+  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
+apply (subgoal_tac "0 \<le> y - x", simp)
+apply (rule LIMSEQ_le_const)
+apply (erule (1) LIMSEQ_diff)
+apply (simp add: le_diff_eq)
+done
+
+
+subsection {* Convergence *}
+
+lemma limI: "X ----> L ==> lim X = L"
+apply (simp add: lim_def)
+apply (blast intro: LIMSEQ_unique)
+done
+
+lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
+by (simp add: convergent_def)
+
+lemma convergentI: "(X ----> L) ==> convergent X"
+by (auto simp add: convergent_def)
+
+lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
+by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def)
+
+lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))"
+apply (simp add: convergent_def)
+apply (auto dest: LIMSEQ_minus)
+apply (drule LIMSEQ_minus, auto)
+done
+
+
+subsection {* Bounded Monotonic Sequences *}
+
+text{*Subsequence (alternative definition, (e.g. Hoskins)*}
+
+lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
+apply (simp add: subseq_def)
+apply (auto dest!: less_imp_Suc_add)
+apply (induct_tac k)
+apply (auto intro: less_trans)
+done
+
+lemma monoseq_Suc:
+   "monoseq X = ((\<forall>n. X n \<le> X (Suc n))
+                 | (\<forall>n. X (Suc n) \<le> X n))"
+apply (simp add: monoseq_def)
+apply (auto dest!: le_imp_less_or_eq)
+apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
+apply (induct_tac "ka")
+apply (auto intro: order_trans)
+apply (erule contrapos_np)
+apply (induct_tac "k")
+apply (auto intro: order_trans)
+done
+
+lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
+by (simp add: monoseq_def)
+
+lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
+by (simp add: monoseq_def)
+
+lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
+by (simp add: monoseq_Suc)
+
+lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
+by (simp add: monoseq_Suc)
+
+text{*Bounded Sequence*}
+
+lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
+by (simp add: Bseq_def)
+
+lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
+by (auto simp add: Bseq_def)
+
+lemma lemma_NBseq_def:
+     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) =
+      (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
+apply auto
+ prefer 2 apply force
+apply (cut_tac x = K in reals_Archimedean2, clarify)
+apply (rule_tac x = n in exI, clarify)
+apply (drule_tac x = na in spec)
+apply (auto simp add: real_of_nat_Suc)
+done
+
+text{* alternative definition for Bseq *}
+lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
+apply (simp add: Bseq_def)
+apply (simp (no_asm) add: lemma_NBseq_def)
+done
+
+lemma lemma_NBseq_def2:
+     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
+apply (subst lemma_NBseq_def, auto)
+apply (rule_tac x = "Suc N" in exI)
+apply (rule_tac [2] x = N in exI)
+apply (auto simp add: real_of_nat_Suc)
+ prefer 2 apply (blast intro: order_less_imp_le)
+apply (drule_tac x = n in spec, simp)
+done
+
+(* yet another definition for Bseq *)
+lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
+by (simp add: Bseq_def lemma_NBseq_def2)
+
+subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
+
+lemma Bseq_isUb:
+  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
+by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
+
+
+text{* Use completeness of reals (supremum property)
+   to show that any bounded sequence has a least upper bound*}
+
+lemma Bseq_isLub:
+  "!!(X::nat=>real). Bseq X ==>
+   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
+by (blast intro: reals_complete Bseq_isUb)
+
+subsubsection{*A Bounded and Monotonic Sequence Converges*}
+
+lemma lemma_converg1:
+     "!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n;
+                  isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma)
+               |] ==> \<forall>n \<ge> ma. X n = X ma"
+apply safe
+apply (drule_tac y = "X n" in isLubD2)
+apply (blast dest: order_antisym)+
+done
+
+text{* The best of both worlds: Easier to prove this result as a standard
+   theorem and then use equivalence to "transfer" it into the
+   equivalent nonstandard form if needed!*}
+
+lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
+apply (simp add: LIMSEQ_def)
+apply (rule_tac x = "X m" in exI, safe)
+apply (rule_tac x = m in exI, safe)
+apply (drule spec, erule impE, auto)
+done
+
+lemma lemma_converg2:
+   "!!(X::nat=>real).
+    [| \<forall>m. X m ~= U;  isLub UNIV {x. \<exists>n. X n = x} U |] ==> \<forall>m. X m < U"
+apply safe
+apply (drule_tac y = "X m" in isLubD2)
+apply (auto dest!: order_le_imp_less_or_eq)
+done
+
+lemma lemma_converg3: "!!(X ::nat=>real). \<forall>m. X m \<le> U ==> isUb UNIV {x. \<exists>n. X n = x} U"
+by (rule setleI [THEN isUbI], auto)
+
+text{* FIXME: @{term "U - T < U"} is redundant *}
+lemma lemma_converg4: "!!(X::nat=> real).
+               [| \<forall>m. X m ~= U;
+                  isLub UNIV {x. \<exists>n. X n = x} U;
+                  0 < T;
+                  U + - T < U
+               |] ==> \<exists>m. U + -T < X m & X m < U"
+apply (drule lemma_converg2, assumption)
+apply (rule ccontr, simp)
+apply (simp add: linorder_not_less)
+apply (drule lemma_converg3)
+apply (drule isLub_le_isUb, assumption)
+apply (auto dest: order_less_le_trans)
+done
+
+text{*A standard proof of the theorem for monotone increasing sequence*}
+
+lemma Bseq_mono_convergent:
+     "[| Bseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> convergent (X::nat=>real)"
+apply (simp add: convergent_def)
+apply (frule Bseq_isLub, safe)
+apply (case_tac "\<exists>m. X m = U", auto)
+apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ)
+(* second case *)
+apply (rule_tac x = U in exI)
+apply (subst LIMSEQ_iff, safe)
+apply (frule lemma_converg2, assumption)
+apply (drule lemma_converg4, auto)
+apply (rule_tac x = m in exI, safe)
+apply (subgoal_tac "X m \<le> X n")
+ prefer 2 apply blast
+apply (drule_tac x=n and P="%m. X m < U" in spec, arith)
+done
+
+lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
+by (simp add: Bseq_def)
+
+text{*Main monotonicity theorem*}
+lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X"
+apply (simp add: monoseq_def, safe)
+apply (rule_tac [2] convergent_minus_iff [THEN ssubst])
+apply (drule_tac [2] Bseq_minus_iff [THEN ssubst])
+apply (auto intro!: Bseq_mono_convergent)
+done
+
+subsubsection{*A Few More Equivalence Theorems for Boundedness*}
+
+text{*alternative formulation for boundedness*}
+lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
+apply (unfold Bseq_def, safe)
+apply (rule_tac [2] x = "k + norm x" in exI)
+apply (rule_tac x = K in exI, simp)
+apply (rule exI [where x = 0], auto)
+apply (erule order_less_le_trans, simp)
+apply (drule_tac x=n in spec, fold diff_def)
+apply (drule order_trans [OF norm_triangle_ineq2])
+apply simp
+done
+
+text{*alternative formulation for boundedness*}
+lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
+apply safe
+apply (simp add: Bseq_def, safe)
+apply (rule_tac x = "K + norm (X N)" in exI)
+apply auto
+apply (erule order_less_le_trans, simp)
+apply (rule_tac x = N in exI, safe)
+apply (drule_tac x = n in spec)
+apply (rule order_trans [OF norm_triangle_ineq], simp)
+apply (auto simp add: Bseq_iff2)
+done
+
+lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
+apply (simp add: Bseq_def)
+apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
+apply (drule_tac x = n in spec, arith)
+done
+
+
+subsection {* Cauchy Sequences *}
+
+lemma CauchyI:
+  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
+by (simp add: Cauchy_def)
+
+lemma CauchyD:
+  "\<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)
+
+subsubsection {* Cauchy Sequences are Bounded *}
+
+text{*A Cauchy sequence is bounded -- this is the standard
+  proof mechanization rather than the nonstandard proof*}
+
+lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
+          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
+apply (clarify, drule spec, drule (1) mp)
+apply (simp only: norm_minus_commute)
+apply (drule order_le_less_trans [OF norm_triangle_ineq2])
+apply simp
+done
+
+lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
+apply (simp add: Cauchy_def)
+apply (drule spec, drule mp, rule zero_less_one, safe)
+apply (drule_tac x="M" in spec, simp)
+apply (drule lemmaCauchy)
+apply (rule_tac k="M" in Bseq_offset)
+apply (simp add: Bseq_def)
+apply (rule_tac x="1 + norm (X M)" in exI)
+apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp)
+apply (simp add: order_less_imp_le)
+done
+
+subsubsection {* Cauchy Sequences are Convergent *}
+
+axclass banach \<subseteq> real_normed_vector
+  Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
+
+theorem LIMSEQ_imp_Cauchy:
+  assumes X: "X ----> a" shows "Cauchy X"
+proof (rule CauchyI)
+  fix e::real assume "0 < e"
+  hence "0 < e/2" by simp
+  with X have "\<exists>N. \<forall>n\<ge>N. norm (X n - a) < e/2" by (rule LIMSEQ_D)
+  then obtain N where N: "\<forall>n\<ge>N. norm (X n - a) < e/2" ..
+  show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < e"
+  proof (intro exI allI impI)
+    fix m assume "N \<le> m"
+    hence m: "norm (X m - a) < e/2" using N by fast
+    fix n assume "N \<le> n"
+    hence n: "norm (X n - a) < e/2" using N by fast
+    have "norm (X m - X n) = norm ((X m - a) - (X n - a))" by simp
+    also have "\<dots> \<le> norm (X m - a) + norm (X n - a)"
+      by (rule norm_triangle_ineq4)
+    also from m n have "\<dots> < e" by(simp add:field_simps)
+    finally show "norm (X m - X n) < e" .
+  qed
+qed
+
+lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
+unfolding convergent_def
+by (erule exE, erule LIMSEQ_imp_Cauchy)
+
+text {*
+Proof that Cauchy sequences converge based on the one from
+http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
+*}
+
+text {*
+  If sequence @{term "X"} is Cauchy, then its limit is the lub of
+  @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
+*}
+
+lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
+by (simp add: isUbI setleI)
+
+lemma real_abs_diff_less_iff:
+  "(\<bar>x - a\<bar> < (r::real)) = (a - r < x \<and> x < a + r)"
+by auto
+
+locale real_Cauchy =
+  fixes X :: "nat \<Rightarrow> real"
+  assumes X: "Cauchy X"
+  fixes S :: "real set"
+  defines S_def: "S \<equiv> {x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
+
+lemma real_CauchyI:
+  assumes "Cauchy X"
+  shows "real_Cauchy X"
+  proof qed (fact assms)
+
+lemma (in real_Cauchy) mem_S: "\<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S"
+by (unfold S_def, auto)
+
+lemma (in real_Cauchy) bound_isUb:
+  assumes N: "\<forall>n\<ge>N. X n < x"
+  shows "isUb UNIV S x"
+proof (rule isUb_UNIV_I)
+  fix y::real assume "y \<in> S"
+  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
+    by (simp add: S_def)
+  then obtain M where "\<forall>n\<ge>M. y < X n" ..
+  hence "y < X (max M N)" by simp
+  also have "\<dots> < x" using N by simp
+  finally show "y \<le> x"
+    by (rule order_less_imp_le)
+qed
+
+lemma (in real_Cauchy) isLub_ex: "\<exists>u. isLub UNIV S u"
+proof (rule reals_complete)
+  obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < 1"
+    using CauchyD [OF X zero_less_one] by fast
+  hence N: "\<forall>n\<ge>N. norm (X n - X N) < 1" by simp
+  show "\<exists>x. x \<in> S"
+  proof
+    from N have "\<forall>n\<ge>N. X N - 1 < X n"
+      by (simp add: real_abs_diff_less_iff)
+    thus "X N - 1 \<in> S" by (rule mem_S)
+  qed
+  show "\<exists>u. isUb UNIV S u"
+  proof
+    from N have "\<forall>n\<ge>N. X n < X N + 1"
+      by (simp add: real_abs_diff_less_iff)
+    thus "isUb UNIV S (X N + 1)"
+      by (rule bound_isUb)
+  qed
+qed
+
+lemma (in real_Cauchy) isLub_imp_LIMSEQ:
+  assumes x: "isLub UNIV S x"
+  shows "X ----> x"
+proof (rule LIMSEQ_I)
+  fix r::real assume "0 < r"
+  hence r: "0 < r/2" by simp
+  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. norm (X n - X m) < r/2"
+    using CauchyD [OF X r] by fast
+  hence "\<forall>n\<ge>N. norm (X n - X N) < r/2" by simp
+  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
+    by (simp only: real_norm_def real_abs_diff_less_iff)
+
+  from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
+  hence "X N - r/2 \<in> S" by (rule mem_S)
+  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
+
+  from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
+  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
+  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
+
+  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
+  proof (intro exI allI impI)
+    fix n assume n: "N \<le> n"
+    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
+    thus "norm (X n - x) < r" using 1 2
+      by (simp add: real_abs_diff_less_iff)
+  qed
+qed
+
+lemma (in real_Cauchy) LIMSEQ_ex: "\<exists>x. X ----> x"
+proof -
+  obtain x where "isLub UNIV S x"
+    using isLub_ex by fast
+  hence "X ----> x"
+    by (rule isLub_imp_LIMSEQ)
+  thus ?thesis ..
+qed
+
+lemma real_Cauchy_convergent:
+  fixes X :: "nat \<Rightarrow> real"
+  shows "Cauchy X \<Longrightarrow> convergent X"
+unfolding convergent_def
+by (rule real_Cauchy.LIMSEQ_ex)
+ (rule real_CauchyI)
+
+instance real :: banach
+by intro_classes (rule real_Cauchy_convergent)
+
+lemma Cauchy_convergent_iff:
+  fixes X :: "nat \<Rightarrow> 'a::banach"
+  shows "Cauchy X = convergent X"
+by (fast intro: Cauchy_convergent convergent_Cauchy)
+
+
+subsection {* Power Sequences *}
+
+text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
+"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
+  also fact that bounded and monotonic sequence converges.*}
+
+lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
+apply (simp add: Bseq_def)
+apply (rule_tac x = 1 in exI)
+apply (simp add: power_abs)
+apply (auto dest: power_mono)
+done
+
+lemma monoseq_realpow: "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
+apply (clarify intro!: mono_SucI2)
+apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
+done
+
+lemma convergent_realpow:
+  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
+by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
+
+lemma LIMSEQ_inverse_realpow_zero_lemma:
+  fixes x :: real
+  assumes x: "0 \<le> x"
+  shows "real n * x + 1 \<le> (x + 1) ^ n"
+apply (induct n)
+apply simp
+apply simp
+apply (rule order_trans)
+prefer 2
+apply (erule mult_left_mono)
+apply (rule add_increasing [OF x], simp)
+apply (simp add: real_of_nat_Suc)
+apply (simp add: ring_distribs)
+apply (simp add: mult_nonneg_nonneg x)
+done
+
+lemma LIMSEQ_inverse_realpow_zero:
+  "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
+proof (rule LIMSEQ_inverse_zero [rule_format])
+  fix y :: real
+  assume x: "1 < x"
+  hence "0 < x - 1" by simp
+  hence "\<forall>y. \<exists>N::nat. y < real N * (x - 1)"
+    by (rule reals_Archimedean3)
+  hence "\<exists>N::nat. y < real N * (x - 1)" ..
+  then obtain N::nat where "y < real N * (x - 1)" ..
+  also have "\<dots> \<le> real N * (x - 1) + 1" by simp
+  also have "\<dots> \<le> (x - 1 + 1) ^ N"
+    by (rule LIMSEQ_inverse_realpow_zero_lemma, cut_tac x, simp)
+  also have "\<dots> = x ^ N" by simp
+  finally have "y < x ^ N" .
+  hence "\<forall>n\<ge>N. y < x ^ n"
+    apply clarify
+    apply (erule order_less_le_trans)
+    apply (erule power_increasing)
+    apply (rule order_less_imp_le [OF x])
+    done
+  thus "\<exists>N. \<forall>n\<ge>N. y < x ^ n" ..
+qed
+
+lemma LIMSEQ_realpow_zero:
+  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
+proof (cases)
+  assume "x = 0"
+  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
+  thus ?thesis by (rule LIMSEQ_imp_Suc)
+next
+  assume "0 \<le> x" and "x \<noteq> 0"
+  hence x0: "0 < x" by simp
+  assume x1: "x < 1"
+  from x0 x1 have "1 < inverse x"
+    by (rule real_inverse_gt_one)
+  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
+    by (rule LIMSEQ_inverse_realpow_zero)
+  thus ?thesis by (simp add: power_inverse)
+qed
+
+lemma LIMSEQ_power_zero:
+  fixes x :: "'a::{real_normed_algebra_1,recpower}"
+  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
+apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
+apply (simp only: LIMSEQ_Zseq_iff, erule Zseq_le)
+apply (simp add: power_abs norm_power_ineq)
+done
+
+lemma LIMSEQ_divide_realpow_zero:
+  "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0"
+apply (cut_tac a = a and x1 = "inverse x" in
+        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero])
+apply (auto simp add: divide_inverse power_inverse)
+apply (simp add: inverse_eq_divide pos_divide_less_eq)
+done
+
+text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
+
+lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----> 0"
+by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
+
+lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----> 0"
+apply (rule LIMSEQ_rabs_zero [THEN iffD1])
+apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs)
+done
+
+end
--- a/src/HOL/Series.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Series.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -10,7 +10,7 @@
 header{*Finite Summation and Infinite Series*}
 
 theory Series
-imports "~~/src/HOL/Hyperreal/SEQ"
+imports SEQ
 begin
 
 definition
--- a/src/HOL/Statespace/DistinctTreeProver.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Statespace/DistinctTreeProver.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
-(*  Title:      DistinctTreeProver.thy
-    ID:         $Id$
+(*  Title:      HOL/Statespace/DistinctTreeProver.thy
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
@@ -453,7 +452,7 @@
   case Tip thus ?case by simp
 next
   case (Node l x d r)
-  have sub: "subtract (Node l x d r) t\<^isub>2 = Some t".
+  have sub: "subtract (Node l x d r) t\<^isub>2 = Some t" by fact
   show ?case
   proof (cases "delete x t\<^isub>2")
     case (Some t\<^isub>2')
@@ -646,9 +645,9 @@
 val const_decls = map (fn i => Syntax.no_syn 
                                  ("const" ^ string_of_int i,Type ("nat",[]))) nums
 
-val consts = sort Term.fast_term_ord 
+val consts = sort TermOrd.fast_term_ord 
               (map (fn i => Const ("DistinctTreeProver.const"^string_of_int i,Type ("nat",[]))) nums)
-val consts' = sort Term.fast_term_ord 
+val consts' = sort TermOrd.fast_term_ord 
               (map (fn i => Const ("DistinctTreeProver.const"^string_of_int i,Type ("nat",[]))) nums')
 
 val t = DistinctTreeProver.mk_tree I (Type ("nat",[])) consts
--- a/src/HOL/Statespace/StateSpaceEx.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Statespace/StateSpaceEx.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -30,6 +30,10 @@
   n::nat
   b::bool
 
+print_locale vars_namespace
+print_locale vars_valuetypes
+print_locale vars
+
 text {* \noindent This resembles a \isacommand{record} definition, 
 but introduces sophisticated locale
 infrastructure instead of HOL type schemes.  The resulting context
@@ -37,7 +41,7 @@
 projection~/ injection functions that convert from abstract values to
 @{typ "nat"} and @{text "bool"}. The logical content of the locale is: *}
 
-locale vars' =
+class_locale vars' =
   fixes n::'name and b::'name
   assumes "distinct [n, b]" 
 
@@ -196,10 +200,19 @@
   by simp
 
 
+text {* Hmm, I hoped this would work now...*}
+
+(*
+locale fooX = foo +
+ assumes "s<a:=i>\<cdot>b = k"
+*)
+
+(* ++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ *)
 text {* There are known problems with syntax-declarations. They currently
 only work, when the context is already built. Hopefully this will be 
 implemented correctly in future Isabelle versions. *}
 
+(*
 lemma 
   assumes "foo f a b c p1 i1 p2 i2 p3 i3 p4 i4"
   shows True
@@ -207,7 +220,7 @@
   interpret foo [f a b c p1 i1 p2 i2 p3 i3 p4 i4] by fact
   term "s<a := i>\<cdot>a = i"
 qed
-
+*)
 (*
 lemma 
   includes foo
--- a/src/HOL/Statespace/distinct_tree_prover.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Statespace/distinct_tree_prover.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,28 +1,24 @@
-(*  Title:      distinct_tree_prover.thy
-    ID:         $Id$
+(*  Title:      HOL/Statespace/distinct_tree_prover.ML
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
 signature DISTINCT_TREE_PROVER =
 sig
   datatype Direction = Left | Right
-  val mk_tree : ('a -> Term.term) -> Term.typ -> 'a list -> Term.term
-  val dest_tree : Term.term -> Term.term list
-  val find_tree :
-       Term.term -> Term.term -> Direction list option 
+  val mk_tree : ('a -> term) -> typ -> 'a list -> term
+  val dest_tree : term -> term list
+  val find_tree : term -> term -> Direction list option 
 
-  val neq_to_eq_False : Thm.thm
-  val distinctTreeProver : Thm.thm -> Direction list -> Direction list -> Thm.thm
-  val neq_x_y :
-       Proof.context -> Term.term -> Term.term -> string -> Thm.thm option   
-  val distinctFieldSolver : string list -> MetaSimplifier.solver
-  val distinctTree_tac :
-       string list -> Proof.context -> Term.term * int -> Tactical.tactic
-  val distinct_implProver : Thm.thm -> Thm.cterm -> Thm.thm
-  val subtractProver : Term.term -> Thm.cterm -> Thm.thm -> Thm.thm
-  val distinct_simproc : string list -> MetaSimplifier.simproc
+  val neq_to_eq_False : thm
+  val distinctTreeProver : thm -> Direction list -> Direction list -> thm
+  val neq_x_y : Proof.context -> term -> term -> string -> thm option   
+  val distinctFieldSolver : string list -> solver
+  val distinctTree_tac : string list -> Proof.context -> term * int -> tactic
+  val distinct_implProver : thm -> cterm -> thm
+  val subtractProver : term -> cterm -> thm -> thm
+  val distinct_simproc : string list -> simproc
   
-  val discharge : Thm.thm list -> Thm.thm -> Thm.thm
+  val discharge : thm list -> thm -> thm
 end;
 
 structure DistinctTreeProver : DISTINCT_TREE_PROVER =
@@ -83,7 +79,7 @@
   | bin_find_tree order e t = raise TERM ("find_tree: input not a tree",[t])
 
 fun find_tree e t =
-  (case bin_find_tree Term.fast_term_ord e t of
+  (case bin_find_tree TermOrd.fast_term_ord e t of
      NONE => lin_find_tree e t
    | x => x);
 
--- a/src/HOL/Statespace/state_fun.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Statespace/state_fun.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Statespace/state_fun.ML
-    ID:         $Id$
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
@@ -8,17 +7,17 @@
   val lookupN : string
   val updateN : string
 
-  val mk_constr : Context.theory -> Term.typ -> Term.term
-  val mk_destr : Context.theory -> Term.typ -> Term.term
+  val mk_constr : theory -> typ -> term
+  val mk_destr : theory -> typ -> term
 
-  val lookup_simproc : MetaSimplifier.simproc
-  val update_simproc : MetaSimplifier.simproc
-  val ex_lookup_eq_simproc : MetaSimplifier.simproc
-  val ex_lookup_ss : MetaSimplifier.simpset
-  val lazy_conj_simproc : MetaSimplifier.simproc
-  val string_eq_simp_tac : int -> Tactical.tactic
+  val lookup_simproc : simproc
+  val update_simproc : simproc
+  val ex_lookup_eq_simproc : simproc
+  val ex_lookup_ss : simpset
+  val lazy_conj_simproc : simproc
+  val string_eq_simp_tac : int -> tactic
 
-  val setup : Context.theory -> Context.theory
+  val setup : theory -> theory
 end;
 
 structure StateFun: STATE_FUN = 
--- a/src/HOL/Statespace/state_space.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Statespace/state_space.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Statespace/state_space.ML
-    ID:         $Id$
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
@@ -17,7 +16,7 @@
     val namespace_definition :
        bstring ->
        typ ->
-       Locale.expr ->
+       Expression.expression ->
        string list -> string list -> theory -> theory
 
     val define_statespace :
@@ -61,7 +60,7 @@
     val update_tr' : Proof.context -> term list -> term
   end;
 
-structure StateSpace: STATE_SPACE =
+structure StateSpace : STATE_SPACE =
 struct
 
 (* Theorems *)
@@ -148,11 +147,25 @@
 
 fun prove_interpretation_in ctxt_tac (name, expr) thy =
    thy
-   |> Locale.interpretation_in_locale I (name, expr)
+   |> Expression.sublocale_cmd name expr
    |> Proof.global_terminal_proof
          (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (ctxt_tac ctxt),Position.none), NONE)
    |> ProofContext.theory_of
 
+fun add_locale name expr elems thy =
+  thy 
+  |> Expression.add_locale name name expr elems
+  |> (fn ((target, notes), ctxt) => TheoryTarget.begin target ctxt |>
+           fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes)
+  |> LocalTheory.exit;
+
+fun add_locale_cmd name expr elems thy =
+  thy 
+  |> Expression.add_locale_cmd name name expr elems
+  |> (fn ((target, notes), ctxt) => TheoryTarget.begin target ctxt |>
+           fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes)
+  |> LocalTheory.exit;
+
 type statespace_info =
  {args: (string * sort) list, (* type arguments *)
   parents: (typ list * string * string option list) list,
@@ -252,7 +265,7 @@
       in EVERY [rtac rule i] st
       end
 
-    fun tac ctxt = EVERY [Locale.intro_locales_tac true ctxt [],
+    fun tac ctxt = EVERY [NewLocale.intro_locales_tac true ctxt [],
                           ALLGOALS (SUBGOAL (solve_tac ctxt))]
 
   in thy
@@ -264,16 +277,11 @@
 fun namespace_definition name nameT parent_expr parent_comps new_comps thy =
   let
     val all_comps = parent_comps @ new_comps;
-    val vars = Locale.Merge
-                (map (fn n => Locale.Rename (Locale.Locale (Locale.intern thy "var")
-                                            ,[SOME (n,NONE)])) all_comps);
-
+    val vars = (map (fn n => (Binding.name n, NONE, NoSyn)) all_comps);
     val full_name = Sign.full_bname thy name;
     val dist_thm_name = distinct_compsN;
-    val dist_thm_full_name =
-        let val prefix = fold1' (fn name => fn prfx => prfx ^ "_" ^ name) all_comps "";
-        in if prefix = "" then dist_thm_name else prefix ^ "." ^ dist_thm_name end;
 
+    val dist_thm_full_name = dist_thm_name;
     fun comps_of_thm thm = prop_of thm
              |> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free);
 
@@ -309,12 +317,10 @@
                       DistinctTreeProver.mk_tree (fn n => Free (n,nameT)) nameT
                                 (sort fast_string_ord all_comps)),
                       ([]))])];
-
   in thy
-     |> Locale.add_locale "" name vars [assumes]
-     ||> ProofContext.theory_of
-     ||> interprete_parent name dist_thm_full_name parent_expr
-     |> #2
+     |> add_locale name ([],vars) [assumes]
+     |> ProofContext.theory_of
+     |> interprete_parent name dist_thm_full_name parent_expr
   end;
 
 fun encode_dot x = if x= #"." then #"_" else x;
@@ -378,11 +384,10 @@
 
     val components' = map (fn (n,T) => (n,(T,full_name))) components;
     val all_comps' = map (fn (n,T) => (n,(T,full_name))) all_comps;
-    fun parent_expr (_,n,rs) = Locale.Rename
-                                (Locale.Locale (suffix namespaceN n),
-                                 map (Option.map (fn s => (s,NONE))) rs);
-    val parents_expr = Locale.Merge (fold (fn p => fn es => parent_expr p::es) parents []);
 
+    fun parent_expr (_,n,rs) = (suffix namespaceN n,((n,false),Expression.Positional rs));
+        (* FIXME: a more specific renaming-prefix (including parameter names) may be nicer *)
+    val parents_expr = map parent_expr parents;
     fun distinct_types Ts =
       let val tab = fold (fn T => fn tab => Typtab.update (T,()) tab) Ts Typtab.empty;
       in map fst (Typtab.dest tab) end;
@@ -396,30 +401,35 @@
     val stateT = nameT --> valueT;
     fun projectT T = valueT --> T;
     fun injectT T = T --> valueT;
-
-    val locs = map (fn T => Locale.Rename (Locale.Locale project_injectL,
-                             [SOME (project_name T,NONE),
-                              SOME (inject_name T ,NONE)])) Ts;
+    val locinsts = map (fn T => (project_injectL,
+                    (("",false),Expression.Positional 
+                             [SOME (Free (project_name T,projectT T)), 
+                              SOME (Free ((inject_name T,injectT T)))]))) Ts;
+    val locs = flat (map (fn T => [(Binding.name (project_name T),NONE,NoSyn),
+                                     (Binding.name (inject_name T),NONE,NoSyn)]) Ts);
     val constrains = List.concat
          (map (fn T => [(project_name T,projectT T),(inject_name T,injectT T)]) Ts);
 
-    fun interprete_parent_valuetypes (Ts, pname, _) =
+    fun interprete_parent_valuetypes (Ts, pname, _) thy =
       let
         val {args,types,...} =
              the (Symtab.lookup (StateSpaceData.get (Context.Theory thy)) pname);
         val inst = map fst args ~~ Ts;
         val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);
         val pars = List.concat (map ((fn T => [project_name T,inject_name T]) o subst) types);
-        val expr = Locale.Rename (Locale.Locale (suffix valuetypesN name),
-                                  map (fn n => SOME (n,NONE)) pars);
-      in prove_interpretation_in (K all_tac)
-           (suffix valuetypesN name, expr) end;
+
+        val expr = ([(suffix valuetypesN name, 
+                     (("",false),Expression.Positional (map SOME pars)))],[]);
+      in
+        prove_interpretation_in (ALLGOALS o solve_tac o Assumption.prems_of)
+          (suffix valuetypesN name, expr) thy
+      end;
 
     fun interprete_parent (_, pname, rs) =
       let
-        val expr = Locale.Rename (Locale.Locale pname, map (Option.map (fn n => (n,NONE))) rs)
+        val expr = ([(pname, (("",false), Expression.Positional rs))],[])
       in prove_interpretation_in
-           (fn ctxt => Locale.intro_locales_tac false ctxt [])
+           (fn ctxt => NewLocale.intro_locales_tac false ctxt [])
            (full_name, expr) end;
 
     fun declare_declinfo updates lthy phi ctxt =
@@ -454,17 +464,16 @@
 
   in thy
      |> namespace_definition
-           (suffix namespaceN name) nameT parents_expr
+           (suffix namespaceN name) nameT (parents_expr,[])
            (map fst parent_comps) (map fst components)
      |> Context.theory_map (add_statespace full_name args parents components [])
-     |> Locale.add_locale "" (suffix valuetypesN name) (Locale.Merge locs)
-            [Element.Constrains constrains]
-     |> ProofContext.theory_of o #2
+     |> add_locale (suffix valuetypesN name) (locinsts,locs) []
+     |> ProofContext.theory_of 
      |> fold interprete_parent_valuetypes parents
-     |> Locale.add_locale_cmd name
-              (Locale.Merge [Locale.Locale (suffix namespaceN full_name)
-                            ,Locale.Locale (suffix valuetypesN full_name)]) fixestate
-     |> ProofContext.theory_of o #2
+     |> add_locale_cmd name
+              ([(suffix namespaceN full_name ,(("",false),Expression.Named [])),
+                (suffix valuetypesN full_name,(("",false),Expression.Named  []))],[]) fixestate
+     |> ProofContext.theory_of 
      |> fold interprete_parent parents
      |> add_declaration (SOME full_name) (declare_declinfo components')
   end;
@@ -572,7 +581,6 @@
              | xs => ["Components already defined in parents: " ^ commas xs]);
     val errs = err_dup_args @ err_dup_components @ err_extra_frees @
                err_dup_types @ err_comp_in_parent;
-
   in if null errs
      then thy |> statespace_definition state_space args' name parents' parent_comps comps'
      else error (cat_lines errs)
--- a/src/HOL/Tools/Groebner_Basis/groebner.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/Groebner_Basis/groebner.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Groebner_Basis/groebner.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -599,14 +598,14 @@
      if length ideal = 2 then ideal else [eq_commute, eq_commute]
   val ring_dest_neg =
     fn t => let val (l,r) = dest_comb t in
-        if could_unify(term_of l,term_of ring_neg_tm) then r else raise CTERM ("ring_dest_neg", [t])
+        if Term.could_unify(term_of l,term_of ring_neg_tm) then r else raise CTERM ("ring_dest_neg", [t])
       end
 
  val ring_mk_neg = fn tm => mk_comb (ring_neg_tm) (tm);
 (*
  fun ring_dest_inv t =
     let val (l,r) = dest_comb t in
-        if could_unify(term_of l, term_of ring_inv_tm) then r else raise CTERM "ring_dest_inv"
+        if Term.could_unify(term_of l, term_of ring_inv_tm) then r else raise CTERM "ring_dest_inv"
     end
 *)
  val ring_dest_add = dest_binary ring_add_tm;
@@ -687,7 +686,7 @@
   val cjs = conjs tm
   val  rawvars = fold_rev (fn eq => fn a =>
                                        grobvars (dest_arg1 eq) (grobvars (dest_arg eq) a)) cjs []
-  val vars = sort (fn (x, y) => Term.term_ord(term_of x,term_of y))
+  val vars = sort (fn (x, y) => TermOrd.term_ord(term_of x,term_of y))
                   (distinct (op aconvc) rawvars)
  in (vars,map (grobify_equation vars) cjs)
  end;
@@ -896,7 +895,7 @@
   val vars = filter (fn v => free_in v eq) evs
   val (qs,p) = isolate_monomials vars eq
   val rs = ideal (qs @ ps) p 
-              (fn (s,t) => Term.term_ord (term_of s, term_of t))
+              (fn (s,t) => TermOrd.term_ord (term_of s, term_of t))
  in (eq,Library.take (length qs, rs) ~~ vars)
  end;
  fun subst_in_poly i p = Thm.rhs_of (ring_normalize_conv (vsubst i p));
--- a/src/HOL/Tools/Groebner_Basis/normalizer.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/Groebner_Basis/normalizer.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Groebner_Basis/normalizer.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -615,7 +614,7 @@
 val nat_exp_ss = HOL_basic_ss addsimps (@{thms nat_number} @ nat_arith @ @{thms arith_simps} @ @{thms rel_simps})
                               addsimps [Let_def, if_False, if_True, @{thm add_0}, @{thm add_Suc}];
 
-fun simple_cterm_ord t u = Term.term_ord (term_of t, term_of u) = LESS;
+fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS;
 
 fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, idom, ideal}, 
                                      {conv, dest_const, mk_const, is_const}) ord =
--- a/src/HOL/Tools/Qelim/cooper.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/Qelim/cooper.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Qelim/cooper.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -515,7 +514,7 @@
   let val _ = ()
   in
    Qelim.gen_qelim_conv pcv postcv pcv (cons o term_of) 
-      (term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt) 
+      (OldTerm.term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt) 
       (cooperex_conv ctxt) p 
   end
   handle  CTERM s => raise COOPER ("Cooper Failed", CTERM s)
@@ -637,7 +636,7 @@
   let
     val thy = Thm.theory_of_cterm ct;
     val t = Thm.term_of ct;
-    val (vs, ps) = pairself (map_index swap) (term_frees t, term_bools [] t);
+    val (vs, ps) = pairself (map_index swap) (OldTerm.term_frees t, term_bools [] t);
   in
     Thm.cterm_of thy (Logic.mk_equals (HOLogic.mk_Trueprop t,
       HOLogic.mk_Trueprop (term_of_qf ps vs (pa (qf_of_term ps vs t)))))
--- a/src/HOL/Tools/Qelim/langford.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/Qelim/langford.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,3 +1,7 @@
+(*  Title:      HOL/Tools/Qelim/langford.ML
+    Author:     Amine Chaieb, TU Muenchen
+*)
+
 signature LANGFORD_QE = 
 sig
   val dlo_tac : Proof.context -> int -> tactic
@@ -114,7 +118,7 @@
       let val (yes,no) = partition f xs 
       in if f x then (x::yes,no) else (yes, x::no) end;
 
-fun contains x ct = member (op aconv) (term_frees (term_of ct)) (term_of x);
+fun contains x ct = member (op aconv) (OldTerm.term_frees (term_of ct)) (term_of x);
 
 fun is_eqx x eq = case term_of eq of
    Const("op =",_)$l$r => l aconv term_of x orelse r aconv term_of x
@@ -211,7 +215,7 @@
  let 
    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
    fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
-   val ts = sort (fn (a,b) => Term.fast_term_ord (term_of a, term_of b)) (f p)
+   val ts = sort (fn (a,b) => TermOrd.fast_term_ord (term_of a, term_of b)) (f p)
    val p' = fold_rev gen ts p
  in implies_intr p' (implies_elim st (fold forall_elim ts (assume p'))) end));
 
--- a/src/HOL/Tools/Qelim/presburger.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/Qelim/presburger.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Qelim/presburger.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -86,8 +85,8 @@
 in 
 fun is_relevant ctxt ct = 
  gen_subset (op aconv) (term_constants (term_of ct) , snd (CooperData.get ctxt))
- andalso forall (fn Free (_,T) => T mem [@{typ "int"}, @{typ nat}]) (term_frees (term_of ct))
- andalso forall (fn Var (_,T) => T mem [@{typ "int"}, @{typ nat}]) (term_vars (term_of ct));
+ andalso forall (fn Free (_,T) => T mem [@{typ "int"}, @{typ nat}]) (OldTerm.term_frees (term_of ct))
+ andalso forall (fn Var (_,T) => T mem [@{typ "int"}, @{typ nat}]) (OldTerm.term_vars (term_of ct));
 
 fun int_nat_terms ctxt ct =
  let 
@@ -104,7 +103,7 @@
  let 
    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
    fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
-   val ts = sort (fn (a,b) => Term.fast_term_ord (term_of a, term_of b)) (f p)
+   val ts = sort (fn (a,b) => TermOrd.fast_term_ord (term_of a, term_of b)) (f p)
    val p' = fold_rev gen ts p
  in implies_intr p' (implies_elim st (fold forall_elim ts (assume p'))) end));
 
--- a/src/HOL/Tools/TFL/casesplit.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/TFL/casesplit.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/TFL/casesplit.ML
-    ID:         $Id$
     Author:     Lucas Dixon, University of Edinburgh
 
 A structure that defines a tactic to program case splits.
@@ -104,7 +103,7 @@
     end;
 
 (*
- val ty = (snd o hd o map Term.dest_Free o Term.term_frees) t;
+ val ty = (snd o hd o map Term.dest_Free o OldTerm.term_frees) t;
 *)
 
 
@@ -122,9 +121,9 @@
       val abs_ct = ctermify abst;
       val free_ct = ctermify x;
 
-      val casethm_vars = rev (Term.term_vars (Thm.concl_of case_thm));
+      val casethm_vars = rev (OldTerm.term_vars (Thm.concl_of case_thm));
 
-      val casethm_tvars = Term.term_tvars (Thm.concl_of case_thm);
+      val casethm_tvars = OldTerm.term_tvars (Thm.concl_of case_thm);
       val (Pv, Dv, type_insts) =
           case (Thm.concl_of case_thm) of
             (_ $ ((Pv as Var(P,Pty)) $ (Dv as Var(D, Dty)))) =>
@@ -170,7 +169,7 @@
       val subgoalth' = atomize_goals subgoalth;
       val gt = Data.dest_Trueprop (Logic.get_goal (Thm.prop_of subgoalth') 1);
 
-      val freets = Term.term_frees gt;
+      val freets = OldTerm.term_frees gt;
       fun getter x =
           let val (n,ty) = Term.dest_Free x in
             (if vstr = n orelse vstr = Name.dest_skolem n
--- a/src/HOL/Tools/TFL/rules.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/TFL/rules.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,7 @@
 (*  Title:      HOL/Tools/TFL/rules.ML
-    ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
-    Copyright   1997  University of Cambridge
 
-Emulation of HOL inference rules for TFL
+Emulation of HOL inference rules for TFL.
 *)
 
 signature RULES =
@@ -173,7 +171,7 @@
  *---------------------------------------------------------------------------*)
 local val thy = Thm.theory_of_thm disjI1
       val prop = Thm.prop_of disjI1
-      val [P,Q] = term_vars prop
+      val [P,Q] = OldTerm.term_vars prop
       val disj1 = Thm.forall_intr (Thm.cterm_of thy Q) disjI1
 in
 fun DISJ1 thm tm = thm RS (forall_elim (D.drop_prop tm) disj1)
@@ -182,7 +180,7 @@
 
 local val thy = Thm.theory_of_thm disjI2
       val prop = Thm.prop_of disjI2
-      val [P,Q] = term_vars prop
+      val [P,Q] = OldTerm.term_vars prop
       val disj2 = Thm.forall_intr (Thm.cterm_of thy P) disjI2
 in
 fun DISJ2 tm thm = thm RS (forall_elim (D.drop_prop tm) disj2)
@@ -278,7 +276,7 @@
 local (* this is fragile *)
       val thy = Thm.theory_of_thm spec
       val prop = Thm.prop_of spec
-      val x = hd (tl (term_vars prop))
+      val x = hd (tl (OldTerm.term_vars prop))
       val cTV = ctyp_of thy (type_of x)
       val gspec = forall_intr (cterm_of thy x) spec
 in
@@ -295,7 +293,7 @@
 (* Not optimized! Too complicated. *)
 local val thy = Thm.theory_of_thm allI
       val prop = Thm.prop_of allI
-      val [P] = add_term_vars (prop, [])
+      val [P] = OldTerm.add_term_vars (prop, [])
       fun cty_theta s = map (fn (i, (S, ty)) => (ctyp_of s (TVar (i, S)), ctyp_of s ty))
       fun ctm_theta s = map (fn (i, (_, tm2)) =>
                              let val ctm2 = cterm_of s tm2
@@ -325,7 +323,7 @@
    let val thy = Thm.theory_of_thm thm
        val prop = Thm.prop_of thm
        val tycheck = cterm_of thy
-       val vlist = map tycheck (add_term_vars (prop, []))
+       val vlist = map tycheck (OldTerm.add_term_vars (prop, []))
   in GENL vlist thm
   end;
 
@@ -365,7 +363,7 @@
 
 local val thy = Thm.theory_of_thm exI
       val prop = Thm.prop_of exI
-      val [P,x] = term_vars prop
+      val [P,x] = OldTerm.term_vars prop
 in
 fun EXISTS (template,witness) thm =
    let val thy = Thm.theory_of_thm thm
@@ -516,7 +514,7 @@
                 val (f,args) = S.strip_comb (get_lhs eq)
                 val (vstrl,_) = S.strip_abs f
                 val names  =
-                  Name.variant_list (add_term_names(body, [])) (map (#1 o dest_Free) vstrl)
+                  Name.variant_list (OldTerm.add_term_names(body, [])) (map (#1 o dest_Free) vstrl)
             in get (rst, n+1, (names,n)::L) end
             handle TERM _ => get (rst, n+1, L)
               | U.ERR _ => get (rst, n+1, L);
@@ -676,7 +674,7 @@
      fun prover used ss thm =
      let fun cong_prover ss thm =
          let val dummy = say "cong_prover:"
-             val cntxt = MetaSimplifier.prems_of_ss ss
+             val cntxt = Simplifier.prems_of_ss ss
              val dummy = print_thms "cntxt:" cntxt
              val dummy = say "cong rule:"
              val dummy = say (Display.string_of_thm thm)
@@ -689,7 +687,7 @@
                      val ants = map tych (Logic.strip_imp_prems imp)
                      val eq = Logic.strip_imp_concl imp
                      val lhs = tych(get_lhs eq)
-                     val ss' = MetaSimplifier.add_prems (map ASSUME ants) ss
+                     val ss' = Simplifier.add_prems (map ASSUME ants) ss
                      val lhs_eq_lhs1 = MetaSimplifier.rewrite_cterm (false,true,false) (prover used) ss' lhs
                        handle U.ERR _ => Thm.reflexive lhs
                      val dummy = print_thms "proven:" [lhs_eq_lhs1]
@@ -710,7 +708,7 @@
                   val Q = get_lhs eq1
                   val QeqQ1 = pbeta_reduce (tych Q)
                   val Q1 = #2(D.dest_eq(cconcl QeqQ1))
-                  val ss' = MetaSimplifier.add_prems (map ASSUME ants1) ss
+                  val ss' = Simplifier.add_prems (map ASSUME ants1) ss
                   val Q1eeqQ2 = MetaSimplifier.rewrite_cterm (false,true,false) (prover used') ss' Q1
                                 handle U.ERR _ => Thm.reflexive Q1
                   val Q2 = #2 (Logic.dest_equals (Thm.prop_of Q1eeqQ2))
@@ -760,12 +758,12 @@
 
         fun restrict_prover ss thm =
           let val dummy = say "restrict_prover:"
-              val cntxt = rev(MetaSimplifier.prems_of_ss ss)
+              val cntxt = rev(Simplifier.prems_of_ss ss)
               val dummy = print_thms "cntxt:" cntxt
               val thy = Thm.theory_of_thm thm
               val Const("==>",_) $ (Const("Trueprop",_) $ A) $ _ = Thm.prop_of thm
               fun genl tm = let val vlist = subtract (op aconv) globals
-                                           (add_term_frees(tm,[]))
+                                           (OldTerm.add_term_frees(tm,[]))
                             in fold_rev Forall vlist tm
                             end
               (*--------------------------------------------------------------
@@ -805,7 +803,7 @@
     (if (is_cong thm) then cong_prover else restrict_prover) ss thm
     end
     val ctm = cprop_of th
-    val names = add_term_names (term_of ctm, [])
+    val names = OldTerm.add_term_names (term_of ctm, [])
     val th1 = MetaSimplifier.rewrite_cterm(false,true,false)
       (prover names) (ss0 addsimps [cut_lemma'] addeqcongs congs) ctm
     val th2 = equal_elim th1 th
--- a/src/HOL/Tools/TFL/tfl.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/TFL/tfl.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      HOL/Tools/TFL/tfl.ML
-    ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
-    Copyright   1997  University of Cambridge
 
 First part of main module.
 *)
@@ -332,7 +330,7 @@
      val dummy = map (no_repeat_vars thy) pats
      val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats,
                               map (fn (t,i) => (t,(i,true))) (enumerate R))
-     val names = foldr add_term_names [] R
+     val names = foldr OldTerm.add_term_names [] R
      val atype = type_of(hd pats)
      and aname = Name.variant names "a"
      val a = Free(aname,atype)
@@ -494,7 +492,7 @@
      val tych = Thry.typecheck thy
      val WFREC_THM0 = R.ISPEC (tych functional) Thms.WFREC_COROLLARY
      val Const("All",_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
-     val R = Free (Name.variant (foldr add_term_names [] eqns) Rname,
+     val R = Free (Name.variant (foldr OldTerm.add_term_names [] eqns) Rname,
                    Rtype)
      val WFREC_THM = R.ISPECL [tych R, tych g] WFREC_THM0
      val ([proto_def, WFR],_) = S.strip_imp(concl WFREC_THM)
@@ -692,7 +690,7 @@
  let val tych = Thry.typecheck thy
      val ty_info = Thry.induct_info thy
  in fn pats =>
- let val names = foldr add_term_names [] pats
+ let val names = foldr OldTerm.add_term_names [] pats
      val T = type_of (hd pats)
      val aname = Name.variant names "a"
      val vname = Name.variant (aname::names) "v"
@@ -845,7 +843,7 @@
     val (pats,TCsl) = ListPair.unzip pat_TCs_list
     val case_thm = complete_cases thy pats
     val domain = (type_of o hd) pats
-    val Pname = Name.variant (foldr (Library.foldr add_term_names)
+    val Pname = Name.variant (foldr (Library.foldr OldTerm.add_term_names)
                               [] (pats::TCsl)) "P"
     val P = Free(Pname, domain --> HOLogic.boolT)
     val Sinduct = R.SPEC (tych P) Sinduction
@@ -856,7 +854,7 @@
     val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats
     val tasks = U.zip3 cases TCl' (R.CONJUNCTS Rinduct_assum)
     val proved_cases = map (prove_case fconst thy) tasks
-    val v = Free (Name.variant (foldr add_term_names [] (map concl proved_cases))
+    val v = Free (Name.variant (foldr OldTerm.add_term_names [] (map concl proved_cases))
                     "v",
                   domain)
     val vtyped = tych v
--- a/src/HOL/Tools/TFL/usyntax.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/TFL/usyntax.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      HOL/Tools/TFL/usyntax.ML
-    ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
-    Copyright   1997  University of Cambridge
 
 Emulation of HOL's abstract syntax functions.
 *)
@@ -114,7 +112,7 @@
 
 val is_vartype = can dest_vtype;
 
-val type_vars  = map mk_prim_vartype o typ_tvars
+val type_vars  = map mk_prim_vartype o OldTerm.typ_tvars
 fun type_varsl L = distinct (op =) (fold (curry op @ o type_vars) L []);
 
 val alpha  = mk_vartype "'a"
@@ -144,7 +142,7 @@
 
 
 
-val type_vars_in_term = map mk_prim_vartype o term_tvars;
+val type_vars_in_term = map mk_prim_vartype o OldTerm.term_tvars;
 
 
 
@@ -321,7 +319,7 @@
 
 
 (* Need to reverse? *)
-fun gen_all tm = list_mk_forall(term_frees tm, tm);
+fun gen_all tm = list_mk_forall(OldTerm.term_frees tm, tm);
 
 (* Destructing a cterm to a list of Terms *)
 fun strip_comb tm =
--- a/src/HOL/Tools/arith_data.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/arith_data.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/arith_data.ML
-    ID:         $Id$
     Author:     Markus Wenzel, Stefan Berghofer, and Tobias Nipkow
 
 Basic arithmetic proof tools.
@@ -7,9 +6,8 @@
 
 signature ARITH_DATA =
 sig
-  val prove_conv: tactic -> (MetaSimplifier.simpset -> tactic)
-    -> MetaSimplifier.simpset -> term * term -> thm
-  val simp_all_tac: thm list -> MetaSimplifier.simpset -> tactic
+  val prove_conv: tactic -> (simpset -> tactic) -> simpset -> term * term -> thm
+  val simp_all_tac: thm list -> simpset -> tactic
 
   val mk_sum: term list -> term
   val mk_norm_sum: term list -> term
--- a/src/HOL/Tools/cnf_funcs.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/cnf_funcs.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,8 +1,6 @@
 (*  Title:      HOL/Tools/cnf_funcs.ML
-    ID:         $Id$
     Author:     Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr)
-    Author:     Tjark Weber
-    Copyright   2005-2006
+    Author:     Tjark Weber, TU Muenchen
 
   FIXME: major overlaps with the code in meson.ML
 
@@ -499,7 +497,7 @@
 					NONE
 		in
 			Int.max (max, getOpt (idx, 0))
-		end) (term_frees simp) 0)
+		end) (OldTerm.term_frees simp) 0)
 	(* finally, convert to definitional CNF (this should preserve the simplification) *)
 	val cnfx_thm = make_cnfx_thm_from_nnf simp
 in
--- a/src/HOL/Tools/datatype_abs_proofs.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/datatype_abs_proofs.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_abs_proofs.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Proofs and defintions independent of concrete representation
@@ -10,8 +9,7 @@
  - characteristic equations for primrec combinators
  - characteristic equations for case combinators
  - equations for splitting "P (case ...)" expressions
-  - "nchotomy" and "case_cong" theorems for TFL
-
+ - "nchotomy" and "case_cong" theorems for TFL
 *)
 
 signature DATATYPE_ABS_PROOFS =
@@ -98,7 +96,7 @@
 
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
     val newTs = Library.take (length (hd descr), recTs);
 
     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
@@ -281,7 +279,7 @@
 
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
     val newTs = Library.take (length (hd descr), recTs);
     val T' = TFree (Name.variant used "'t", HOLogic.typeS);
 
@@ -425,8 +423,8 @@
         val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
         val cert = cterm_of thy;
         val nchotomy' = nchotomy RS spec;
-        val nchotomy'' = cterm_instantiate
-          [(cert (hd (add_term_vars (concl_of nchotomy', []))), cert Ma)] nchotomy'
+        val [v] = Term.add_vars (concl_of nchotomy') [];
+        val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
       in
         SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
           (fn {prems, ...} => 
--- a/src/HOL/Tools/datatype_aux.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/datatype_aux.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_aux.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Auxiliary functions for defining datatypes.
@@ -131,7 +130,7 @@
     val flt = if null indnames then I else
       filter (fn Free (s, _) => s mem indnames | _ => false);
     fun abstr (t1, t2) = (case t1 of
-        NONE => (case flt (term_frees t2) of
+        NONE => (case flt (OldTerm.term_frees t2) of
             [Free (s, T)] => SOME (absfree (s, T, t2))
           | _ => NONE)
       | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))))
--- a/src/HOL/Tools/datatype_case.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/datatype_case.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,6 @@
 (*  Title:      HOL/Tools/datatype_case.ML
-    ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
-                Stefan Berghofer, TU Muenchen
+    Author:     Stefan Berghofer, TU Muenchen
 
 Nested case expressions on datatypes.
 *)
@@ -55,15 +54,12 @@
  * b=false --> i = ~1
  *---------------------------------------------------------------------------*)
 
-fun pattern_map f (tm,x) = (f tm, x);
-
-fun pattern_subst theta = pattern_map (subst_free theta);
+fun pattern_subst theta (tm, x) = (subst_free theta tm, x);
 
 fun row_of_pat x = fst (snd x);
 
-fun add_row_used ((prfx, pats), (tm, tag)) used =
-  foldl add_term_free_names (foldl add_term_free_names
-    (add_term_free_names (tm, used)) pats) prfx;
+fun add_row_used ((prfx, pats), (tm, tag)) =
+  fold Term.add_free_names (tm :: pats @ prfx);
 
 (* try to preserve names given by user *)
 fun default_names names ts =
@@ -140,7 +136,7 @@
                     let
                       val Ts = map type_of rstp;
                       val xs = Name.variant_list
-                        (foldl add_term_free_names used' gvars)
+                        (fold Term.add_free_names gvars used')
                         (replicate (length rstp) "x")
                     in
                       [((prfx, gvars @ map Free (xs ~~ Ts)),
@@ -222,7 +218,7 @@
               | SOME {case_name, constructors} =>
                 let
                   val pty = body_type cT;
-                  val used' = foldl add_term_free_names used rstp;
+                  val used' = fold Term.add_free_names rstp used;
                   val nrows = maps (expand constructors used' pty) rows;
                   val subproblems = partition ty_match ty_inst type_of used'
                     constructors pty range_ty nrows;
@@ -336,7 +332,7 @@
         | prep_pat t used = case_error ("Bad pattern: " ^ Syntax.string_of_term ctxt t);
       fun dest_case1 (t as Const ("_case1", _) $ l $ r) =
             let val (l', cnstrts) = strip_constraints l
-            in ((fst (prep_pat l' (add_term_free_names (t, []))), r), cnstrts)
+            in ((fst (prep_pat l' (Term.add_free_names t [])), r), cnstrts)
             end
         | dest_case1 t = case_error "dest_case1";
       fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
@@ -366,7 +362,7 @@
             val _ = if length zs < i then raise CASE_ERROR ("", 0) else ();
             val (xs, ys) = chop i zs;
             val u = list_abs (ys, strip_abs_body t);
-            val xs' = map Free (Name.variant_list (add_term_names (u, used))
+            val xs' = map Free (Name.variant_list (OldTerm.add_term_names (u, used))
               (map fst xs) ~~ map snd xs)
           in (xs', subst_bounds (rev xs', u)) end;
         fun is_dependent i t =
@@ -424,11 +420,11 @@
 (* destruct nested patterns *)
 
 fun strip_case' dest (pat, rhs) =
-  case dest (add_term_free_names (pat, [])) rhs of
+  case dest (Term.add_free_names pat []) rhs of
     SOME (exp as Free _, clauses) =>
-      if member op aconv (term_frees pat) exp andalso
+      if member op aconv (OldTerm.term_frees pat) exp andalso
         not (exists (fn (_, rhs') =>
-          member op aconv (term_frees rhs') exp) clauses)
+          member op aconv (OldTerm.term_frees rhs') exp) clauses)
       then
         maps (strip_case' dest) (map (fn (pat', rhs') =>
           (subst_free [(exp, pat')] pat, rhs')) clauses)
@@ -451,7 +447,7 @@
     val thy = ProofContext.theory_of ctxt;
     val consts = ProofContext.consts_of ctxt;
     fun mk_clause (pat, rhs) =
-      let val xs = term_frees pat
+      let val xs = Term.add_frees pat []
       in
         Syntax.const "_case1" $
           map_aterms
@@ -459,8 +455,8 @@
               | Const (s, _) => Const (Consts.extern_early consts s, dummyT)
               | t => t) pat $
           map_aterms
-            (fn x as Free (s, _) =>
-                  if member op aconv xs x then Syntax.mark_bound s else x
+            (fn x as Free (s, T) =>
+                  if member (op =) xs (s, T) then Syntax.mark_bound s else x
               | t => t) rhs
       end
   in case strip_case' (tab_of thy) true (list_comb (Syntax.const cname, ts)) of
--- a/src/HOL/Tools/datatype_codegen.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/datatype_codegen.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*  Title:      HOL/Tools/datatype_codegen.ML
-    ID:         $Id$
-    Author:     Stefan Berghofer & Florian Haftmann, TU Muenchen
+    Author:     Stefan Berghofer and Florian Haftmann, TU Muenchen
 
 Code generator facilities for inductive datatypes.
 *)
@@ -217,8 +216,8 @@
       let
         val ts1 = Library.take (i, ts);
         val t :: ts2 = Library.drop (i, ts);
-        val names = foldr add_term_names
-          (map (fst o fst o dest_Var) (foldr add_term_vars [] ts1)) ts1;
+        val names = foldr OldTerm.add_term_names
+          (map (fst o fst o dest_Var) (foldr OldTerm.add_term_vars [] ts1)) ts1;
         val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
 
         fun pcase [] [] [] gr = ([], gr)
@@ -414,7 +413,7 @@
     val inject2 = (#inject o DatatypePackage.the_datatype thy) dtco;
     val ctxt = ProofContext.init thy;
     val simpset = Simplifier.context ctxt
-      (MetaSimplifier.empty_ss addsimprocs [DatatypePackage.distinct_simproc]);
+      (Simplifier.empty_ss addsimprocs [DatatypePackage.distinct_simproc]);
     val cos = map (fn (co, tys) =>
         (Const (co, tys ---> Type (dtco, map TFree vs)), tys)) cs;
     val tac = ALLGOALS (simp_tac simpset)
--- a/src/HOL/Tools/datatype_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/datatype_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_package.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Datatype package for Isabelle/HOL.
@@ -492,7 +491,7 @@
       ^ Syntax.string_of_typ_global thy T);
     fun type_of_constr (cT as (_, T)) =
       let
-        val frees = typ_tfrees T;
+        val frees = OldTerm.typ_tfrees T;
         val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
           handle TYPE _ => no_constr cT
         val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
@@ -583,7 +582,7 @@
         fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
           let
             val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
-            val _ = (case fold (curry add_typ_tfree_names) cargs' [] \\ tvs of
+            val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
                 [] => ()
               | vs => error ("Extra type variables on rhs: " ^ commas vs))
           in (constrs @ [((if flat_names then Sign.full_bname tmp_thy else
--- a/src/HOL/Tools/datatype_prop.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/datatype_prop.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_prop.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Characteristic properties of datatypes.
@@ -206,7 +205,7 @@
   let
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
 
     val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr sorts used;
 
@@ -256,7 +255,7 @@
   let
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
     val newTs = Library.take (length (hd descr), recTs);
     val T' = TFree (Name.variant used "'t", HOLogic.typeS);
 
@@ -303,7 +302,7 @@
   let
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used' = foldr add_typ_tfree_names [] recTs;
+    val used' = foldr OldTerm.add_typ_tfree_names [] recTs;
     val newTs = Library.take (length (hd descr), recTs);
     val T' = TFree (Name.variant used' "'t", HOLogic.typeS);
     val P = Free ("P", T' --> HOLogic.boolT);
--- a/src/HOL/Tools/datatype_rep_proofs.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/datatype_rep_proofs.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_rep_proofs.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Definitional introduction of datatypes
@@ -8,7 +7,6 @@
  - injectivity of constructors
  - distinctness of constructors
  - induction theorem
-
 *)
 
 signature DATATYPE_REP_PROOFS =
@@ -85,7 +83,7 @@
     val branchT = if null branchTs then HOLogic.unitT
       else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) branchTs;
     val arities = get_arities descr' \ 0;
-    val unneeded_vars = hd tyvars \\ foldr add_typ_tfree_names [] (leafTs' @ branchTs);
+    val unneeded_vars = hd tyvars \\ foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs);
     val leafTs = leafTs' @ (map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars);
     val recTs = get_rec_types descr' sorts;
     val newTs = Library.take (length (hd descr), recTs);
@@ -369,7 +367,7 @@
         val prop = Thm.prop_of thm;
         val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
           (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
-        val used = add_term_tfree_names (a, []);
+        val used = OldTerm.add_term_tfree_names (a, []);
 
         fun mk_thm i =
           let
--- a/src/HOL/Tools/function_package/context_tree.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/function_package/context_tree.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,12 +1,10 @@
 (*  Title:      HOL/Tools/function_package/context_tree.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
 A package for general recursive function definitions. 
 Builds and traverses trees of nested contexts along a term.
 *)
 
-
 signature FUNDEF_CTXTREE =
 sig
     type ctxt = (string * typ) list * thm list (* poor man's contexts: fixes + assumes *)
@@ -82,7 +80,7 @@
     let 
       val t' = snd (dest_all_all t)
       val (assumes, concl) = Logic.strip_horn t'
-    in (fold (curry add_term_vars) assumes [], term_vars concl)
+    in (fold Term.add_vars assumes [], Term.add_vars concl [])
     end
 
 fun cong_deps crule =
@@ -91,7 +89,9 @@
     in
       IntGraph.empty
         |> fold (fn (i,_)=> IntGraph.new_node (i,i)) num_branches
-        |> fold_product (fn (i,(c1,_)) => fn (j,(_, t2)) => if i = j orelse null (c1 inter t2) then I else IntGraph.add_edge_acyclic (i,j))
+        |> fold_product (fn (i, (c1, _)) => fn (j, (_, t2)) => 
+               if i = j orelse null (c1 inter t2) 
+               then I else IntGraph.add_edge_acyclic (i,j))
              num_branches num_branches
     end
     
--- a/src/HOL/Tools/function_package/decompose.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/function_package/decompose.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -19,7 +19,7 @@
 structure Decompose : DECOMPOSE =
 struct
 
-structure TermGraph = GraphFun(type key = term val ord = Term.fast_term_ord);
+structure TermGraph = GraphFun(type key = term val ord = TermOrd.fast_term_ord);
 
 
 fun derive_chains ctxt chain_tac cont D = Termination.CALLS (fn (cs, i) =>
--- a/src/HOL/Tools/function_package/fundef_core.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/function_package/fundef_core.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,8 @@
 (*  Title:      HOL/Tools/function_package/fundef_core.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
-A package for general recursive function definitions.
-Main functionality
+A package for general recursive function definitions:
+Main functionality.
 *)
 
 signature FUNDEF_CORE =
@@ -437,7 +436,7 @@
                                   |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
                                   |> forall_intr (cterm_of thy x)
                                   |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
-                                  |> (fn it => fold (forall_intr o cterm_of thy) (term_vars (prop_of it)) it)
+                                  |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
 
         val goalstate =  Conjunction.intr graph_is_function complete
                           |> Thm.close_derivation
--- a/src/HOL/Tools/function_package/fundef_datatype.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/function_package/fundef_datatype.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/function_package/fundef_datatype.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
 A package for general recursive function definitions.
@@ -64,9 +63,10 @@
 
 fun inst_case_thm thy x P thm =
     let
-        val [Pv, xv] = term_vars (prop_of thm)
+        val [Pv, xv] = Term.add_vars (prop_of thm) []
     in
-        cterm_instantiate [(cterm_of thy xv, cterm_of thy x), (cterm_of thy Pv, cterm_of thy P)] thm
+        cterm_instantiate [(cterm_of thy (Var xv), cterm_of thy x), 
+                           (cterm_of thy (Var Pv), cterm_of thy P)] thm
     end
 
 
--- a/src/HOL/Tools/function_package/fundef_lib.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/function_package/fundef_lib.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/function_package/fundef_lib.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
 A package for general recursive function definitions. 
@@ -163,7 +162,7 @@
           else if js = []
             then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
             else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
-     (K (MetaSimplifier.rewrite_goals_tac ac
+     (K (rewrite_goals_tac ac
          THEN rtac Drule.reflexive_thm 1))
  end
 
--- a/src/HOL/Tools/function_package/termination.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/function_package/termination.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,4 +1,4 @@
-(*  Title:       HOL/Tools/function_package/termination_data.ML
+(*  Title:       HOL/Tools/function_package/termination.ML
     Author:      Alexander Krauss, TU Muenchen
 
 Context data for termination proofs
@@ -50,9 +50,9 @@
 
 open FundefLib
 
-val term2_ord = prod_ord Term.fast_term_ord Term.fast_term_ord
+val term2_ord = prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord
 structure Term2tab = TableFun(type key = term * term val ord = term2_ord);
-structure Term3tab = TableFun(type key = term * (term * term) val ord = prod_ord Term.fast_term_ord term2_ord);
+structure Term3tab = TableFun(type key = term * (term * term) val ord = prod_ord TermOrd.fast_term_ord term2_ord);
 
 (** Analyzing binary trees **)
 
--- a/src/HOL/Tools/inductive_codegen.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/inductive_codegen.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/inductive_codegen.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Code generator for inductive predicates.
@@ -58,7 +57,7 @@
       | _ => (warn thm; thy))
     | SOME (Const (s, _), _) =>
         let
-          val cs = foldr add_term_consts [] (prems_of thm);
+          val cs = fold Term.add_const_names (Thm.prems_of thm) [];
           val rules = Symtab.lookup_list intros s;
           val nparms = (case optnparms of
             SOME k => k
@@ -129,7 +128,7 @@
   cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
     string_of_mode ms)) modes));
 
-val term_vs = map (fst o fst o dest_Var) o term_vars;
+val term_vs = map (fst o fst o dest_Var) o OldTerm.term_vars;
 val terms_vs = distinct (op =) o List.concat o (map term_vs);
 
 (** collect all Vars in a term (with duplicates!) **)
@@ -491,7 +490,7 @@
       | SOME (names, thyname, nparms, intrs) =>
           mk_ind_def thy defs gr dep names (if_library thyname module)
             [] (prep_intrs intrs) nparms))
-            (gr, foldr add_term_consts [] ts)
+            (gr, fold Term.add_const_names ts [])
 
 and mk_ind_def thy defs gr dep names module modecs intrs nparms =
   add_edge_acyclic (hd names, dep) gr handle
--- a/src/HOL/Tools/inductive_realizer.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/inductive_realizer.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,8 @@
 (*  Title:      HOL/Tools/inductive_realizer.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Porgram extraction from proofs involving inductive predicates:
-Realizers for induction and elimination rules
+Realizers for induction and elimination rules.
 *)
 
 signature INDUCTIVE_REALIZER =
@@ -50,7 +49,7 @@
       t $ strip_all' used names Q
   | strip_all' _ _ t = t;
 
-fun strip_all t = strip_all' (add_term_free_names (t, [])) [] t;
+fun strip_all t = strip_all' (Term.add_free_names t []) [] t;
 
 fun strip_one name (Const ("all", _) $ Abs (s, T, Const ("==>", _) $ P $ Q)) =
       (subst_bound (Free (name, T), P), subst_bound (Free (name, T), Q))
@@ -60,11 +59,11 @@
       (Var ((a, i), T), vs) => (case strip_type T of
         (_, Type (s, _)) => if s mem ["bool"] then (a, T) :: vs else vs
       | _ => vs)
-    | (_, vs) => vs) [] (term_vars prop);
+    | (_, vs) => vs) [] (OldTerm.term_vars prop);
 
 fun dt_of_intrs thy vs nparms intrs =
   let
-    val iTs = term_tvars (prop_of (hd intrs));
+    val iTs = OldTerm.term_tvars (prop_of (hd intrs));
     val Tvs = map TVar iTs;
     val (Const (s, _), ts) = strip_comb (HOLogic.dest_Trueprop
       (Logic.strip_imp_concl (prop_of (hd intrs))));
@@ -101,7 +100,7 @@
 fun mk_realizes_eqn n vs nparms intrs =
   let
     val concl = HOLogic.dest_Trueprop (concl_of (hd intrs));
-    val iTs = term_tvars concl;
+    val iTs = OldTerm.term_tvars concl;
     val Tvs = map TVar iTs;
     val (h as Const (s, T), us) = strip_comb concl;
     val params = List.take (us, nparms);
@@ -147,7 +146,7 @@
     val conclT = Extraction.etype_of thy vs [] (Logic.strip_imp_concl rule');
     val used = map (fst o dest_Free) args;
 
-    fun is_rec t = not (null (term_consts t inter rsets));
+    val is_rec = exists_Const (fn (c, _) => member (op =) rsets c);
 
     fun is_meta (Const ("all", _) $ Abs (s, _, P)) = is_meta P
       | is_meta (Const ("==>", _) $ _ $ Q) = is_meta Q
@@ -276,7 +275,7 @@
   let
     val qualifier = NameSpace.qualifier (name_of_thm induct);
     val inducts = PureThy.get_thms thy (NameSpace.qualified qualifier "inducts");
-    val iTs = term_tvars (prop_of (hd intrs));
+    val iTs = OldTerm.term_tvars (prop_of (hd intrs));
     val ar = length vs + length iTs;
     val params = InductivePackage.params_of raw_induct;
     val arities = InductivePackage.arities_of raw_induct;
@@ -371,7 +370,7 @@
           (vs' @ Ps) rec_names rss' intrs dummies;
         val rlzs = map (fn (r, ind) => Extraction.realizes_of thy (vs' @ Ps) r
           (prop_of ind)) (rs ~~ inducts);
-        val used = foldr add_term_free_names [] rlzs;
+        val used = fold Term.add_free_names rlzs [];
         val rnames = Name.variant_list used (replicate (length inducts) "r");
         val rnames' = Name.variant_list
           (used @ rnames) (replicate (length intrs) "s");
--- a/src/HOL/Tools/inductive_set_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/inductive_set_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -169,7 +169,7 @@
     ({to_set_simps = to_set_simps1, to_pred_simps = to_pred_simps1,
       set_arities = set_arities1, pred_arities = pred_arities1},
      {to_set_simps = to_set_simps2, to_pred_simps = to_pred_simps2,
-      set_arities = set_arities2, pred_arities = pred_arities2}) =
+      set_arities = set_arities2, pred_arities = pred_arities2}) : T =
     {to_set_simps = Thm.merge_thms (to_set_simps1, to_set_simps2),
      to_pred_simps = Thm.merge_thms (to_pred_simps1, to_pred_simps2),
      set_arities = Symtab.merge_list op = (set_arities1, set_arities2),
--- a/src/HOL/Tools/int_arith.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/int_arith.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
-(*  Title:      HOL/int_arith1.ML
-    ID:         $Id$
+(*  Title:      HOL/Tools/int_arith1.ML
     Authors:    Larry Paulson and Tobias Nipkow
 
 Simprocs and decision procedure for linear arithmetic.
@@ -66,12 +65,12 @@
     EQUAL => int_ord (Int.sign i, Int.sign j) 
   | ord => ord);
 
-(*This resembles Term.term_ord, but it puts binary numerals before other
+(*This resembles TermOrd.term_ord, but it puts binary numerals before other
   non-atomic terms.*)
 local open Term 
 in 
 fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) =
-      (case numterm_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
+      (case numterm_ord (t, u) of EQUAL => TermOrd.typ_ord (T, U) | ord => ord)
   | numterm_ord
      (Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) =
      num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w)
@@ -81,7 +80,7 @@
       (case int_ord (size_of_term t, size_of_term u) of
         EQUAL =>
           let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
-            (case hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord)
+            (case TermOrd.hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord)
           end
       | ord => ord)
 and numterms_ord (ts, us) = list_ord numterm_ord (ts, us)
@@ -164,7 +163,7 @@
 (*Express t as a product of (possibly) a numeral with other sorted terms*)
 fun dest_coeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_coeff (~sign) t
   | dest_coeff sign t =
-    let val ts = sort Term.term_ord (dest_prod t)
+    let val ts = sort TermOrd.term_ord (dest_prod t)
         val (n, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, ts)
     in (sign*n, mk_prod (Term.fastype_of t) ts') end;
--- a/src/HOL/Tools/lin_arith.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/lin_arith.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*  Title:      HOL/Tools/lin_arith.ML
-    ID:         $Id$
-    Author:     Tjark Weber and Tobias Nipkow
+    Author:     Tjark Weber and Tobias Nipkow, TU Muenchen
 
 HOL setup for linear arithmetic (see Provers/Arith/fast_lin_arith.ML).
 *)
@@ -99,8 +98,9 @@
             tactics: arith_tactic list};
   val empty = {splits = [], inj_consts = [], discrete = [], tactics = []};
   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}) =
+  fun merge _
+   ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1, tactics= tactics1},
+    {splits= splits2, inj_consts= inj_consts2, discrete= discrete2, tactics= tactics2}) : 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),
--- a/src/HOL/Tools/meson.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/meson.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,11 +1,8 @@
 (*  Title:      HOL/Tools/meson.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1992  University of Cambridge
 
 The MESON resolution proof procedure for HOL.
-
-When making clauses, avoids using the rewriter -- instead uses RS recursively
+When making clauses, avoids using the rewriter -- instead uses RS recursively.
 *)
 
 signature MESON =
@@ -400,7 +397,7 @@
   Also rejects functions whose arguments are Booleans or other functions.*)
 fun is_fol_term thy t =
     Term.is_first_order ["all","All","Ex"] t andalso
-    not (exists (has_bool o fastype_of) (term_vars t)  orelse
+    not (exists_subterm (fn Var (_, T) => has_bool T | _ => false) t  orelse
          has_bool_arg_const t  orelse
          exists_Const (higher_inst_const thy) t orelse
          has_meta_conn t);
@@ -699,4 +696,3 @@
   handle Subscript => Seq.empty;
 
 end;
-
--- a/src/HOL/Tools/metis_tools.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/metis_tools.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -2,7 +2,7 @@
     Author:     Kong W. Susanto and Lawrence C. Paulson, CU Computer Laboratory
     Copyright   Cambridge University 2007
 
-HOL setup for the Metis prover (version 2.0 from 2007).
+HOL setup for the Metis prover.
 *)
 
 signature METIS_TOOLS =
@@ -280,9 +280,10 @@
 
   fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx));
 
-  fun inst_excluded_middle th thy i_atm =
-    let val vx = hd (term_vars (prop_of th))
-        val substs = [(cterm_of thy vx, cterm_of thy i_atm)]
+  fun inst_excluded_middle thy i_atm =
+    let val th = EXCLUDED_MIDDLE
+        val [vx] = Term.add_vars (prop_of th) []
+        val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)]
     in  cterm_instantiate substs th  end;
 
   (* INFERENCE RULE: AXIOM *)
@@ -291,7 +292,7 @@
 
   (* INFERENCE RULE: ASSUME *)
   fun assume_inf ctxt atm =
-    inst_excluded_middle EXCLUDED_MIDDLE
+    inst_excluded_middle
       (ProofContext.theory_of ctxt)
       (singleton (fol_terms_to_hol ctxt) (Metis.Term.Fn atm));
 
@@ -301,12 +302,12 @@
   fun inst_inf ctxt thpairs fsubst th =    
     let val thy = ProofContext.theory_of ctxt
         val i_th   = lookth thpairs th
-        val i_th_vars = term_vars (prop_of i_th)
-        fun find_var x = valOf (List.find (fn Term.Var((a,_),_) => a=x) i_th_vars)
+        val i_th_vars = Term.add_vars (prop_of i_th) []
+        fun find_var x = valOf (List.find (fn ((a,_),_) => a=x) i_th_vars)
         fun subst_translation (x,y) =
               let val v = find_var x
                   val t = fol_term_to_hol_RAW ctxt y (*we call infer_types below*)
-              in  SOME (cterm_of thy v, t)  end
+              in  SOME (cterm_of thy (Var v), t)  end
               handle Option =>
                   (Output.debug (fn() => "List.find failed for the variable " ^ x ^
                                          " in " ^ Display.string_of_thm i_th);
@@ -370,7 +371,7 @@
     end;
 
   (* INFERENCE RULE: REFL *)
-  val refl_x = cterm_of (the_context ()) (hd (term_vars (prop_of REFL_THM)));
+  val refl_x = cterm_of (the_context ()) (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
   val refl_idx = 1 + Thm.maxidx_of REFL_THM;
 
   fun refl_inf ctxt t =
@@ -429,7 +430,7 @@
         val subst' = incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
         val _ = Output.debug (fn () => "subst' " ^ Display.string_of_thm subst')
         val eq_terms = map (pairself (cterm_of thy))
-                           (ListPair.zip (term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
+                           (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
     in  cterm_instantiate eq_terms subst'  end;
 
   val factor = Seq.hd o distinct_subgoals_tac;
--- a/src/HOL/Tools/nat_simprocs.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/nat_simprocs.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
-(*  Title:      HOL/nat_simprocs.ML
-    ID:         $Id$
+(*  Title:      HOL/Tools/nat_simprocs.ML
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2000  University of Cambridge
 
 Simprocs for nat numerals.
 *)
@@ -81,7 +79,7 @@
 
 (*Express t as a product of (possibly) a numeral with other factors, sorted*)
 fun dest_coeff t =
-    let val ts = sort Term.term_ord (dest_prod t)
+    let val ts = sort TermOrd.term_ord (dest_prod t)
         val (n, _, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, one, ts)
     in (n, mk_prod ts') end;
--- a/src/HOL/Tools/numeral_syntax.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/numeral_syntax.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/numeral_syntax.ML
-    ID:         $Id$
     Authors:    Markus Wenzel, TU Muenchen
 
 Concrete syntax for generic numerals -- preserves leading zeros/ones.
@@ -19,12 +18,11 @@
 
 fun mk_bin num =
   let
-    val {leading_zeros = z, value, ...} = Syntax.read_xnum num;
     fun bit b bs = HOLogic.mk_bit b $ bs;
-    fun mk 0 = (* FIXME funpow z (bit 0) *) (Syntax.const @{const_name Int.Pls})
-      | mk ~1 = (* FIXME funpow z (bit 1) *) (Syntax.const @{const_name Int.Min})
+    fun mk 0 = Syntax.const @{const_name Int.Pls}
+      | mk ~1 = Syntax.const @{const_name Int.Min}
       | mk i = let val (q, r) = Integer.div_mod i 2 in bit r (mk q) end;
-  in mk value end;
+  in mk (#value (Syntax.read_xnum num)) end;
 
 in
 
@@ -65,15 +63,18 @@
     else sign ^ implode (replicate z "0") ^ num
   end;
 
+fun syntax_numeral t =
+  Syntax.const "_Numeral" $ (Syntax.const "_numeral" $ Syntax.free (dest_bin_str t));
+
 in
 
 fun numeral_tr' show_sorts (*"number_of"*) (Type ("fun", [_, T])) (t :: ts) =
-      let val t' = Syntax.const "_Numeral" $ Syntax.free (dest_bin_str t) in
-        if not (! show_types) andalso can Term.dest_Type T then t'
-        else Syntax.const Syntax.constrainC $ t' $ Syntax.term_of_typ show_sorts T
-      end
+      let val t' =
+        if not (! show_types) andalso can Term.dest_Type T then syntax_numeral t
+        else Syntax.const Syntax.constrainC $ syntax_numeral t $ Syntax.term_of_typ show_sorts T
+      in list_comb (t', ts) end
   | numeral_tr' _ (*"number_of"*) T (t :: ts) =
-      if T = dummyT then Syntax.const "_Numeral" $ Syntax.free (dest_bin_str t)
+      if T = dummyT then list_comb (syntax_numeral t, ts)
       else raise Match
   | numeral_tr' _ (*"number_of"*) _ _ = raise Match;
 
--- a/src/HOL/Tools/old_primrec_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/old_primrec_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Tools/old_primrec_package.ML
-    ID:         $Id$
-    Author:     Stefan Berghofer, TU Muenchen and Norbert Voelker, FernUni Hagen
+    Author:     Norbert Voelker, FernUni Hagen
+    Author:     Stefan Berghofer, TU Muenchen
 
 Package for defining functions on datatypes by primitive recursion.
 *)
@@ -34,14 +34,13 @@
   same type in all introduction rules*)
 fun unify_consts thy cs intr_ts =
   (let
-    val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
     fun varify (t, (i, ts)) =
       let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify [] t))
       in (maxidx_of_term t', t'::ts) end;
     val (i, cs') = foldr varify (~1, []) cs;
     val (i', intr_ts') = foldr varify (i, []) intr_ts;
-    val rec_consts = fold add_term_consts_2 cs' [];
-    val intr_consts = fold add_term_consts_2 intr_ts' [];
+    val rec_consts = fold Term.add_consts cs' [];
+    val intr_consts = fold Term.add_consts intr_ts' [];
     fun unify (cname, cT) =
       let val consts = map snd (filter (fn (c, _) => c = cname) intr_consts)
       in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
@@ -58,7 +57,7 @@
 fun process_eqn thy eq rec_fns =
   let
     val (lhs, rhs) =
-      if null (term_vars eq) then
+      if null (Term.add_vars eq []) then
         HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
         handle TERM _ => raise RecError "not a proper equation"
       else raise RecError "illegal schematic variable(s)";
@@ -91,7 +90,7 @@
     else
      (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
       check_vars "extra variables on rhs: "
-        (map dest_Free (term_frees rhs) \\ lfrees);
+        (map dest_Free (OldTerm.term_frees rhs) \\ lfrees);
       case AList.lookup (op =) rec_fns fnameT of
         NONE =>
           (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns
@@ -161,7 +160,7 @@
               val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
               val rargs = map fst recs;
               val subs = map (rpair dummyT o fst)
-                (rev (rename_wrt_term rhs rargs));
+                (rev (Term.rename_wrt_term rhs rargs));
               val (rhs', (fnameTs'', fnss'')) =
                   (subst (map (fn ((x, y), z) =>
                                (Free x, (body_index y, Free z)))
--- a/src/HOL/Tools/primrec_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/primrec_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -69,7 +69,7 @@
     else
      (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
       check_vars "extra variables on rhs: "
-        (map dest_Free (term_frees rhs) |> subtract (op =) lfrees
+        (map dest_Free (OldTerm.term_frees rhs) |> subtract (op =) lfrees
           |> filter_out (is_fixed o fst));
       case AList.lookup (op =) rec_fns fname of
         NONE =>
@@ -141,7 +141,7 @@
               val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
               val rargs = map fst recs;
               val subs = map (rpair dummyT o fst)
-                (rev (rename_wrt_term rhs rargs));
+                (rev (Term.rename_wrt_term rhs rargs));
               val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
                 (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss')
                   handle PrimrecError (s, NONE) => primrec_error_eqn s eq
--- a/src/HOL/Tools/recfun_codegen.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/recfun_codegen.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/recfun_codegen.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Code generator for recursive functions.
@@ -46,7 +45,7 @@
   | expand_eta thy (thms as thm :: _) =
       let
         val (_, ty) = Code_Unit.const_typ_eqn thm;
-      in if (null o Term.typ_tvars) ty orelse (null o fst o strip_type) ty
+      in if null (Term.add_tvarsT ty []) orelse (null o fst o strip_type) ty
         then thms
         else map (Code_Unit.expand_eta thy 1) thms
       end;
--- a/src/HOL/Tools/record_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/record_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/record_package.ML
-    ID:         $Id$
     Author:     Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen
 
 Extensible records with structural subtyping in HOL.
@@ -934,7 +933,7 @@
                  then (let
                         val P=cterm_of thy (abstract_over_fun_app trm);
                         val thm = mk_fun_apply_eq trm thy;
-                        val PV = cterm_of thy (hd (term_vars (prop_of thm)));
+                        val PV = cterm_of thy (hd (OldTerm.term_vars (prop_of thm)));
                         val thm' = cterm_instantiate [(PV,P)] thm;
                        in SOME  thm' end handle TERM _ => NONE)
                 else NONE)
@@ -1305,7 +1304,7 @@
         | _ => false);
 
     val goal = nth (Thm.prems_of st) (i - 1);
-    val frees = List.filter (is_recT o type_of) (term_frees goal);
+    val frees = List.filter (is_recT o type_of) (OldTerm.term_frees goal);
 
     fun mk_split_free_tac free induct_thm i =
         let val cfree = cterm_of thy free;
@@ -1384,14 +1383,14 @@
   let
     val ctxt' = fold (Variable.declare_typ o TFree) env ctxt;
     val T = Syntax.read_typ ctxt' raw_T;
-    val env' = Term.add_typ_tfrees (T, env);
+    val env' = OldTerm.add_typ_tfrees (T, env);
   in (T, env') end;
 
 fun cert_typ ctxt raw_T env =
   let
     val thy = ProofContext.theory_of ctxt;
     val T = Type.no_tvars (Sign.certify_typ thy raw_T) handle TYPE (msg, _, _) => error msg;
-    val env' = Term.add_typ_tfrees (T, env);
+    val env' = OldTerm.add_typ_tfrees (T, env);
   in (T, env') end;
 
 
@@ -1426,7 +1425,7 @@
           (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true)
       | [x] => (head_of x, false));
     val rule'' = cterm_instantiate (map (pairself cert) (case (rev params) of
-        [] => (case AList.lookup (op =) (map dest_Free (term_frees (prop_of st))) s of
+        [] => (case AList.lookup (op =) (map dest_Free (OldTerm.term_frees (prop_of st))) s of
           NONE => sys_error "try_param_tac: no such variable"
         | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl),
             (x, Free (s, T))])
@@ -1777,7 +1776,7 @@
     val names = map fst fields;
     val extN = full bname;
     val types = map snd fields;
-    val alphas_fields = foldr add_typ_tfree_names [] types;
+    val alphas_fields = 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);
@@ -2226,7 +2225,7 @@
     val init_env =
       (case parent of
         NONE => []
-      | SOME (types, _) => foldr Term.add_typ_tfrees [] types);
+      | SOME (types, _) => foldr OldTerm.add_typ_tfrees [] types);
 
 
     (* fields *)
--- a/src/HOL/Tools/refute.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/refute.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      HOL/Tools/refute.ML
-    ID:         $Id$
-    Author:     Tjark Weber
-    Copyright   2003-2007
+    Author:     Tjark Weber, TU Muenchen
 
 Finite model generation for HOL formulas, using a SAT solver.
 *)
@@ -391,11 +389,6 @@
 (* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL       *)
 (* ------------------------------------------------------------------------- *)
 
-  (* (''a * 'b) list -> ''a -> 'b *)
-
-  fun lookup xs key =
-    Option.valOf (AList.lookup (op =) xs key);
-
 (* ------------------------------------------------------------------------- *)
 (* typ_of_dtyp: converts a data type ('DatatypeAux.dtyp') into a type        *)
 (*              ('Term.typ'), given type parameters for the data type's type *)
@@ -407,12 +400,12 @@
 
   fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) =
     (* replace a 'DtTFree' variable by the associated type *)
-    lookup typ_assoc (DatatypeAux.DtTFree a)
+    the (AList.lookup (op =) typ_assoc (DatatypeAux.DtTFree a))
     | typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) =
     Type (s, map (typ_of_dtyp descr typ_assoc) ds)
     | typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) =
     let
-      val (s, ds, _) = lookup descr i
+      val (s, ds, _) = the (AList.lookup (op =) descr i)
     in
       Type (s, map (typ_of_dtyp descr typ_assoc) ds)
     end;
@@ -426,7 +419,7 @@
   fun close_form t =
   let
     (* (Term.indexname * Term.typ) list *)
-    val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
+    val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t))
   in
     Library.foldl (fn (t', ((x, i), T)) =>
       (Term.all T) $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
@@ -792,7 +785,7 @@
       (* replace the (at most one) schematic type variable in each axiom *)
       (* by the actual type 'T'                                          *)
       val monomorphic_class_axioms = map (fn (axname, ax) =>
-        (case Term.term_tvars ax of
+        (case Term.add_tvars ax [] of
           [] =>
           (axname, ax)
         | [(idx, S)] =>
@@ -852,14 +845,14 @@
       | Const ("The", T)                =>
         let
           val ax = specialize_type thy ("The", T)
-            (lookup axioms "HOL.the_eq_trivial")
+            (the (AList.lookup (op =) axioms "HOL.the_eq_trivial"))
         in
           collect_this_axiom ("HOL.the_eq_trivial", ax) axs
         end
       | Const ("Hilbert_Choice.Eps", T) =>
         let
           val ax = specialize_type thy ("Hilbert_Choice.Eps", T)
-            (lookup axioms "Hilbert_Choice.someI")
+            (the (AList.lookup (op =) axioms "Hilbert_Choice.someI"))
         in
           collect_this_axiom ("Hilbert_Choice.someI", ax) axs
         end
@@ -944,23 +937,20 @@
 (*               and all mutually recursive IDTs are considered.             *)
 (* ------------------------------------------------------------------------- *)
 
-  (* theory -> Term.term -> Term.typ list *)
-
   fun ground_types thy t =
   let
-    (* Term.typ * Term.typ list -> Term.typ list *)
-    fun collect_types (T, acc) =
+    fun collect_types T acc =
       (case T of
-        Type ("fun", [T1, T2]) => collect_types (T1, collect_types (T2, acc))
+        Type ("fun", [T1, T2]) => collect_types T1 (collect_types T2 acc)
       | Type ("prop", [])      => acc
-      | Type ("set", [T1])     => collect_types (T1, acc)
+      | Type ("set", [T1])     => collect_types T1 acc
       | Type (s, Ts)           =>
         (case DatatypePackage.get_datatype thy s of
           SOME info =>  (* inductive datatype *)
           let
             val index        = #index info
             val descr        = #descr info
-            val (_, typs, _) = lookup descr index
+            val (_, typs, _) = the (AList.lookup (op =) descr index)
             val typ_assoc    = typs ~~ Ts
             (* sanity check: every element in 'dtyps' must be a *)
             (* 'DtTFree'                                        *)
@@ -978,15 +968,15 @@
             in
               case d of
                 DatatypeAux.DtTFree _ =>
-                collect_types (dT, acc)
+                collect_types dT acc
               | DatatypeAux.DtType (_, ds) =>
-                collect_types (dT, foldr collect_dtyp acc ds)
+                collect_types dT (foldr collect_dtyp acc ds)
               | DatatypeAux.DtRec i =>
                 if dT mem acc then
                   acc  (* prevent infinite recursion *)
                 else
                   let
-                    val (_, dtyps, dconstrs) = lookup descr i
+                    val (_, dtyps, dconstrs) = the (AList.lookup (op =) descr i)
                     (* if the current type is a recursive IDT (i.e. a depth *)
                     (* is required), add it to 'acc'                        *)
                     val acc_dT = if Library.exists (fn (_, ds) =>
@@ -1010,11 +1000,11 @@
         | NONE =>
           (* not an inductive datatype, e.g. defined via "typedef" or *)
           (* "typedecl"                                               *)
-          insert (op =) T (foldr collect_types acc Ts))
+          insert (op =) T (fold collect_types Ts acc))
       | TFree _                => insert (op =) T acc
       | TVar _                 => insert (op =) T acc)
   in
-    it_term_types collect_types (t, [])
+    fold_types collect_types t []
   end;
 
 (* ------------------------------------------------------------------------- *)
@@ -1170,7 +1160,7 @@
             let
               val index           = #index info
               val descr           = #descr info
-              val (_, _, constrs) = lookup descr index
+              val (_, _, constrs) = the (AList.lookup (op =) descr index)
             in
               (* recursive datatype? *)
               Library.exists (fn (_, ds) =>
@@ -1289,12 +1279,12 @@
     (* terms containing them (their logical meaning is that there EXISTS a *)
     (* type s.t. ...; to refute such a formula, we would have to show that *)
     (* for ALL types, not ...)                                             *)
-    val _ = null (term_tvars t) orelse
+    val _ = null (Term.add_tvars t []) orelse
       error "Term to be refuted contains schematic type variables"
 
     (* existential closure over schematic variables *)
     (* (Term.indexname * Term.typ) list *)
-    val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
+    val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t))
     (* Term.term *)
     val ex_closure = Library.foldl (fn (t', ((x, i), T)) =>
       (HOLogic.exists_const T) $
@@ -1662,7 +1652,7 @@
       fun interpret_groundtype () =
       let
         (* the model must specify a size for ground types *)
-        val size = (if T = Term.propT then 2 else lookup typs T)
+        val size = if T = Term.propT then 2 else the (AList.lookup (op =) typs T)
         val next = next_idx+size
         (* check if 'maxvars' is large enough *)
         val _    = (if next-1>maxvars andalso maxvars>0 then
@@ -2025,7 +2015,7 @@
             let
               val index               = #index info
               val descr               = #descr info
-              val (_, dtyps, constrs) = lookup descr index
+              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
               val typ_assoc           = dtyps ~~ Ts
               (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
               val _ = if Library.exists (fn d =>
@@ -2149,7 +2139,7 @@
             let
               val index               = #index info
               val descr               = #descr info
-              val (_, dtyps, constrs) = lookup descr index
+              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
               val typ_assoc           = dtyps ~~ Ts
               (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
               val _ = if Library.exists (fn d =>
@@ -2209,7 +2199,7 @@
             let
               val index               = #index info
               val descr               = #descr info
-              val (_, dtyps, constrs) = lookup descr index
+              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
               val typ_assoc           = dtyps ~~ Ts'
               (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
               val _ = if Library.exists (fn d =>
@@ -2405,7 +2395,7 @@
                   (* the index of some mutually recursive IDT               *)
                   val index         = #index info
                   val descr         = #descr info
-                  val (_, dtyps, _) = lookup descr index
+                  val (_, dtyps, _) = the (AList.lookup (op =) descr index)
                   (* sanity check: we assume that the order of constructors *)
                   (*               in 'descr' is the same as the order of   *)
                   (*               corresponding parameters, otherwise the  *)
@@ -2464,7 +2454,7 @@
                         end
                       | DatatypeAux.DtRec i =>
                         let
-                          val (_, ds, _) = lookup descr i
+                          val (_, ds, _) = the (AList.lookup (op =) descr i)
                           val (_, Ts)    = dest_Type T
                         in
                           rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)
@@ -2478,10 +2468,10 @@
                         raise REFUTE ("IDT_recursion_interpreter",
                           "different type associations for the same dtyp"))
                   (* (DatatypeAux.dtyp * Term.typ) list *)
-                  val typ_assoc = List.filter
+                  val typ_assoc = filter
                     (fn (DatatypeAux.DtTFree _, _) => true | (_, _) => false)
                     (rec_typ_assoc []
-                      (#2 (lookup descr idt_index) ~~ (snd o dest_Type) IDT))
+                      (#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT))
                   (* sanity check: typ_assoc must associate types to the   *)
                   (*               elements of 'dtyps' (and only to those) *)
                   val _ = if not (Library.eq_set (dtyps, map fst typ_assoc))
@@ -2605,7 +2595,7 @@
                           SOME args => (n, args)
                         | NONE      => get_cargs_rec (n+1, xs))
                     in
-                      get_cargs_rec (0, lookup mc_intrs idx)
+                      get_cargs_rec (0, the (AList.lookup (op =) mc_intrs idx))
                     end
                   (* computes one entry in 'REC_OPERATORS', and recursively *)
                   (* all entries needed for it, where 'idx' gives the       *)
@@ -2613,7 +2603,7 @@
                   (* int -> int -> interpretation *)
                   fun compute_array_entry idx elem =
                   let
-                    val arr          = lookup REC_OPERATORS idx
+                    val arr          = the (AList.lookup (op =) REC_OPERATORS idx)
                     val (flag, intr) = Array.sub (arr, elem)
                   in
                     if flag then
@@ -2627,7 +2617,7 @@
                         val (c, args) = get_cargs idx elem
                         (* find the indices of the constructor's /recursive/ *)
                         (* arguments                                         *)
-                        val (_, _, constrs) = lookup descr idx
+                        val (_, _, constrs) = the (AList.lookup (op =) descr idx)
                         val (_, dtyps)      = List.nth (constrs, c)
                         val rec_dtyps_args  = List.filter
                           (DatatypeAux.is_rec_type o fst) (dtyps ~~ args)
@@ -2679,7 +2669,7 @@
                         result
                       end
                   end
-                  val idt_size = Array.length (lookup REC_OPERATORS idt_index)
+                  val idt_size = Array.length (the (AList.lookup (op =) REC_OPERATORS idt_index))
                   (* sanity check: the size of 'IDT' should be 'idt_size' *)
                   val _ = if idt_size <> size_of_type thy (typs, []) IDT then
                         raise REFUTE ("IDT_recursion_interpreter",
@@ -2978,8 +2968,8 @@
         (* nat -> nat -> interpretation *)
         fun append m n =
           let
-            val (len_m, off_m) = lookup lenoff_lists m
-            val (len_n, off_n) = lookup lenoff_lists n
+            val (len_m, off_m) = the (AList.lookup (op =) lenoff_lists m)
+            val (len_n, off_n) = the (AList.lookup (op =) lenoff_lists n)
             val len_elem = len_m + len_n
             val off_elem = off_m * power (size_elem, len_n) + off_n
           in
@@ -3267,7 +3257,7 @@
           val (typs, _)           = model
           val index               = #index info
           val descr               = #descr info
-          val (_, dtyps, constrs) = lookup descr index
+          val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
           val typ_assoc           = dtyps ~~ Ts
           (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
           val _ = if Library.exists (fn d =>
--- a/src/HOL/Tools/res_atp.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/res_atp.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,4 @@
-(*  Author: Jia Meng, Cambridge University Computer Laboratory, NICTA
-    ID: $Id$
-    Copyright 2004 University of Cambridge
-*)
+(*  Author: Jia Meng, Cambridge University Computer Laboratory, NICTA *)
 
 signature RES_ATP =
 sig
@@ -443,11 +440,11 @@
 fun delete_type cset = Symtab.delete_safe "HOL.type" cset;
 
 fun tvar_classes_of_terms ts =
-  let val sorts_list = map (map #2 o term_tvars) ts
+  let val sorts_list = map (map #2 o OldTerm.term_tvars) ts
   in  Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list))  end;
 
 fun tfree_classes_of_terms ts =
-  let val sorts_list = map (map #2 o term_tfrees) ts
+  let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts
   in  Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list))  end;
 
 (*fold type constructors*)
@@ -507,11 +504,8 @@
 fun is_taut (Const ("Trueprop", _) $ Const ("True", _)) = true
   | is_taut _ = false;
 
-(*True if the term contains a variable whose (atomic) type is in the given list.*)
-fun has_typed_var tycons =
-  let fun var_tycon (Var (_, Type(a,_))) = a mem_string tycons
-        | var_tycon _ = false
-  in  exists var_tycon o term_vars  end;
+fun has_typed_var tycons = exists_subterm
+  (fn Var (_, Type (a, _)) => member (op =) tycons a | _ => false);
 
 (*Clauses are forbidden to contain variables of these types. The typical reason is that
   they lead to unsoundness. Note that "unit" satisfies numerous equations like ?X=().
--- a/src/HOL/Tools/res_axioms.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/res_axioms.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,4 @@
 (*  Author: Jia Meng, Cambridge University Computer Laboratory
-    ID: $Id$
-    Copyright 2004 University of Cambridge
 
 Transformation of axiom rules (elim/intro/etc) into CNF forms.
 *)
@@ -75,7 +73,7 @@
           (*Existential: declare a Skolem function, then insert into body and continue*)
           let
             val cname = "sko_" ^ s ^ "_" ^ Int.toString (inc nref)
-            val args0 = term_frees xtp  (*get the formal parameter list*)
+            val args0 = OldTerm.term_frees xtp  (*get the formal parameter list*)
             val Ts = map type_of args0
             val extraTs = rhs_extra_types (Ts ---> T) xtp
             val argsx = map (fn T => Free (gensym "vsk", T)) extraTs
@@ -91,7 +89,7 @@
           in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end
       | dec_sko (Const ("All", _) $ (xtp as Abs (a, T, p))) thx =
           (*Universal quant: insert a free variable into body and continue*)
-          let val fname = Name.variant (add_term_names (p, [])) a
+          let val fname = Name.variant (OldTerm.add_term_names (p, [])) a
           in dec_sko (subst_bound (Free (fname, T), p)) thx end
       | dec_sko (Const ("op &", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
       | dec_sko (Const ("op |", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
@@ -105,7 +103,7 @@
       fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) defs =
             (*Existential: declare a Skolem function, then insert into body and continue*)
             let val skos = map (#1 o Logic.dest_equals) defs  (*existing sko fns*)
-                val args = term_frees xtp \\ skos  (*the formal parameters*)
+                val args = OldTerm.term_frees xtp \\ skos  (*the formal parameters*)
                 val Ts = map type_of args
                 val cT = Ts ---> T
                 val id = "sko_" ^ s ^ "_" ^ Int.toString (inc sko_count)
@@ -119,7 +117,7 @@
             end
         | dec_sko (Const ("All",_) $ (xtp as Abs(a,T,p))) defs =
             (*Universal quant: insert a free variable into body and continue*)
-            let val fname = Name.variant (add_term_names (p,[])) a
+            let val fname = Name.variant (OldTerm.add_term_names (p,[])) a
             in dec_sko (subst_bound (Free(fname,T), p)) defs end
         | dec_sko (Const ("op &", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
         | dec_sko (Const ("op |", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
@@ -158,9 +156,9 @@
 
 val monomorphic = not o Term.exists_type (Term.exists_subtype Term.is_TVar);
 
-val [f_B,g_B] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_B}));
-val [g_C,f_C] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_C}));
-val [f_S,g_S] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_S}));
+val [f_B,g_B] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_B}));
+val [g_C,f_C] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_C}));
+val [f_S,g_S] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_S}));
 
 (*FIXME: requires more use of cterm constructors*)
 fun abstract ct =
@@ -495,8 +493,8 @@
       val defs = filter (is_absko o Thm.term_of) newhyps
       val remaining_hyps = filter_out (member (op aconv) (map Thm.term_of defs))
                                       (map Thm.term_of hyps)
-      val fixed = term_frees (concl_of st) @
-                  foldl (gen_union (op aconv)) [] (map term_frees remaining_hyps)
+      val fixed = OldTerm.term_frees (concl_of st) @
+                  foldl (gen_union (op aconv)) [] (map OldTerm.term_frees remaining_hyps)
   in Seq.of_list [LocalDefs.expand (filter (is_okdef fixed o Thm.term_of) defs) st] end;
 
 
--- a/src/HOL/Tools/res_reconstruct.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/res_reconstruct.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,4 @@
-(*  ID:         $Id$
-    Author:     L C Paulson and Claire Quigley
-    Copyright   2004  University of Cambridge
-*)
+(*  Author:     L C Paulson and Claire Quigley, Cambridge University Computer Laboratory *)
 
 (***************************************************************************)
 (*  Code to deal with the transfer of proofs from a prover process         *)
@@ -243,7 +240,7 @@
     singleton (Syntax.check_terms ctxt) (TypeInfer.constrain HOLogic.boolT (fix_sorts vt dt))
   end;
 
-fun gen_all_vars t = fold_rev Logic.all (term_vars t) t;
+fun gen_all_vars t = fold_rev Logic.all (OldTerm.term_vars t) t;
 
 fun ints_of_stree_aux (Int n, ns) = n::ns
   | ints_of_stree_aux (Br(_,ts), ns) = foldl ints_of_stree_aux ns ts;
@@ -253,7 +250,7 @@
 fun decode_tstp vt0 (name, role, ts, annots) ctxt =
   let val deps = case annots of NONE => [] | SOME (source,_) => ints_of_stree source
       val cl = clause_of_strees ctxt vt0 ts
-  in  ((name, role, cl, deps), fold Variable.declare_term (term_frees cl) ctxt)  end;
+  in  ((name, role, cl, deps), fold Variable.declare_term (OldTerm.term_frees cl) ctxt)  end;
 
 fun dest_tstp ((((name, role), ts), annots), chs) =
   case chs of
@@ -408,8 +405,8 @@
 fun add_wanted_prfline ctxt ((lno, t, []), (nlines, lines)) =
       (nlines, (lno, t, []) :: lines)   (*conjecture clauses must be kept*)
   | add_wanted_prfline ctxt ((lno, t, deps), (nlines, lines)) =
-      if eq_types t orelse not (null (term_tvars t)) orelse
-         exists bad_free (term_frees t) orelse
+      if eq_types t orelse not (null (Term.add_tvars t [])) orelse
+         exists_subterm bad_free t orelse
          (not (null lines) andalso   (*final line can't be deleted for these reasons*)
           (length deps < 2 orelse nlines mod (Config.get ctxt modulus) <> 0))   
       then (nlines+1, map (replace_deps (lno, deps)) lines) (*Delete line*)
--- a/src/HOL/Tools/sat_funcs.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/sat_funcs.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,6 @@
 (*  Title:      HOL/Tools/sat_funcs.ML
-    ID:         $Id$
     Author:     Stephan Merz and Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr)
-    Author:     Tjark Weber
-    Copyright   2005-2006
-
+    Author:     Tjark Weber, TU Muenchen
 
 Proof reconstruction from SAT solvers.
 
@@ -323,7 +320,7 @@
 	(* Conjunction.intr_balanced and fold Thm.weaken below, is quadratic for *)
 	(* terms sorted in descending order, while only linear for terms      *)
 	(* sorted in ascending order                                          *)
-	val sorted_clauses = sort (Term.fast_term_ord o pairself Thm.term_of) nontrivial_clauses
+	val sorted_clauses = sort (TermOrd.fast_term_ord o pairself Thm.term_of) nontrivial_clauses
 	val _ = if !trace_sat then
 			tracing ("Sorted non-trivial clauses:\n" ^ cat_lines (map Display.string_of_cterm sorted_clauses))
 		else ()
--- a/src/HOL/Tools/specification_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/specification_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/specification_package.ML
-    ID:         $Id$
     Author:     Sebastian Skalberg, TU Muenchen
 
 Package for defining constants by specification.
@@ -118,7 +117,7 @@
 
         fun proc_single prop =
             let
-                val frees = Term.term_frees prop
+                val frees = OldTerm.term_frees prop
                 val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees
                   orelse error "Specificaton: Only free variables of sort 'type' allowed"
                 val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees)
--- a/src/HOL/Tools/split_rule.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/split_rule.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/split_rule.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, David von Oheimb, and Markus Wenzel, TU Muenchen
 
 Some tools for managing tupled arguments and abstractions in rules.
@@ -105,7 +104,7 @@
 
 (*curries ALL function variables occurring in a rule's conclusion*)
 fun split_rule rl =
-  fold_rev split_rule_var' (Term.term_vars (concl_of rl)) rl
+  fold_rev split_rule_var' (OldTerm.term_vars (concl_of rl)) rl
   |> remove_internal_split
   |> Drule.standard;
 
--- a/src/HOL/Tools/string_syntax.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/string_syntax.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/string_syntax.ML
-    ID:         $Id$
     Author:     Makarius
 
 Concrete syntax for hex chars and strings.
@@ -43,8 +42,8 @@
 fun dest_char (Syntax.Appl [Syntax.Constant "Char", c1, c2]) = dest_chr c1 c2
   | dest_char _ = raise Match;
 
-fun syntax_xstr cs =
-  Syntax.Appl [Syntax.Constant "_xstr", Syntax.Variable (Syntax.implode_xstr cs)];
+fun syntax_string cs =
+  Syntax.Appl [Syntax.Constant "_inner_string", Syntax.Variable (Syntax.implode_xstr cs)];
 
 
 fun char_ast_tr [Syntax.Variable xstr] =
@@ -53,7 +52,7 @@
     | _ => error ("Single character expected: " ^ xstr))
   | char_ast_tr asts = raise AST ("char_ast_tr", asts);
 
-fun char_ast_tr' [c1, c2] = Syntax.Appl [Syntax.Constant "_Char", syntax_xstr [dest_chr c1 c2]]
+fun char_ast_tr' [c1, c2] = Syntax.Appl [Syntax.Constant "_Char", syntax_string [dest_chr c1 c2]]
   | char_ast_tr' _ = raise Match;
 
 
@@ -70,7 +69,7 @@
   | string_ast_tr asts = raise AST ("string_tr", asts);
 
 fun list_ast_tr' [args] = Syntax.Appl [Syntax.Constant "_String",
-        syntax_xstr (map dest_char (Syntax.unfold_ast "_args" args))]
+        syntax_string (map dest_char (Syntax.unfold_ast "_args" args))]
   | list_ast_tr' ts = raise Match;
 
 
--- a/src/HOL/Tools/typecopy_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/typecopy_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/typecopy_package.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Introducing copies of types using trivial typedefs; datatype-like abstraction.
@@ -75,7 +74,8 @@
 fun add_typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
   let
     val ty = Sign.certify_typ thy raw_ty;
-    val vs = AList.make (the_default HOLogic.typeS o AList.lookup (op =) (typ_tfrees ty)) raw_vs;
+    val vs =
+      AList.make (the_default HOLogic.typeS o AList.lookup (op =) (Term.add_tfreesT ty [])) raw_vs;
     val tac = Tactic.rtac UNIV_witness 1;
     fun add_info tyco ( { abs_type = ty_abs, rep_type = ty_rep, Abs_name = c_abs,
       Rep_name = c_rep, Abs_inject = inject,
--- a/src/HOL/Tools/typedef_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Tools/typedef_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -176,7 +176,7 @@
     fun show_names pairs = commas_quote (map fst pairs);
 
     val illegal_vars =
-      if null (term_vars set) andalso null (term_tvars set) then []
+      if null (Term.add_vars set []) andalso null (Term.add_tvars set []) then []
       else ["Illegal schematic variable(s) on rhs"];
 
     val dup_lhs_tfrees =
@@ -188,8 +188,8 @@
       | extras => ["Extra type variables on rhs: " ^ show_names extras]);
 
     val illegal_frees =
-      (case term_frees set of [] => []
-      | xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]);
+      (case Term.add_frees set [] of [] => []
+      | xs => ["Illegal variables on rhs: " ^ show_names xs]);
 
     val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
     val _ = if null errs then () else error (cat_lines errs);
--- a/src/HOL/Univ_Poly.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Univ_Poly.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -2,7 +2,7 @@
     Author:      Amine Chaieb
 *)
 
-header{*Univariate Polynomials*}
+header {* Univariate Polynomials *}
 
 theory Univ_Poly
 imports Plain List
@@ -368,12 +368,13 @@
   "\<not> (finite (UNIV:: 'a set))" 
 proof
   assume F: "finite (UNIV :: 'a set)"
-  have th0: "of_nat ` UNIV \<subseteq> UNIV" by simp
-  from finite_subset[OF th0] have th: "finite (of_nat ` UNIV :: 'a set)" .
-  have th': "inj_on (of_nat::nat \<Rightarrow> 'a) (UNIV)"
-    unfolding inj_on_def by auto
-  from finite_imageD[OF th th'] UNIV_nat_infinite 
-  show False by blast
+  have "finite (UNIV :: nat set)"
+  proof (rule finite_imageD)
+    have "of_nat ` UNIV \<subseteq> UNIV" by simp
+    then show "finite (of_nat ` UNIV :: 'a set)" using F by (rule finite_subset)
+    show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: inj_on_def)
+  qed
+  with UNIV_nat_infinite show False ..
 qed
 
 lemma (in idom_char_0) poly_roots_finite: "(poly p \<noteq> poly []) = 
@@ -580,8 +581,8 @@
 next
   case (Suc n p)
   {assume p0: "poly p a = 0"
-    from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by blast
-    hence pN: "p \<noteq> []" by - (rule ccontr, simp)
+    from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by auto
+    hence pN: "p \<noteq> []" by auto
     from p0[unfolded poly_linear_divides] pN  obtain q where 
       q: "p = [-a, 1] *** q" by blast
     from q h p0 have qh: "length q = n" "poly q \<noteq> poly []" 
--- a/src/HOL/Word/TdThs.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Word/TdThs.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -90,7 +90,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	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Word/WordArith.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -22,7 +22,7 @@
 proof
 qed (unfold word_sle_def word_sless_def, auto)
 
-interpretation signed: linorder ["word_sle" "word_sless"] 
+class_interpretation signed: linorder ["word_sle" "word_sless"] 
   by (rule signed_linorder)
 
 lemmas word_arith_wis = 
@@ -858,11 +858,11 @@
 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:
-  td_ext ["unat::'a::len word => nat" 
-          of_nat 
-          "unats (len_of TYPE('a::len))"
-          "%i. i mod 2 ^ len_of TYPE('a::len)"]
+interpretation word_unat!:
+  td_ext "unat::'a::len word => nat" 
+         of_nat 
+         "unats (len_of TYPE('a::len))"
+         "%i. i mod 2 ^ len_of TYPE('a::len)"
   by (rule td_ext_unat)
 
 lemmas td_unat = word_unat.td_thm
--- a/src/HOL/Word/WordBitwise.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Word/WordBitwise.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -344,11 +344,11 @@
 
 lemmas td_ext_nth = td_ext_nth' [OF refl refl refl, unfolded word_size]
 
-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)}"
-          "(\<lambda>h i. h i \<and> i < len_of TYPE('a::len0))"]
+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)}"
+         "(\<lambda>h i. h i \<and> i < len_of TYPE('a::len0))"
   by (rule td_ext_nth)
 
 declare test_bit.Rep' [simp del]
--- a/src/HOL/Word/WordDefinition.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Word/WordDefinition.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -356,11 +356,11 @@
 
 lemmas int_word_uint = td_ext_uint [THEN td_ext.eq_norm, standard]
 
-interpretation word_uint: 
-  td_ext ["uint::'a::len0 word \<Rightarrow> int" 
-          word_of_int 
-          "uints (len_of TYPE('a::len0))"
-          "\<lambda>w. w mod 2 ^ len_of TYPE('a::len0)"]
+interpretation word_uint!:
+  td_ext "uint::'a::len0 word \<Rightarrow> int" 
+         word_of_int 
+         "uints (len_of TYPE('a::len0))"
+         "\<lambda>w. w mod 2 ^ len_of TYPE('a::len0)"
   by (rule td_ext_uint)
   
 lemmas td_uint = word_uint.td_thm
@@ -368,11 +368,11 @@
 lemmas td_ext_ubin = td_ext_uint 
   [simplified len_gt_0 no_bintr_alt1 [symmetric]]
 
-interpretation word_ubin:
-  td_ext ["uint::'a::len0 word \<Rightarrow> int" 
-          word_of_int 
-          "uints (len_of TYPE('a::len0))"
-          "bintrunc (len_of TYPE('a::len0))"]
+interpretation word_ubin!:
+  td_ext "uint::'a::len0 word \<Rightarrow> int" 
+         word_of_int 
+         "uints (len_of TYPE('a::len0))"
+         "bintrunc (len_of TYPE('a::len0))"
   by (rule td_ext_ubin)
 
 lemma sint_sbintrunc': 
@@ -423,19 +423,19 @@
    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:
-  td_ext ["sint ::'a::len word => int" 
+interpretation word_sint!:
+  td_ext "sint ::'a::len word => int" 
           word_of_int 
           "sints (len_of TYPE('a::len))"
           "%w. (w + 2^(len_of TYPE('a::len) - 1)) mod 2^len_of TYPE('a::len) -
-               2 ^ (len_of TYPE('a::len) - 1)"]
+               2 ^ (len_of TYPE('a::len) - 1)"
   by (rule td_ext_sint)
 
-interpretation word_sbin:
-  td_ext ["sint ::'a::len word => int" 
+interpretation word_sbin!:
+  td_ext "sint ::'a::len word => int" 
           word_of_int 
           "sints (len_of TYPE('a::len))"
-          "sbintrunc (len_of TYPE('a::len) - 1)"]
+          "sbintrunc (len_of TYPE('a::len) - 1)"
   by (rule td_ext_sbin)
 
 lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm, standard]
@@ -635,10 +635,10 @@
   apply simp
   done
 
-interpretation word_bl:
-  type_definition ["to_bl :: 'a::len0 word => bool list"
-                   of_bl  
-                   "{bl. length bl = len_of TYPE('a::len0)}"]
+interpretation word_bl!:
+  type_definition "to_bl :: 'a::len0 word => bool list"
+                  of_bl  
+                  "{bl. length bl = len_of TYPE('a::len0)}"
   by (rule td_bl)
 
 lemma word_size_bl: "size w == size (to_bl w)"
--- a/src/HOL/Word/WordGenLib.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/Word/WordGenLib.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -107,16 +107,16 @@
   apply (rule word_or_not)
   done
 
-interpretation word_bool_alg:
-  boolean ["op AND" "op OR" bitNOT 0 max_word]
+interpretation word_bool_alg!:
+  boolean "op AND" "op OR" bitNOT 0 max_word
   by (rule word_boolean)
 
 lemma word_xor_and_or:
   "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:
-  boolean_xor ["op AND" "op OR" bitNOT 0 max_word "op XOR"]
+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)
   apply (rule boolean_xor_axioms.intro)
@@ -363,7 +363,7 @@
    apply (erule contrapos_pn, simp)
    apply (drule arg_cong[where f=of_nat])
    apply simp
-   apply (subst (asm) word_unat.Rep_Abs_A.Rep_inverse[of n])
+   apply (subst (asm) word_unat.Rep_inverse[of n])
    apply simp
   apply simp
   done
--- a/src/HOL/ex/Abstract_NAT.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/Abstract_NAT.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*
-    ID:         $Id$
     Author:     Makarius
 *)
 
@@ -131,7 +130,7 @@
 
 text {* \medskip Just see that our abstract specification makes sense \dots *}
 
-interpretation NAT [0 Suc]
+interpretation NAT 0 Suc
 proof (rule NAT.intro)
   fix m n
   show "(Suc m = Suc n) = (m = n)" by simp
--- a/src/HOL/ex/LocaleTest2.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/LocaleTest2.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ex/LocaleTest2.thy
-    ID:         $Id$
     Author:     Clemens Ballarin
     Copyright (c) 2007 by Clemens Ballarin
 
@@ -433,8 +432,7 @@
 end
 
 
-interpretation dlo < dlat
-(* TODO: definition syntax is unavailable *)
+sublocale dlo < dlat
 proof
   fix x y
   from total have "is_inf x y (if x \<sqsubseteq> y then x else y)" by (auto simp: is_inf_def)
@@ -445,7 +443,7 @@
   then show "EX sup. is_sup x y sup" by blast
 qed
 
-interpretation dlo < ddlat
+sublocale dlo < ddlat
 proof
   fix x y z
   show "x \<sqinter> (y \<squnion> z) = x \<sqinter> y \<squnion> x \<sqinter> z" (is "?l = ?r")
@@ -470,13 +468,13 @@
 
 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 -
   show "dpo (op <= :: [int, int] => bool)"
     proof qed auto
-  then interpret int: dpo ["op <= :: [int, int] => bool"] .
+  then interpret int: dpo "op <= :: [int, int] => bool" .
     txt {* Gives interpreted version of @{text less_def} (without condition). *}
   show "(dpo.less (op <=) (x::int) y) = (x < y)"
     by (unfold int.less_def) auto
@@ -489,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 -
@@ -498,7 +496,7 @@
     apply (unfold int.is_inf_def int.is_sup_def)
     apply arith+
     done
-  then interpret int: dlat ["op <= :: [int, int] => bool"] .
+  then interpret int: dlat "op <= :: [int, int] => bool" .
   txt {* Interpretation to ease use of definitions, which are
     conditional in general but unconditional after interpretation. *}
   show "dlat.meet (op <=) (x::int) y = min x y"
@@ -513,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. *}
@@ -526,13 +524,13 @@
 
 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 -
   show "dpo (op <= :: [nat, nat] => bool)"
     proof qed auto
-  then interpret nat: dpo ["op <= :: [nat, nat] => bool"] .
+  then interpret nat: dpo "op <= :: [nat, nat] => bool" .
     txt {* Gives interpreted version of @{text less_def} (without condition). *}
   show "dpo.less (op <=) (x::nat) y = (x < y)"
     apply (unfold nat.less_def)
@@ -540,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 -
@@ -549,7 +547,7 @@
     apply (unfold nat.is_inf_def nat.is_sup_def)
     apply arith+
     done
-  then interpret nat: dlat ["op <= :: [nat, nat] => bool"] .
+  then interpret nat: dlat "op <= :: [nat, nat] => bool" .
   txt {* Interpretation to ease use of definitions, which are
     conditional in general but unconditional after interpretation. *}
   show "dlat.meet (op <=) (x::nat) y = min x y"
@@ -564,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. *}
@@ -577,13 +575,13 @@
 
 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 -
   show "dpo (op dvd :: [nat, nat] => bool)"
     proof qed (auto simp: dvd_def)
-  then interpret nat_dvd: dpo ["op dvd :: [nat, nat] => bool"] .
+  then interpret nat_dvd: dpo "op dvd :: [nat, nat] => bool" .
     txt {* Gives interpreted version of @{text less_def} (without condition). *}
   show "dpo.less (op dvd) (x::nat) y = (x dvd y & x ~= y)"
     apply (unfold nat_dvd.less_def)
@@ -591,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 -
@@ -603,7 +601,7 @@
     apply (rule_tac x = "lcm x y" in exI)
     apply (auto intro: lcm_dvd1 lcm_dvd2 lcm_least)
     done
-  then interpret nat_dvd: dlat ["op dvd :: [nat, nat] => bool"] .
+  then interpret nat_dvd: dlat "op dvd :: [nat, nat] => bool" .
   txt {* Interpretation to ease use of definitions, which are
     conditional in general but unconditional after interpretation. *}
   show "dlat.meet (op dvd) (x::nat) y = gcd x y"
@@ -819,7 +817,8 @@
 end
 
 
-locale Dhom = Dgrp prod (infixl "**" 65) one + Dgrp sum (infixl "+++" 60) zero +
+locale Dhom = prod: Dgrp prod one + sum: Dgrp sum zero
+    for prod (infixl "**" 65) and one and sum (infixl "+++" 60) and zero +
   fixes hom
   assumes hom_mult [simp]: "hom (x ** y) = hom x +++ hom y"
 
@@ -838,14 +837,14 @@
 
 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 -
   show "Dmonoid op o (id :: 'a => 'a)" proof qed (simp_all add: o_assoc)
   note Dmonoid = this
 (*
-  from this interpret Dmonoid ["op o" "id :: 'a => 'a"] .
+  from this interpret Dmonoid "op o" "id :: 'a => 'a" .
 *)
   show "Dmonoid.unit (op o) (id :: 'a => 'a) f = bij f"
     apply (unfold Dmonoid.unit_def [OF Dmonoid])
@@ -888,7 +887,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/MIR.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/MIR.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -5899,7 +5899,7 @@
   let 
     val thy = Thm.theory_of_cterm ct;
     val t = Thm.term_of ct;
-    val fs = term_frees t;
+    val fs = OldTerm.term_frees t;
     val vs = fs ~~ (0 upto (length fs - 1));
     val qe = if proofs then @{code mirlfrqe} else @{code mircfrqe};
     val t' = (term_of_fm vs o qe o fm_of_term vs) t;
--- a/src/HOL/ex/ReflectedFerrack.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/ReflectedFerrack.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,4 +1,4 @@
-(*  Title:      Complex/ex/ReflectedFerrack.thy
+(*  Title:      HOL/ex/ReflectedFerrack.thy
     Author:     Amine Chaieb
 *)
 
@@ -2070,7 +2070,7 @@
   let 
     val thy = Thm.theory_of_cterm ct;
     val t = Thm.term_of ct;
-    val fs = term_frees t;
+    val fs = OldTerm.term_frees t;
     val vs = fs ~~ (0 upto (length fs - 1));
     val res = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, term_of_fm vs (@{code linrqe} (fm_of_term vs t))));
   in Thm.cterm_of thy res end
--- a/src/HOL/ex/Reflected_Presburger.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/Reflected_Presburger.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -2039,7 +2039,7 @@
   let
     val thy = Thm.theory_of_cterm ct;
     val t = Thm.term_of ct;
-    val fs = term_frees t;
+    val fs = OldTerm.term_frees t;
     val bs = term_bools [] t;
     val vs = fs ~~ (0 upto (length fs - 1))
     val ps = bs ~~ (0 upto (length bs - 1))
--- a/src/HOL/ex/Tarski.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/Tarski.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -120,7 +120,7 @@
 locale CL = S +
   assumes cl_co:  "cl : CompleteLattice"
 
-interpretation CL < PO
+sublocale CL < PO
 apply (simp_all add: A_def r_def)
 apply unfold_locales
 using cl_co unfolding CompleteLattice_def by auto
@@ -131,7 +131,7 @@
   assumes f_cl:  "(cl,f) : CLF_set" (*was the equivalent "f : CLF_set``{cl}"*)
   defines P_def: "P == fix f A"
 
-interpretation CLF < CL
+sublocale CLF < CL
 apply (simp_all add: A_def r_def)
 apply unfold_locales
 using f_cl unfolding CLF_set_def by auto
--- a/src/HOL/ex/coopertac.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/coopertac.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -47,7 +47,7 @@
     val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
       (foldr HOLogic.mk_imp c rhs, np) ps
     val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
-      (term_frees fm' @ term_vars fm');
+      (OldTerm.term_frees fm' @ OldTerm.term_vars fm');
     val fm2 = foldr mk_all2 fm' vs
   in (fm2, np + length vs, length rhs) end;
 
--- a/src/HOL/ex/linrtac.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/linrtac.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -54,7 +54,7 @@
     val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
       (foldr HOLogic.mk_imp c rhs, np) ps
     val (vs, _) = List.partition (fn t => q orelse (type_of t) = HOLogic.natT)
-      (term_frees fm' @ term_vars fm');
+      (OldTerm.term_frees fm' @ OldTerm.term_vars fm');
     val fm2 = foldr mk_all2 fm' vs
   in (fm2, np + length vs, length rhs) end;
 
--- a/src/HOL/ex/mirtac.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/mirtac.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
-(*  Title:      HOL/Complex/ex/mirtac.ML
-    ID:         $Id$
+(*  Title:      HOL/ex/mirtac.ML
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -74,7 +73,7 @@
     val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
       (foldr HOLogic.mk_imp c rhs, np) ps
     val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
-      (term_frees fm' @ term_vars fm');
+      (OldTerm.term_frees fm' @ OldTerm.term_vars fm');
     val fm2 = foldr mk_all2 fm' vs
   in (fm2, np + length vs, length rhs) end;
 
--- a/src/HOL/ex/reflection.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/reflection.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,19 +1,18 @@
-(*
-    ID:         $Id$
-    Author:     Amine Chaieb, TU Muenchen
+(*  Author:     Amine Chaieb, TU Muenchen
 
 A trial for automatical reification.
 *)
 
-signature REFLECTION = sig
+signature REFLECTION =
+sig
   val genreify_tac: Proof.context -> thm list -> term option -> int -> tactic
   val reflection_tac: Proof.context -> thm list -> thm list -> term option -> int -> tactic
   val gen_reflection_tac: Proof.context -> (cterm -> thm)
     -> thm list -> thm list -> term option -> int -> tactic
 end;
 
-structure Reflection : REFLECTION
-= struct
+structure Reflection : REFLECTION =
+struct
 
 val ext2 = thm "ext2";
 val nth_Cons_0 = thm "nth_Cons_0";
@@ -38,10 +37,10 @@
    val (((_,_),[th']), ctxt') = Variable.import_thms true [th] ctxt
    val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th'))
    fun add_fterms (t as t1 $ t2) = 
-       if exists (fn f => could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t
+       if exists (fn f => Term.could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t
        else add_fterms t1 #> add_fterms t2
      | add_fterms (t as Abs(xn,xT,t')) = 
-       if (fN mem (term_consts t)) then (fn _ => [t]) else (fn _ => [])
+       if exists_Const (fn (c, _) => c = fN) t then (fn _ => [t]) else (fn _ => [])
      | add_fterms _ = I
    val fterms = add_fterms rhs []
    val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt'
@@ -139,7 +138,7 @@
         val (tyenv, tmenv) =
         Pattern.match thy
         ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t)
-        (Envir.type_env (Envir.empty 0),Term.Vartab.empty)
+        (Envir.type_env (Envir.empty 0), Vartab.empty)
         val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv)
         val (fts,its) = 
 	    (map (snd o snd) fnvs,
@@ -191,7 +190,7 @@
    val rhs_P = subst_free subst rhs
    val (tyenv, tmenv) = Pattern.match 
 	                    thy (rhs_P, t)
-	                    (Envir.type_env (Envir.empty 0),Term.Vartab.empty)
+	                    (Envir.type_env (Envir.empty 0), Vartab.empty)
    val sbst = Envir.subst_vars (tyenv, tmenv)
    val sbsT = Envir.typ_subst_TVars tyenv
    val subst_ty = map (fn (n,(s,t)) => (certT (TVar (n, s)), certT t)) 
--- a/src/HOL/ex/svc_funcs.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/ex/svc_funcs.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,8 @@
 (*  Title:      HOL/ex/svc_funcs.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson
     Copyright   1999  University of Cambridge
 
-Translation functions for the interface to SVC
+Translation functions for the interface to SVC.
 
 Based upon the work of Soren T. Heilmann
 
@@ -126,7 +125,7 @@
    pos ["positive"]: true if an assumption, false if a goal*)
  fun expr_of pos t =
   let
-    val params = rev (rename_wrt_term t (Term.strip_all_vars t))
+    val params = rev (Term.rename_wrt_term t (Term.strip_all_vars t))
     and body   = Term.strip_all_body t
     val nat_vars = ref ([] : string list)
     (*translation of a variable: record all natural numbers*)
--- a/src/HOL/main.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/main.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,2 @@
-(*  Title:      HOL/main.ML
-    ID:         $Id$
- 
-Classical Higher-order Logic -- only "Main".
-*)
-
+(*side-entry for HOL-Main*)
 use_thy "Main";
--- a/src/HOL/plain.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOL/plain.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,2 @@
-(*  Title:      HOL/plain.ML
-    ID:         $Id$
- 
-Classical Higher-order Logic -- plain Tool bootstrap.
-*)
-
+(*side-entry for HOL-Plain*)
 use_thy "Plain";
--- a/src/HOLCF/Algebraic.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/Algebraic.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -160,7 +160,7 @@
   assumes f: "\<And>x. f\<cdot>x \<sqsubseteq> x"
   shows "pre_deflation (d oo f)"
 proof
-  interpret d: finite_deflation [d] by fact
+  interpret d: finite_deflation d by fact
   fix x
   show "\<And>x. (d oo f)\<cdot>x \<sqsubseteq> x"
     by (simp, rule trans_less [OF d.less f])
@@ -173,9 +173,9 @@
   assumes f: "\<And>x. f\<cdot>x \<sqsubseteq> x"
   shows "eventual (\<lambda>n. iterate n\<cdot>(d oo f))\<cdot>x = x \<longleftrightarrow> d\<cdot>x = x \<and> f\<cdot>x = x"
 proof -
-  interpret d: finite_deflation [d] by fact
+  interpret d: finite_deflation d by fact
   let ?e = "d oo f"
-  interpret e: pre_deflation ["d oo f"]
+  interpret e: pre_deflation "d oo f"
     using `finite_deflation d` f
     by (rule pre_deflation_d_f)
   let ?g = "eventual (\<lambda>n. iterate n\<cdot>?e)"
@@ -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,8 +370,8 @@
 unfolding alg_defl_principal_def
 by (simp add: Abs_alg_defl_inverse sq_le.ideal_principal)
 
-interpretation alg_defl:
-  ideal_completion [sq_le fd_take alg_defl_principal Rep_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)
 apply (erule Rep_alg_defl_lub)
@@ -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"
@@ -485,7 +485,7 @@
   constrains e :: "'a::profinite \<rightarrow> 'b::profinite"
   shows "\<exists>d. cast\<cdot>d = e oo p"
 proof
-  interpret ep_pair [e p] by fact
+  interpret ep_pair e p by fact
   let ?a = "\<lambda>i. e oo approx i oo p"
   have a: "\<And>i. finite_deflation (?a i)"
     apply (rule finite_deflation_e_d_p)
@@ -516,7 +516,7 @@
     "\<And>i. finite_deflation (a i)"
     "(\<Squnion>i. a i) = ID"
 proof
-  interpret ep_pair [e p] by fact
+  interpret ep_pair e p by fact
   let ?a = "\<lambda>i. p oo cast\<cdot>(approx i\<cdot>d) oo e"
   show "\<And>i. finite_deflation (?a i)"
     apply (rule finite_deflation_p_d_e)
--- a/src/HOLCF/Bifinite.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/Bifinite.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -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	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/CompactBasis.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -46,8 +46,8 @@
 
 lemmas approx_Rep_compact_basis = Rep_compact_take [symmetric]
 
-interpretation compact_basis:
-  basis_take [sq_le compact_take]
+interpretation compact_basis!:
+  basis_take sq_le compact_take
 proof
   fix n :: nat and a :: "'a compact_basis"
   show "compact_take n a \<sqsubseteq> a"
@@ -92,8 +92,8 @@
   approximants :: "'a \<Rightarrow> 'a compact_basis set" where
   "approximants = (\<lambda>x. {a. Rep_compact_basis a \<sqsubseteq> x})"
 
-interpretation compact_basis:
-  ideal_completion [sq_le compact_take Rep_compact_basis approximants]
+interpretation compact_basis!:
+  ideal_completion sq_le compact_take Rep_compact_basis approximants
 proof
   fix w :: 'a
   show "preorder.ideal sq_le (approximants w)"
@@ -244,7 +244,7 @@
   assumes "ab_semigroup_idem_mult f"
   shows "fold_pd g f (PDPlus t u) = f (fold_pd g f t) (fold_pd g f u)"
 proof -
-  interpret ab_semigroup_idem_mult [f] by fact
+  class_interpret ab_semigroup_idem_mult [f] by fact
   show ?thesis unfolding fold_pd_def Rep_PDPlus by (simp add: image_Un fold1_Un2)
 qed
 
--- a/src/HOLCF/Completion.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/Completion.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -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	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/ConvexPD.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -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,8 +178,8 @@
 unfolding convex_principal_def
 by (simp add: Abs_convex_pd_inverse convex_le.ideal_principal)
 
-interpretation convex_pd:
-  ideal_completion [convex_le pd_take convex_principal Rep_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)
 apply (rule pd_take_idem)
@@ -296,7 +296,7 @@
 apply (simp add: PDPlus_absorb)
 done
 
-interpretation aci_convex_plus: ab_semigroup_idem_mult ["op +\<natural>"]
+class_interpretation aci_convex_plus: ab_semigroup_idem_mult ["op +\<natural>"]
   by unfold_locales
     (rule convex_plus_assoc convex_plus_commute convex_plus_absorb)+
 
--- a/src/HOLCF/Deflation.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/Deflation.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -81,10 +81,10 @@
   assumes "deflation g"
   shows "f \<sqsubseteq> g \<Longrightarrow> f\<cdot>(g\<cdot>x) = f\<cdot>x"
 proof (rule antisym_less)
-  interpret g: deflation [g] by fact
+  interpret g: deflation g by fact
   from g.less show "f\<cdot>(g\<cdot>x) \<sqsubseteq> f\<cdot>x" by (rule monofun_cfun_arg)
 next
-  interpret f: deflation [f] by fact
+  interpret f: deflation f by fact
   assume "f \<sqsubseteq> g" hence "f\<cdot>x \<sqsubseteq> g\<cdot>x" by (rule monofun_cfun_fun)
   hence "f\<cdot>(f\<cdot>x) \<sqsubseteq> f\<cdot>(g\<cdot>x)" by (rule monofun_cfun_arg)
   also have "f\<cdot>(f\<cdot>x) = f\<cdot>x" by (rule f.idem)
@@ -219,7 +219,7 @@
   assumes "deflation d"
   shows "deflation (e oo d oo p)"
 proof
-  interpret deflation [d] by fact
+  interpret deflation d by fact
   fix x :: 'b
   show "(e oo d oo p)\<cdot>((e oo d oo p)\<cdot>x) = (e oo d oo p)\<cdot>x"
     by (simp add: idem)
@@ -231,7 +231,7 @@
   assumes "finite_deflation d"
   shows "finite_deflation (e oo d oo p)"
 proof
-  interpret finite_deflation [d] by fact
+  interpret finite_deflation d by fact
   fix x :: 'b
   show "(e oo d oo p)\<cdot>((e oo d oo p)\<cdot>x) = (e oo d oo p)\<cdot>x"
     by (simp add: idem)
@@ -250,7 +250,7 @@
   assumes d: "\<And>x. d\<cdot>x \<sqsubseteq> e\<cdot>(p\<cdot>x)"
   shows "deflation (p oo d oo e)"
 proof -
-  interpret d: deflation [d] by fact
+  interpret d: deflation d by fact
   {
     fix x
     have "d\<cdot>(e\<cdot>x) \<sqsubseteq> e\<cdot>x"
@@ -287,7 +287,7 @@
   assumes d: "\<And>x. d\<cdot>x \<sqsubseteq> e\<cdot>(p\<cdot>x)"
   shows "finite_deflation (p oo d oo e)"
 proof -
-  interpret d: finite_deflation [d] by fact
+  interpret d: finite_deflation d by fact
   show ?thesis
   proof (intro_locales)
     have "deflation d" ..
@@ -316,8 +316,8 @@
   assumes "ep_pair e1 p" and "ep_pair e2 p"
   shows "e1 \<sqsubseteq> e2"
 proof (rule less_cfun_ext)
-  interpret e1: ep_pair [e1 p] by fact
-  interpret e2: ep_pair [e2 p] by fact
+  interpret e1: ep_pair e1 p by fact
+  interpret e2: ep_pair e2 p by fact
   fix x
   have "e1\<cdot>(p\<cdot>(e2\<cdot>x)) \<sqsubseteq> e2\<cdot>x"
     by (rule e1.e_p_less)
@@ -333,8 +333,8 @@
   assumes "ep_pair e p1" and "ep_pair e p2"
   shows "p1 \<sqsubseteq> p2"
 proof (rule less_cfun_ext)
-  interpret p1: ep_pair [e p1] by fact
-  interpret p2: ep_pair [e p2] by fact
+  interpret p1: ep_pair e p1 by fact
+  interpret p2: ep_pair e p2 by fact
   fix x
   have "e\<cdot>(p1\<cdot>x) \<sqsubseteq> x"
     by (rule p1.e_p_less)
@@ -357,8 +357,8 @@
   assumes "ep_pair e1 p1" and "ep_pair e2 p2"
   shows "ep_pair (e2 oo e1) (p1 oo p2)"
 proof
-  interpret ep1: ep_pair [e1 p1] by fact
-  interpret ep2: ep_pair [e2 p2] by fact
+  interpret ep1: ep_pair e1 p1 by fact
+  interpret ep2: ep_pair e2 p2 by fact
   fix x y
   show "(p1 oo p2)\<cdot>((e2 oo e1)\<cdot>x) = x"
     by simp
--- a/src/HOLCF/Fixrec.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/Fixrec.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -216,19 +216,19 @@
 
 syntax
   "_pat" :: "'a"
-  "_var" :: "'a"
+  "_variable" :: "'a"
   "_noargs" :: "'a"
 
 translations
-  "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_var p r)"
-  "_var (_args x y) r" => "CONST csplit\<cdot>(_var x (_var y r))"
-  "_var _noargs r" => "CONST unit_when\<cdot>r"
+  "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)"
+  "_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))"
+  "_variable _noargs r" => "CONST unit_when\<cdot>r"
 
 parse_translation {*
 (* rewrites (_pat x) => (return) *)
-(* rewrites (_var x t) => (Abs_CFun (%x. t)) *)
+(* rewrites (_variable x t) => (Abs_CFun (%x. t)) *)
   [("_pat", K (Syntax.const "Fixrec.return")),
-   mk_binder_tr ("_var", "Abs_CFun")];
+   mk_binder_tr ("_variable", "Abs_CFun")];
 *}
 
 text {* Printing Case expressions *}
@@ -250,7 +250,7 @@
             val abs = case t of Abs abs => abs
                 | _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0);
             val (x, t') = atomic_abs_tr' abs;
-          in (Syntax.const "_var" $ x, t') end
+          in (Syntax.const "_variable" $ x, t') end
     |   dest_LAM _ = raise Match; (* too few vars: abort translation *)
 
     fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] =
@@ -261,7 +261,7 @@
 *}
 
 translations
-  "x" <= "_match Fixrec.return (_var x)"
+  "x" <= "_match Fixrec.return (_variable x)"
 
 
 subsection {* Pattern combinators for data constructors *}
@@ -320,18 +320,18 @@
 
 text {* Parse translations (variables) *}
 translations
-  "_var (XCONST cpair\<cdot>x\<cdot>y) r" => "_var (_args x y) r"
-  "_var (XCONST spair\<cdot>x\<cdot>y) r" => "_var (_args x y) r"
-  "_var (XCONST sinl\<cdot>x) r" => "_var x r"
-  "_var (XCONST sinr\<cdot>x) r" => "_var x r"
-  "_var (XCONST up\<cdot>x) r" => "_var x r"
-  "_var (XCONST TT) r" => "_var _noargs r"
-  "_var (XCONST FF) r" => "_var _noargs r"
-  "_var (XCONST ONE) r" => "_var _noargs r"
+  "_variable (XCONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+  "_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+  "_variable (XCONST sinl\<cdot>x) r" => "_variable x r"
+  "_variable (XCONST sinr\<cdot>x) r" => "_variable x r"
+  "_variable (XCONST up\<cdot>x) r" => "_variable x r"
+  "_variable (XCONST TT) r" => "_variable _noargs r"
+  "_variable (XCONST FF) r" => "_variable _noargs r"
+  "_variable (XCONST ONE) r" => "_variable _noargs r"
 
 translations
-  "_var (CONST cpair\<cdot>x\<cdot>y) r" => "_var (_args x y) r"
-  "_var (CONST spair\<cdot>x\<cdot>y) r" => "_var (_args x y) r"
+  "_variable (CONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+  "_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
 
 text {* Print translations *}
 translations
@@ -437,14 +437,14 @@
 
 text {* Parse translations (variables) *}
 translations
-  "_var _ r" => "_var _noargs r"
-  "_var (_as_pat x y) r" => "_var (_args x y) r"
-  "_var (_lazy_pat x) r" => "_var x r"
+  "_variable _ r" => "_variable _noargs r"
+  "_variable (_as_pat x y) r" => "_variable (_args x y) r"
+  "_variable (_lazy_pat x) r" => "_variable x r"
 
 text {* Print translations *}
 translations
   "_" <= "_match (CONST wild_pat) _noargs"
-  "_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_var x) v)"
+  "_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_variable x) v)"
   "_lazy_pat (_match p v)" <= "_match (CONST lazy_pat p) v"
 
 text {* Lazy patterns in lambda abstractions *}
--- a/src/HOLCF/LowerPD.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/LowerPD.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -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,8 +133,8 @@
 unfolding lower_principal_def
 by (simp add: Abs_lower_pd_inverse lower_le.ideal_principal)
 
-interpretation lower_pd:
-  ideal_completion [lower_le pd_take lower_principal Rep_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)
 apply (rule pd_take_idem)
@@ -250,7 +250,7 @@
 apply (simp add: PDPlus_absorb)
 done
 
-interpretation aci_lower_plus: ab_semigroup_idem_mult ["op +\<flat>"]
+class_interpretation aci_lower_plus: ab_semigroup_idem_mult ["op +\<flat>"]
   by unfold_locales
     (rule lower_plus_assoc lower_plus_commute lower_plus_absorb)+
 
--- a/src/HOLCF/Tools/domain/domain_syntax.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/Tools/domain/domain_syntax.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -89,7 +89,7 @@
     fun arg1 n (con,_,args) = foldr cabs (expvar n) (argvars n args);
     fun when1 n m = if n = m then arg1 n else K (Constant "UU");
 
-    fun app_var x = mk_appl (Constant "_var") [x, Variable "rhs"];
+    fun app_var x = mk_appl (Constant "_variable") [x, Variable "rhs"];
     fun app_pat x = mk_appl (Constant "_pat") [x];
     fun args_list [] = Constant "_noargs"
     |   args_list xs = foldr1 (app "_args") xs;
--- a/src/HOLCF/Universal.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/Universal.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -226,13 +226,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)
@@ -281,8 +281,8 @@
 unfolding udom_principal_def
 by (simp add: Abs_udom_inverse udom.ideal_principal)
 
-interpretation udom:
-  ideal_completion [ubasis_le ubasis_take udom_principal Rep_udom]
+interpretation udom!:
+  ideal_completion ubasis_le ubasis_take udom_principal Rep_udom
 apply unfold_locales
 apply (rule ideal_Rep_udom)
 apply (erule Rep_udom_lub)
--- a/src/HOLCF/UpperPD.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/HOLCF/UpperPD.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -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,8 +131,8 @@
 unfolding upper_principal_def
 by (simp add: Abs_upper_pd_inverse upper_le.ideal_principal)
 
-interpretation upper_pd:
-  ideal_completion [upper_le pd_take upper_principal Rep_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)
 apply (rule pd_take_idem)
@@ -248,7 +248,7 @@
 apply (simp add: PDPlus_absorb)
 done
 
-interpretation aci_upper_plus: ab_semigroup_idem_mult ["op +\<sharp>"]
+class_interpretation aci_upper_plus: ab_semigroup_idem_mult ["op +\<sharp>"]
   by unfold_locales
     (rule upper_plus_assoc upper_plus_commute upper_plus_absorb)+
 
--- a/src/Provers/Arith/cancel_factor.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Provers/Arith/cancel_factor.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Provers/Arith/cancel_factor.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 Cancel common constant factor from balanced exression (e.g. =, <=, < of sums).
@@ -36,7 +35,7 @@
       if t aconv u then (u, k + 1) :: uks
       else (t, 1) :: (u, k) :: uks;
 
-fun count_terms ts = foldr ins_term (sort Term.term_ord ts, []);
+fun count_terms ts = foldr ins_term (sort TermOrd.term_ord ts, []);
 
 
 (* divide sum *)
--- a/src/Provers/Arith/cancel_sums.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Provers/Arith/cancel_sums.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Provers/Arith/cancel_sums.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 Cancel common summands of balanced expressions:
@@ -38,11 +37,11 @@
 fun cons2 y (x, ys, z) = (x, y :: ys, z);
 fun cons12 x y (xs, ys, z) = (x :: xs, y :: ys, z);
 
-(*symmetric difference of multisets -- assumed to be sorted wrt. Logic.term_ord*)
+(*symmetric difference of multisets -- assumed to be sorted wrt. TermOrd.term_ord*)
 fun cancel ts [] vs = (ts, [], vs)
   | cancel [] us vs = ([], us, vs)
   | cancel (t :: ts) (u :: us) vs =
-      (case Term.term_ord (t, u) of
+      (case TermOrd.term_ord (t, u) of
         EQUAL => cancel ts us (t :: vs)
       | LESS => cons1 t (cancel ts (u :: us) vs)
       | GREATER => cons2 u (cancel (t :: ts) us vs));
@@ -64,7 +63,7 @@
   | SOME bal =>
       let
         val thy = ProofContext.theory_of (Simplifier.the_context ss);
-        val (ts, us) = pairself (sort Term.term_ord o Data.dest_sum) bal;
+        val (ts, us) = pairself (sort TermOrd.term_ord o Data.dest_sum) bal;
         val (ts', us', vs) = cancel ts us [];
       in
         if null vs then NONE
--- a/src/Provers/classical.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Provers/classical.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      Provers/classical.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1992  University of Cambridge
 
 Theorem prover for classical reasoning, including predicate calculus, set
 theory, etc.
@@ -810,9 +808,7 @@
     (fn (prem,i) =>
       let val deti =
           (*No Vars in the goal?  No need to backtrack between goals.*)
-          case term_vars prem of
-              []        => DETERM
-            | _::_      => I
+          if exists_subterm (fn Var _ => true | _ => false) prem then DETERM else I
       in  SELECT_GOAL (TRY (safe_tac cs) THEN
                        DEPTH_SOLVE (deti (depth_tac cs m 1))) i
       end);
--- a/src/Provers/coherent.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Provers/coherent.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,6 @@
 (*  Title:      Provers/coherent.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
-                Marc Bezem, Institutt for Informatikk, Universitetet i Bergen 
+    Author:     Marc Bezem, Institutt for Informatikk, Universitetet i Bergen 
 
 Prover for coherent logic, see e.g.
 
@@ -39,7 +38,7 @@
   ClPrf of thm * (Type.tyenv * Envir.tenv) * ((indexname * typ) * term) list *
   int list * (term list * cl_prf) list;
 
-fun is_atomic t = null (term_consts t inter Data.operator_names);
+val is_atomic = not o exists_Const (member (op =) Data.operator_names o #1);
 
 local open Conv in
 
--- a/src/Provers/eqsubst.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Provers/eqsubst.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,13 +1,12 @@
 (*  Title:      Provers/eqsubst.ML
-    ID:         $Id$
-    Author:     Lucas Dixon, University of Edinburgh, lucas.dixon@ed.ac.uk
+    Author:     Lucas Dixon, University of Edinburgh
 
 A proof method to perform a substiution using an equation.
 *)
 
 signature EQSUBST =
 sig
-  (* a type abriviation for match information *)
+  (* a type abbreviation for match information *)
   type match =
        ((indexname * (sort * typ)) list (* type instantiations *)
         * (indexname * (typ * term)) list) (* term instantiations *)
@@ -224,7 +223,7 @@
       (* is it OK to ignore the type instantiation info?
          or should I be using it? *)
       val typs_unify =
-          SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Term.Vartab.empty, ix))
+          SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Vartab.empty, ix))
             handle Type.TUNIFY => NONE;
     in
       case typs_unify of
--- a/src/Provers/order.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Provers/order.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,8 +1,6 @@
-(*
-  Title:	Transitivity reasoner for partial and linear orders
-  Id:		$Id$
-  Author:	Oliver Kutter
-  Copyright:	TU Muenchen
+(*  Author:     Oliver Kutter, TU Muenchen
+
+Transitivity reasoner for partial and linear orders.
 *)
 
 (* TODO: reduce number of input thms *)
@@ -1234,7 +1232,7 @@
 
     SUBGOAL (fn (A, n, sign) =>
      let
-      val rfrees = map Free (rename_wrt_term A (Logic.strip_params A))
+      val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
       val Hs = map prop_of prems @ map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
       val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
       val lesss = List.concat (ListPair.map (mkasm decomp less_thms sign) (Hs, 0 upto (length Hs - 1)))
--- a/src/Provers/quasi.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Provers/quasi.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,7 @@
-(* 
-   Title:	Reasoner for simple transitivity and quasi orders. 
-   Id:		$Id$
-   Author:	Oliver Kutter
-   Copyright:	TU Muenchen
- *)
+(*  Author:     Oliver Kutter, TU Muenchen
+
+Reasoner for simple transitivity and quasi orders.
+*)
 
 (* 
  
@@ -557,7 +555,7 @@
 (* trans_tac - solves transitivity chains over <= *)
 val trans_tac  =  SUBGOAL (fn (A, n, sign) =>
  let
-  val rfrees = map Free (rename_wrt_term A (Logic.strip_params A))
+  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
   val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
   val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
   val lesss = List.concat (ListPair.map (mkasm_trans  sign) (Hs, 0 upto (length Hs - 1)))
@@ -578,7 +576,7 @@
 (* quasi_tac - solves quasi orders *)
 val quasi_tac = SUBGOAL (fn (A, n, sign) =>
  let
-  val rfrees = map Free (rename_wrt_term A (Logic.strip_params A))
+  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
   val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
   val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
   val lesss = List.concat (ListPair.map (mkasm_quasi sign) (Hs, 0 upto (length Hs - 1)))
--- a/src/Pure/Concurrent/future.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Concurrent/future.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -116,7 +116,7 @@
   ConditionVar.wait (cond, lock);
 
 fun wait_timeout timeout = (*requires SYNCHRONIZED*)
-  ConditionVar.waitUntil (cond, lock, Time.+ (Time.now (), timeout));
+  ignore (ConditionVar.waitUntil (cond, lock, Time.+ (Time.now (), timeout)));
 
 fun notify_all () = (*requires SYNCHRONIZED*)
   ConditionVar.broadcast cond;
@@ -139,6 +139,9 @@
 
 (* execute *)
 
+fun do_cancel group = (*requires SYNCHRONIZED*)
+  change canceled (insert Task_Queue.eq_group group);
+
 fun execute name (task, group, run) =
   let
     val _ = trace_active ();
@@ -147,7 +150,7 @@
      (change queue (Task_Queue.finish task);
       if ok then ()
       else if Task_Queue.cancel (! queue) group then ()
-      else change canceled (cons group);
+      else do_cancel group;
       notify_all ()));
   in () end;
 
@@ -205,7 +208,8 @@
     val _ = if continue then () else scheduler := NONE;
 
     val _ = notify_all ();
-    val _ = wait_timeout (Time.fromSeconds 3);
+    val _ = interruptible (fn () => wait_timeout (Time.fromSeconds 1)) ()
+      handle Exn.Interrupt => List.app do_cancel (Task_Queue.cancel_all (! queue));
   in continue end;
 
 fun scheduler_loop () =
@@ -306,7 +310,7 @@
 (*cancel: present and future group members will be interrupted eventually*)
 fun cancel x =
  (scheduler_check "cancel check";
-  SYNCHRONIZED "cancel" (fn () => (change canceled (cons (group_of x)); notify_all ())));
+  SYNCHRONIZED "cancel" (fn () => (do_cancel (group_of x); notify_all ())));
 
 
 (*global join and shutdown*)
--- a/src/Pure/Concurrent/task_queue.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Concurrent/task_queue.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -6,11 +6,12 @@
 
 signature TASK_QUEUE =
 sig
-  eqtype task
+  type task
   val new_task: int -> task
   val pri_of_task: task -> int
   val str_of_task: task -> string
-  eqtype group
+  type group
+  val eq_group: group * group -> bool
   val new_group: unit -> group
   val invalidate_group: group -> unit
   val str_of_group: group -> string
@@ -24,11 +25,12 @@
     (((task * group * (unit -> bool)) * task list) option * queue)
   val interrupt: queue -> task -> unit
   val interrupt_external: queue -> string -> unit
+  val cancel: queue -> group -> bool
+  val cancel_all: queue -> group list
   val finish: task -> queue -> queue
-  val cancel: queue -> group -> bool
 end;
 
-structure Task_Queue: TASK_QUEUE =
+structure Task_Queue:> TASK_QUEUE =
 struct
 
 (* tasks *)
@@ -46,6 +48,7 @@
 (* groups *)
 
 datatype group = Group of serial * bool ref;
+fun eq_group (Group (gid1, _), Group (gid2, _)) = gid1 = gid2;
 
 fun new_group () = Group (serial (), ref true);
 
@@ -150,16 +153,26 @@
   | NONE => ());
 
 
-(* misc operations *)
+(* termination *)
 
 fun cancel (Queue {groups, jobs, ...}) (group as Group (gid, _)) =
   let
     val _ = invalidate_group group;
     val tasks = Inttab.lookup_list groups gid;
-    val running = fold (get_job jobs #> (fn Running thread => cons thread | _ => I)) tasks [];
+    val running = fold (get_job jobs #> (fn Running t => insert Thread.equal t | _ => I)) tasks [];
     val _ = List.app SimpleThread.interrupt running;
   in null running end;
 
+fun cancel_all (Queue {jobs, ...}) =
+  let
+    fun cancel_job (group, job) (groups, running) =
+      (invalidate_group group;
+        (case job of Running t => (insert eq_group group groups, insert Thread.equal t running)
+        | _ => (groups, running)));
+    val (groups, running) = Task_Graph.fold (cancel_job o #1 o #2) jobs ([], []);
+    val _ = List.app SimpleThread.interrupt running;
+  in groups end;
+
 fun finish task (Queue {groups, jobs}) =
   let
     val Group (gid, _) = get_group jobs task;
--- a/src/Pure/General/binding.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/binding.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -26,6 +26,8 @@
   val base_name: T -> string
   val pos_of: T -> Position.T
   val dest: T -> (string * bool) list * string
+  val separator: string
+  val is_qualified: string -> bool
   val display: T -> string
 end
 
@@ -39,6 +41,17 @@
 val unique_names = ref true;
 
 
+(** qualification **)
+
+val separator = ".";
+val is_qualified = exists_string (fn s => s = separator);
+
+fun reject_qualified kind s =
+  if is_qualified s then
+    error ("Attempt to declare qualified " ^ kind ^ " " ^ quote s)
+  else s;
+
+
 (** binding representation **)
 
 datatype T = Binding of ((string * bool) list * string) * Position.T;
@@ -54,13 +67,14 @@
 
 fun qualify_base path name =
   if path = "" orelse name = "" then name
-  else path ^ "." ^ name;
+  else path ^ separator ^ name;
 
 val qualify = map_base o qualify_base;
   (*FIXME should all operations on bare names move here from name_space.ML ?*)
 
-fun add_prefix sticky prfx b = if prfx = "" then error "empty name prefix"
-  else (map_binding o apfst) (cons (prfx, sticky)) b;
+fun add_prefix sticky "" b = b
+  | add_prefix sticky prfx b = (map_binding o apfst)
+      (cons ((*reject_qualified "prefix"*) prfx, sticky)) b;
 
 fun map_prefix f (Binding ((prefix, name), pos)) =
   f prefix (name_pos (name, pos));
--- a/src/Pure/General/buffer.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/buffer.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/buffer.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Efficient text buffers.
@@ -10,7 +9,6 @@
   type T
   val empty: T
   val add: string -> T -> T
-  val add_substring: substring -> T -> T
   val markup: Markup.T -> (T -> T) -> T -> T
   val content: T -> string
   val output: T -> TextIO.outstream -> unit
@@ -19,46 +17,18 @@
 structure Buffer: BUFFER =
 struct
 
-(* datatype *)
-
-datatype T =
-    EmptyBuffer
-  | String of string * T
-  | Substring of substring * T;
+datatype T = Buffer of string list;
 
-val empty = EmptyBuffer;
-
+val empty = Buffer [];
 
-(* add content *)
-
-fun add s buf = if s = "" then buf else String (s, buf);
-fun add_substring s buf = if Substring.isEmpty s then buf else Substring (s, buf);
+fun add "" buf = buf
+  | add x (Buffer xs) = Buffer (x :: xs);
 
 fun markup m body =
   let val (bg, en) = Markup.output m
   in add bg #> body #> add en end;
 
-
-(* cumulative content *)
-
-fun rev_content EmptyBuffer acc = acc
-  | rev_content (String (s, buf)) acc = rev_content buf (s :: acc)
-  | rev_content (Substring (s, buf)) acc = rev_content buf (Substring.string s :: acc);
-
-fun content buf = implode (rev_content buf []);
-
-
-(* file output *)
-
-fun rev_buffer EmptyBuffer acc = acc
-  | rev_buffer (String (s, buf)) acc = rev_buffer buf (String (s, acc))
-  | rev_buffer (Substring (s, buf)) acc = rev_buffer buf (Substring (s, acc));
-
-fun output buffer stream =
-  let
-    fun rev_output EmptyBuffer = ()
-      | rev_output (String (s, buf)) = (TextIO.output (stream, s); rev_output buf)
-      | rev_output (Substring (s, buf)) = (TextIO.outputSubstr (stream, s); rev_output buf);
-  in rev_output (rev_buffer buffer empty) end;
+fun content (Buffer xs) = implode (rev xs);
+fun output (Buffer xs) stream = List.app (fn s => TextIO.output (stream, s)) (rev xs);
 
 end;
--- a/src/Pure/General/event_bus.scala	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/event_bus.scala	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,8 @@
 /*  Title:      Pure/General/event_bus.scala
     Author:     Makarius
 
-Generic event bus.
+Generic event bus with multiple handlers and optional exception
+logging.
 */
 
 package isabelle
@@ -9,7 +10,10 @@
 import scala.collection.mutable.ListBuffer
 
 
-class EventBus[Event] {
+class EventBus[Event]
+{
+  /* event handlers */
+
   type Handler = Event => Unit
   private val handlers = new ListBuffer[Handler]
 
@@ -19,5 +23,16 @@
   def -= (h: Handler) = synchronized { handlers -= h }
   def - (h: Handler) = { this -= h; this }
 
-  def event(e: Event) = synchronized { handlers.toList } foreach (h => h(e))
+
+  /* event invocation */
+
+  var logger: Throwable => Unit = throw _
+
+  def event(x: Event) = {
+    val log = logger
+    for (h <- synchronized { handlers.toList }) {
+      try { h(x) }
+      catch { case e: Throwable => log(e) }
+    }
+  }
 }
--- a/src/Pure/General/markup.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/markup.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/markup.ML
-    ID:         $Id$
     Author:     Makarius
 
 Common markup elements.
@@ -52,11 +51,9 @@
   val skolemN: string val skolem: T
   val boundN: string val bound: T
   val varN: string val var: T
-  val numN: string val num: T
-  val floatN: string val float: T
-  val xnumN: string val xnum: T
-  val xstrN: string val xstr: T
+  val numeralN: string val numeral: T
   val literalN: string val literal: T
+  val inner_stringN: string val inner_string: T
   val inner_commentN: string val inner_comment: T
   val sortN: string val sort: T
   val typN: string val typ: T
@@ -106,6 +103,7 @@
   val debugN: string
   val initN: string
   val statusN: string
+  val no_output: output * output
   val default_output: T -> output * output
   val add_mode: string -> (T -> output * output) -> unit
   val output: T -> output * output
@@ -203,11 +201,9 @@
 val (skolemN, skolem) = markup_elem "skolem";
 val (boundN, bound) = markup_elem "bound";
 val (varN, var) = markup_elem "var";
-val (numN, num) = markup_elem "num";
-val (floatN, float) = markup_elem "float";
-val (xnumN, xnum) = markup_elem "xnum";
-val (xstrN, xstr) = markup_elem "xstr";
+val (numeralN, numeral) = markup_elem "numeral";
 val (literalN, literal) = markup_elem "literal";
+val (inner_stringN, inner_string) = markup_elem "inner_string";
 val (inner_commentN, inner_comment) = markup_elem "inner_comment";
 
 val (sortN, sort) = markup_elem "sort";
@@ -292,7 +288,8 @@
 
 (* print mode operations *)
 
-fun default_output (_: T) = ("", "");
+val no_output = ("", "");
+fun default_output (_: T) = no_output;
 
 local
   val default = {output = default_output};
@@ -304,7 +301,7 @@
     the_default default (Library.get_first (Symtab.lookup (! modes)) (print_mode_value ()));
 end;
 
-fun output m = if is_none m then ("", "") else #output (get_mode ()) m;
+fun output m = if is_none m then no_output else #output (get_mode ()) m;
 
 val enclose = output #-> Library.enclose;
 
--- a/src/Pure/General/markup.scala	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/markup.scala	Mon Jan 05 07:54:16 2009 -0800
@@ -25,6 +25,8 @@
   val FILE = "file"
   val ID = "id"
 
+  val POSITION_PROPERTIES = Set(LINE, COLUMN, OFFSET, END_LINE, END_COLUMN, END_OFFSET, FILE, ID)
+
   val POSITION = "position"
   val LOCATION = "location"
 
@@ -120,6 +122,8 @@
   val PID = "pid"
   val SESSION = "session"
 
+  val MESSAGE = "message"
+
 
   /* content */
 
--- a/src/Pure/General/name_space.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/name_space.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -59,10 +59,10 @@
 (** long identifiers **)
 
 fun hidden name = "??." ^ name;
-val is_hidden = String.isPrefix "??."
+val is_hidden = String.isPrefix "??.";
 
-val separator = ".";
-val is_qualified = exists_string (fn s => s = separator);
+val separator = Binding.separator;
+val is_qualified = Binding.is_qualified;
 
 val implode_name = space_implode separator;
 val explode_name = space_explode separator;
--- a/src/Pure/General/position.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/position.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,8 +1,7 @@
 (*  Title:      Pure/General/position.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
-Source positions: counting Isabelle symbols -- starting from 1.
+Source positions: counting Isabelle symbols, starting from 1.
 *)
 
 signature POSITION =
@@ -19,6 +18,7 @@
   val file: string -> T
   val line: int -> T
   val line_file: int -> string -> T
+  val id: string -> T
   val get_id: T -> string option
   val put_id: string -> T -> T
   val of_properties: Properties.T -> T
@@ -86,6 +86,7 @@
 val none = Pos ((0, 0, 0), []);
 val start = Pos ((1, 1, 1), []);
 
+
 fun file_name "" = []
   | file_name name = [(Markup.fileN, name)];
 
@@ -95,11 +96,14 @@
 fun line i = line_file i "";
 
 
-(* markup properties *)
+fun id id = Pos ((0, 0, 1), [(Markup.idN, id)]);
 
 fun get_id (Pos (_, props)) = Properties.get props Markup.idN;
 fun put_id id (Pos (count, props)) = Pos (count, Properties.put (Markup.idN, id) props);
 
+
+(* markup properties *)
+
 fun of_properties props =
   let
     fun get name =
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/swing.scala	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,18 @@
+/*  Title:      Pure/General/swing.scala
+    Author:     Makarius
+
+Swing utilities.
+*/
+
+package isabelle
+
+import javax.swing.SwingUtilities
+
+object Swing
+{
+  def now(body: => Unit) =
+    SwingUtilities.invokeAndWait(new Runnable { def run = body })
+
+  def later(body: => Unit) =
+    SwingUtilities.invokeLater(new Runnable { def run = body })
+}
--- a/src/Pure/General/symbol.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/symbol.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -65,6 +65,7 @@
   val bump_string: string -> string
   val length: symbol list -> int
   val xsymbolsN: string
+  val output: string -> output * int
 end;
 
 structure Symbol: SYMBOL =
@@ -175,21 +176,22 @@
   ord space <= ord c andalso ord c <= ord "~" andalso c <> "." andalso c <> ">"
   orelse ord c >= 128;
 
-fun encode_raw str =
-  let
-    val raw0 = enclose "\\<^raw:" ">";
-    val raw1 = raw0 o implode;
-    val raw2 = enclose "\\<^raw" ">" o string_of_int o ord;
-
-    fun encode cs = enc (Library.take_prefix raw_chr cs)
-    and enc ([], []) = []
-      | enc (cs, []) = [raw1 cs]
-      | enc ([], d :: ds) = raw2 d :: encode ds
-      | enc (cs, d :: ds) = raw1 cs :: raw2 d :: encode ds;
-  in
-    if exists_string (not o raw_chr) str then implode (encode (explode str))
-    else raw0 str
-  end;
+fun encode_raw "" = ""
+  | encode_raw str =
+      let
+        val raw0 = enclose "\\<^raw:" ">";
+        val raw1 = raw0 o implode;
+        val raw2 = enclose "\\<^raw" ">" o string_of_int o ord;
+    
+        fun encode cs = enc (Library.take_prefix raw_chr cs)
+        and enc ([], []) = []
+          | enc (cs, []) = [raw1 cs]
+          | enc ([], d :: ds) = raw2 d :: encode ds
+          | enc (cs, d :: ds) = raw1 cs :: raw2 d :: encode ds;
+      in
+        if exists_string (not o raw_chr) str then implode (encode (explode str))
+        else raw0 str
+      end;
 
 
 (* diagnostics *)
@@ -519,9 +521,9 @@
 
 
 
-(** xsymbols **)
+(** symbol output **)
 
-val xsymbolsN = "xsymbols";
+(* length *)
 
 fun sym_len s =
   if not (is_printable s) then (0: int)
@@ -532,8 +534,16 @@
 
 fun sym_length ss = fold (fn s => fn n => sym_len s + n) ss 0;
 
+
+(* print mode *)
+
+val xsymbolsN = "xsymbols";
+
+fun output s = (s, sym_length (sym_explode s));
+
+
 (*final declarations of this structure!*)
+val explode = sym_explode;
 val length = sym_length;
-val explode = sym_explode;
 
 end;
--- a/src/Pure/General/xml.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/xml.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -79,7 +79,7 @@
   end;
 
 fun output_markup (markup as (name, atts)) =
-  if Markup.is_none markup then ("", "")
+  if Markup.is_none markup then Markup.no_output
   else (enclose "<" ">" (elem name atts), enclose "</" ">" name);
 
 
--- a/src/Pure/General/xml.scala	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/xml.scala	Mon Jan 05 07:54:16 2009 -0800
@@ -10,16 +10,56 @@
 import javax.xml.parsers.DocumentBuilderFactory
 
 
-object XML {
+object XML
+{
   /* datatype representation */
 
   type Attributes = List[(String, String)]
 
-  abstract class Tree
+  abstract class Tree {
+    override def toString = {
+      val s = new StringBuilder
+      append_tree(this, s)
+      s.toString
+    }
+  }
   case class Elem(name: String, attributes: Attributes, body: List[Tree]) extends Tree
   case class Text(content: String) extends Tree
 
 
+  /* string representation */
+
+  private def append_text(text: String, s: StringBuilder) {
+    for (c <- text.elements) c match {
+      case '<' => s.append("&lt;")
+      case '>' => s.append("&gt;")
+      case '&' => s.append("&amp;")
+      case '"' => s.append("&quot;")
+      case '\'' => s.append("&apos;")
+      case _ => s.append(c)
+    }
+  }
+
+  private def append_elem(name: String, atts: Attributes, s: StringBuilder) {
+    s.append(name)
+    for ((a, x) <- atts) {
+      s.append(" "); s.append(a); s.append("=\""); append_text(x, s); s.append("\"")
+    }
+  }
+
+  private def append_tree(tree: Tree, s: StringBuilder) {
+    tree match {
+      case Elem(name, atts, Nil) =>
+        s.append("<"); append_elem(name, atts, s); s.append("/>")
+      case Elem(name, atts, ts) =>
+        s.append("<"); append_elem(name, atts, s); s.append(">")
+        for (t <- ts) append_tree(t, s)
+        s.append("</"); s.append(name); s.append(">")
+      case Text(text) => append_text(text, s)
+    }
+  }
+
+
   /* iterate over content */
 
   private type State = Option[(String, List[Tree])]
--- a/src/Pure/General/yxml.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/General/yxml.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -42,7 +42,7 @@
 (* output *)
 
 fun output_markup (markup as (name, atts)) =
-  if Markup.is_none markup then ("", "")
+  if Markup.is_none markup then Markup.no_output
   else (XY ^ name ^ implode (map (fn (a, x) => Y ^ a ^ "=" ^ x) atts) ^ X, XYX);
 
 fun element name atts body =
--- a/src/Pure/IsaMakefile	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/IsaMakefile	Mon Jan 05 07:54:16 2009 -0800
@@ -26,6 +26,7 @@
   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/markup.ML			\
   General/name_space.ML General/ord_list.ML General/output.ML		\
@@ -47,15 +48,15 @@
   Isar/proof_display.ML Isar/proof_node.ML Isar/rule_cases.ML		\
   Isar/rule_insts.ML Isar/session.ML Isar/skip_proof.ML			\
   Isar/spec_parse.ML Isar/specification.ML Isar/subclass.ML		\
-  Isar/theory_target.ML Isar/toplevel.ML ML-Systems/alice.ML		\
-  ML-Systems/exn.ML ML-Systems/install_pp_polyml.ML			\
-  ML-Systems/ml_name_space.ML ML-Systems/multithreading.ML		\
-  ML-Systems/mosml.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_common.ML ML-Systems/polyml.ML			\
-  ML-Systems/polyml_old_compiler4.ML					\
+  Isar/theory_target.ML Isar/toplevel.ML Isar/value_parse.ML		\
+  ML-Systems/alice.ML ML-Systems/exn.ML					\
+  ML-Systems/install_pp_polyml.ML ML-Systems/ml_name_space.ML		\
+  ML-Systems/multithreading.ML ML-Systems/mosml.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_common.ML			\
+  ML-Systems/polyml.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/time_limit.ML ML-Systems/thread_dummy.ML			\
@@ -73,17 +74,18 @@
   Syntax/printer.ML Syntax/simple_syntax.ML Syntax/syn_ext.ML		\
   Syntax/syn_trans.ML Syntax/syntax.ML Syntax/type_ext.ML Thy/html.ML	\
   Thy/latex.ML Thy/present.ML Thy/term_style.ML Thy/thm_deps.ML		\
-  Thy/thy_edit.ML Thy/thy_header.ML Thy/thy_info.ML Thy/thy_load.ML	\
-  Thy/thy_output.ML Tools/ROOT.ML Tools/invoke.ML			\
+  Thy/thy_header.ML Thy/thy_info.ML Thy/thy_load.ML Thy/thy_output.ML	\
+  Thy/thy_syntax.ML Tools/ROOT.ML Tools/invoke.ML			\
   Tools/isabelle_process.ML Tools/named_thms.ML Tools/xml_syntax.ML	\
   assumption.ML axclass.ML codegen.ML config.ML conjunction.ML		\
   consts.ML context.ML context_position.ML conv.ML defs.ML display.ML	\
   drule.ML envir.ML facts.ML goal.ML interpretation.ML library.ML	\
   logic.ML meta_simplifier.ML more_thm.ML morphism.ML name.ML net.ML	\
-  old_goals.ML pattern.ML primitive_defs.ML proofterm.ML pure_setup.ML	\
-  pure_thy.ML search.ML sign.ML simplifier.ML sorts.ML subgoal.ML	\
-  tactic.ML tctical.ML term.ML term_subst.ML theory.ML thm.ML type.ML	\
-  type_infer.ML unify.ML variable.ML ../Tools/quickcheck.ML
+  old_goals.ML old_term.ML pattern.ML primitive_defs.ML proofterm.ML	\
+  pure_setup.ML pure_thy.ML search.ML sign.ML simplifier.ML sorts.ML	\
+  subgoal.ML tactic.ML tctical.ML term.ML term_ord.ML term_subst.ML	\
+  theory.ML thm.ML type.ML type_infer.ML unify.ML variable.ML		\
+  ../Tools/quickcheck.ML
 	@./mk
 
 
@@ -121,11 +123,11 @@
 
 ## Scala material
 
-SCALA_FILES = General/event_bus.scala General/markup.scala	\
-  General/position.scala General/symbol.scala General/xml.scala	\
-  General/yxml.scala Isar/isar.scala Thy/thy_header.scala	\
-  Tools/isabelle_process.scala Tools/isabelle_syntax.scala	\
-  Tools/isabelle_system.scala
+SCALA_FILES = General/event_bus.scala General/markup.scala		\
+  General/position.scala General/swing.scala General/symbol.scala	\
+  General/xml.scala General/yxml.scala Isar/isar.scala			\
+  Thy/thy_header.scala Tools/isabelle_process.scala			\
+  Tools/isabelle_syntax.scala Tools/isabelle_system.scala
 
 
 SCALA_TARGET = $(ISABELLE_HOME)/lib/classes/Pure.jar
--- a/src/Pure/Isar/ROOT.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/ROOT.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/ROOT.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Isar -- Intelligible Semi-Automated Reasoning for Isabelle.
@@ -21,8 +20,9 @@
 
 (*outer syntax*)
 use "outer_lex.ML";
+use "outer_keyword.ML";
 use "outer_parse.ML";
-use "outer_keyword.ML";
+use "value_parse.ML";
 use "args.ML";
 use "antiquote.ML";
 use "../ML/ml_context.ML";
@@ -51,14 +51,13 @@
 use "obtain.ML";
 
 (*local theories and targets*)
-val new_locales = ref false;
 use "local_theory.ML";
 use "overloading.ML";
 use "locale.ML";
 use "new_locale.ML";
-use "expression.ML";
 use "class.ML";
 use "theory_target.ML";
+use "expression.ML";
 use "instance.ML";
 use "subclass.ML";
 
@@ -81,7 +80,7 @@
 (*theory syntax*)
 use "../Thy/term_style.ML";
 use "../Thy/thy_output.ML";
-use "../Thy/thy_edit.ML";
+use "../Thy/thy_syntax.ML";
 use "../old_goals.ML";
 use "outer_syntax.ML";
 use "../Thy/thy_info.ML";
--- a/src/Pure/Isar/class.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/class.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -85,7 +85,7 @@
 
 end;
 
-structure New_Locale =
+(*structure New_Locale =
 struct
 
 val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
@@ -106,7 +106,7 @@
 val parameters_of = NewLocale.params_of; (*why typ option?*)
 val add_locale = Expression.add_locale;
 
-end;
+end;*)
 
 structure Locale = Old_Locale;
 
--- a/src/Pure/Isar/code.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/code.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/code.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Abstract executable content of theory.  Management of data dependent on
@@ -16,8 +15,8 @@
   val del_eqn: thm -> theory -> theory
   val del_eqns: string -> theory -> theory
   val add_eqnl: string * (thm * bool) list lazy -> theory -> theory
-  val map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory
-  val map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory
+  val map_pre: (simpset -> simpset) -> theory -> theory
+  val map_post: (simpset -> simpset) -> theory -> theory
   val add_inline: thm -> theory -> theory
   val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory
   val del_functrans: string -> theory -> theory
@@ -186,8 +185,8 @@
 (* pre- and postprocessor *)
 
 datatype thmproc = Thmproc of {
-  pre: MetaSimplifier.simpset,
-  post: MetaSimplifier.simpset,
+  pre: simpset,
+  post: simpset,
   functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list
 };
 
@@ -198,8 +197,8 @@
 fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
   Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
     let
-      val pre = MetaSimplifier.merge_ss (pre1, pre2);
-      val post = MetaSimplifier.merge_ss (post1, post2);
+      val pre = Simplifier.merge_ss (pre1, pre2);
+      val post = Simplifier.merge_ss (post1, post2);
       val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2);
     in mk_thmproc ((pre, post), functrans) end;
 
@@ -221,7 +220,7 @@
     val thmproc = merge_thmproc (thmproc1, thmproc2);
     val spec = merge_spec (spec1, spec2);
   in mk_exec (thmproc, spec) end;
-val empty_exec = mk_exec (mk_thmproc ((MetaSimplifier.empty_ss, MetaSimplifier.empty_ss), []),
+val empty_exec = mk_exec (mk_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []),
   mk_spec ((false, Symtab.empty), (Symtab.empty, (Symtab.empty, Symtab.empty))));
 
 fun the_thmproc (Exec { thmproc = Thmproc x, ...}) = x;
@@ -417,12 +416,12 @@
       Pretty.block [
         Pretty.str "preprocessing simpset:",
         Pretty.fbrk,
-        MetaSimplifier.pretty_ss pre
+        Simplifier.pretty_ss pre
       ],
       Pretty.block [
         Pretty.str "postprocessing simpset:",
         Pretty.fbrk,
-        MetaSimplifier.pretty_ss post
+        Simplifier.pretty_ss post
       ],
       Pretty.block (
         Pretty.str "function transformers:"
--- a/src/Pure/Isar/code_unit.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/code_unit.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/code_unit.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Basic notions of code generation.  Auxiliary.
@@ -230,7 +229,7 @@
   val empty = ([], []);
   val copy = I;
   val extend = I;
-  fun merge _ ((alias1, classes1), (alias2, classes2)) =
+  fun merge _ ((alias1, classes1), (alias2, classes2)) : T =
     (Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2),
       Library.merge (op =) (classes1, classes2));
 );
@@ -286,7 +285,7 @@
       ^ " :: " ^ string_of_typ thy ty);
     fun last_typ c_ty ty =
       let
-        val frees = typ_tfrees ty;
+        val frees = OldTerm.typ_tfrees ty;
         val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty
           handle TYPE _ => no_constr c_ty
         val _ = if has_duplicates (eq_fst (op =)) vs then no_constr c_ty else ();
--- a/src/Pure/Isar/element.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/element.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -23,6 +23,10 @@
   val facts_map: (('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt) ->
    (Attrib.binding * ('fact * Attrib.src list) list) list ->
    (Attrib.binding * ('c * Attrib.src list) list) list
+  val map_ctxt': {binding: Binding.T -> Binding.T,
+    var: Binding.T * mixfix -> Binding.T * mixfix,
+    typ: 'typ -> 'a, term: 'term -> 'b, pat: 'term -> 'b, fact: 'fact -> 'c,
+    attrib: Attrib.src -> Attrib.src} -> ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
   val map_ctxt: {binding: Binding.T -> Binding.T,
     var: Binding.T * mixfix -> Binding.T * mixfix,
     typ: 'typ -> 'a, term: 'term -> 'b, fact: 'fact -> 'c,
@@ -74,10 +78,9 @@
   val generalize_facts: Proof.context -> Proof.context ->
     (Attrib.binding * (thm list * Attrib.src list) list) list ->
     (Attrib.binding * (thm list * Attrib.src list) list) list
-  val activate: (typ, term, Facts.ref) ctxt list -> Proof.context ->
-    (context_i list * (Binding.T * Thm.thm list) list) * Proof.context
-  val activate_i: context_i list -> Proof.context ->
-    (context_i list * (Binding.T * Thm.thm list) list) * Proof.context
+  val transfer_morphism: theory -> morphism
+  val activate: (typ, term, Facts.ref) ctxt list -> Proof.context -> context_i list * Proof.context
+  val activate_i: context_i list -> Proof.context -> context_i list * Proof.context
 end;
 
 structure Element: ELEMENT =
@@ -109,18 +112,22 @@
 
 fun facts_map f facts = Notes ("", facts) |> f |> (fn Notes (_, facts') => facts');
 
-fun map_ctxt {binding, var, typ, term, fact, attrib} =
+fun map_ctxt' {binding, var, typ, term, pat, fact, attrib} =
   fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) =>
        let val (x', mx') = var (x, mx) in (x', Option.map typ T, mx') end))
    | Constrains xs => Constrains (xs |> map (fn (x, T) =>
        let val x' = Binding.base_name (#1 (var (Binding.name x, NoSyn))) in (x', typ T) end))
    | Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) =>
-      ((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map term ps)))))
+      ((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map pat ps)))))
    | Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) =>
-      ((binding a, map attrib atts), (term t, map term ps))))
+      ((binding a, map attrib atts), (term t, map pat ps))))
    | Notes (kind, facts) => Notes (kind, facts |> map (fn ((a, atts), bs) =>
       ((binding a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts)))));
 
+fun map_ctxt {binding, var, typ, term, fact, attrib} =
+  map_ctxt' {binding = binding, var = var, typ = typ, term = term, pat = term,
+    fact = fact, attrib = attrib}
+
 fun map_ctxt_attrib attrib =
   map_ctxt {binding = I, var = I, typ = I, term = I, fact = I, attrib = attrib};
 
@@ -531,14 +538,23 @@
   in facts_map (morph_ctxt morphism) facts end;
 
 
+(* transfer to theory using closure *)
+
+fun transfer_morphism thy =
+  let val thy_ref = Theory.check_thy thy in
+    Morphism.morphism {binding = I, var = I, typ = I, term = I,
+      fact = map (fn th => transfer (Theory.deref thy_ref) th)}
+  end;
+
+
 (** activate in context, return elements and facts **)
 
 local
 
 fun activate_elem (Fixes fixes) ctxt =
-      ([], ctxt |> ProofContext.add_fixes_i fixes |> snd)
+      ctxt |> ProofContext.add_fixes_i fixes |> snd
   | activate_elem (Constrains _) ctxt =
-      ([], ctxt)
+      ctxt
   | activate_elem (Assumes asms) ctxt =
       let
         val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) asms;
@@ -546,7 +562,7 @@
         val (_, ctxt') =
           ctxt |> fold Variable.auto_fixes ts
           |> ProofContext.add_assms_i Assumption.presume_export asms';
-      in ([], ctxt') end
+      in ctxt' end
   | activate_elem (Defines defs) ctxt =
       let
         val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs;
@@ -556,19 +572,20 @@
         val (_, ctxt') =
           ctxt |> fold (Variable.auto_fixes o #1) asms
           |> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms);
-      in ([], ctxt') end
+      in ctxt' end
   | activate_elem (Notes (kind, facts)) ctxt =
       let
         val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts;
         val (res, ctxt') = ctxt |> ProofContext.note_thmss_i kind facts';
-      in ((map (#1 o #1) facts' ~~ map #2 res), ctxt') end;
+      in ctxt' end;
 
 fun gen_activate prep_facts raw_elems ctxt =
   let
-    val elems = map (map_ctxt_attrib Args.assignable o prep_facts ctxt) raw_elems;
-    val (res, ctxt') = fold_map activate_elem elems (ProofContext.qualified_names ctxt);
-    val elems' = elems |> map (map_ctxt_attrib Args.closure);
-  in ((elems', flat res), ProofContext.restore_naming ctxt ctxt') end;
+    fun activate elem ctxt =
+      let val elem' = (map_ctxt_attrib Args.assignable o prep_facts ctxt) elem
+      in (elem', activate_elem elem' ctxt) end
+    val (elems, ctxt') = fold_map activate raw_elems (ProofContext.qualified_names ctxt);
+  in (elems |> map (map_ctxt_attrib Args.closure), ProofContext.restore_naming ctxt ctxt') end;
 
 fun check_name name =
   if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name)
--- a/src/Pure/Isar/expression.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/expression.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -7,7 +7,7 @@
 signature EXPRESSION =
 sig
   datatype 'term map = Positional of 'term option list | Named of (string * 'term) list;
-  type 'term expr = (string * (string * 'term map)) list;
+  type 'term expr = (string * ((string * bool) * 'term map)) list;
   type expression = string expr * (Binding.T * string option * mixfix) list;
   type expression_i = term expr * (Binding.T * typ option * mixfix) list;
 
@@ -28,8 +28,8 @@
   (* Interpretation *)
   val sublocale_cmd: string -> expression -> theory -> Proof.state;
   val sublocale: string -> expression_i -> theory -> Proof.state;
-  val interpretation_cmd: expression -> theory -> Proof.state;
-  val interpretation: expression_i -> theory -> Proof.state;
+  val interpretation_cmd: expression -> (Attrib.binding * string) list -> theory -> Proof.state;
+  val interpretation: expression_i -> (Attrib.binding * term) list -> theory -> Proof.state;
   val interpret_cmd: expression -> bool -> Proof.state -> Proof.state;
   val interpret: expression_i -> bool -> Proof.state -> Proof.state;
 end;
@@ -47,7 +47,7 @@
   Positional of 'term option list |
   Named of (string * 'term) list;
 
-type 'term expr = (string * (string * 'term map)) list;
+type 'term expr = (string * ((string * bool) * 'term map)) list;
 
 type expression = string expr * (Binding.T * string option * mixfix) list;
 type expression_i = term expr * (Binding.T * typ option * mixfix) list;
@@ -74,7 +74,7 @@
       end;
 
     fun match_bind (n, b) = (n = Binding.base_name b);
-    fun parm_eq ((b1, mx1), (b2, mx2)) =
+    fun parm_eq ((b1, mx1: mixfix), (b2, mx2)) =
       (* FIXME: cannot compare bindings for equality, instead check for equal name and syntax *)
       (Binding.base_name b1 = Binding.base_name b2) andalso
       (if mx1 = mx2 then true
@@ -154,7 +154,7 @@
 
 (* Instantiation morphism *)
 
-fun inst_morph (parm_names, parm_types) (prfx, insts') ctxt =
+fun inst_morph (parm_names, parm_types) ((prfx, strict), insts') ctxt =
   let
     (* parameters *)
     val type_parm_names = fold Term.add_tfreesT parm_types [] |> map fst;
@@ -175,7 +175,7 @@
     val inst = Symtab.make insts'';
   in
     (Element.inst_morphism (ProofContext.theory_of ctxt) (instT, inst) $>
-      Morphism.binding_morphism (Binding.qualify prfx), ctxt')
+      Morphism.binding_morphism (Binding.add_prefix strict prfx), ctxt')
   end;
 
 
@@ -184,12 +184,15 @@
 (** Parsing **)
 
 fun parse_elem prep_typ prep_term ctxt elem =
-  Element.map_ctxt {binding = I, var = I, typ = prep_typ ctxt,
-    term = prep_term ctxt, fact = I, attrib = I} elem;
+  Element.map_ctxt' {binding = I, var = I, typ = prep_typ ctxt,
+  term = prep_term (ProofContext.set_mode ProofContext.mode_schematic ctxt), (* FIXME ?? *)
+  pat = prep_term (ProofContext.set_mode ProofContext.mode_pattern ctxt),
+  fact = I, attrib = I} elem;
 
 fun parse_concl prep_term ctxt concl =
   (map o map) (fn (t, ps) =>
-    (prep_term ctxt, map (prep_term ctxt) ps)) concl;
+    (prep_term (ProofContext.set_mode ProofContext.mode_schematic ctxt) t, (* FIXME ?? *)
+      map (prep_term (ProofContext.set_mode ProofContext.mode_pattern ctxt)) ps)) concl;
 
 
 (** Simultaneous type inference: instantiations + elements + conclusion **)
@@ -268,38 +271,7 @@
   | declare_elem _ (Defines _) ctxt = ctxt
   | declare_elem _ (Notes _) ctxt = ctxt;
 
-(** Finish locale elements, extract specification text **)
-
-val norm_term = Envir.beta_norm oo Term.subst_atomic;
-
-fun bind_def ctxt eq (xs, env, eqs) =
-  let
-    val _ = LocalDefs.cert_def ctxt eq;
-    val ((y, T), b) = LocalDefs.abs_def eq;
-    val b' = norm_term env b;
-    fun err msg = error (msg ^ ": " ^ quote y);
-  in
-    exists (fn (x, _) => x = y) xs andalso
-      err "Attempt to define previously specified variable";
-    exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso
-      err "Attempt to redefine variable";
-    (Term.add_frees b' xs, (Free (y, T), b') :: env, eq :: eqs)
-  end;
-
-fun eval_text _ _ (Fixes _) text = text
-  | eval_text _ _ (Constrains _) text = text
-  | eval_text _ is_ext (Assumes asms)
-        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
-      let
-        val ts = maps (map #1 o #2) asms;
-        val ts' = map (norm_term env) ts;
-        val spec' =
-          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
-          else ((exts, exts'), (ints @ ts, ints' @ ts'));
-      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
-  | eval_text ctxt _ (Defines defs) (spec, binds) =
-      (spec, fold (bind_def ctxt o #1 o #2) defs binds)
-  | eval_text _ _ (Notes _) text = text;
+(** Finish locale elements **)
 
 fun closeup _ _ false elem = elem
   | closeup ctxt parms true elem =
@@ -334,36 +306,24 @@
   | finish_primitive _ close (Defines defs) = close (Defines defs)
   | finish_primitive _ _ (Notes facts) = Notes facts;
 
-fun finish_inst ctxt parms do_close (loc, (prfx, inst)) text =
+fun finish_inst ctxt parms do_close (loc, (prfx, inst)) =
   let
     val thy = ProofContext.theory_of ctxt;
     val (parm_names, parm_types) = NewLocale.params_of thy loc |>
       map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
-    val (asm, defs) = NewLocale.specification_of thy loc;
     val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
-    val asm' = Option.map (Morphism.term morph) asm;
-    val defs' = map (Morphism.term morph) defs;
-    val text' = text |>
-      (if is_some asm
-        then eval_text ctxt false (Assumes [(Attrib.empty_binding, [(the asm', [])])])
-        else I) |>
-      (if not (null defs)
-        then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
-        else I)
-(* FIXME clone from new_locale.ML *)
-  in ((loc, morph), text') end;
+  in (loc, morph) end;
 
-fun finish_elem ctxt parms do_close elem text =
+fun finish_elem ctxt parms do_close elem =
   let
     val elem' = finish_primitive parms (closeup ctxt parms do_close) elem;
-    val text' = eval_text ctxt true elem' text;
-  in (elem', text') end
+  in elem' end
   
-fun finish ctxt parms do_close insts elems text =
+fun finish ctxt parms do_close insts elems =
   let
-    val (deps, text') = fold_map (finish_inst ctxt parms do_close) insts text;
-    val (elems', text'') = fold_map (finish_elem ctxt parms do_close) elems text';
-  in (deps, elems', text'') end;
+    val deps = map (finish_inst ctxt parms do_close) insts;
+    val elems' = map (finish_elem ctxt parms do_close) elems;
+  in (deps, elems') end;
 
 
 (** Process full context statement: instantiations + elements + conclusion **)
@@ -398,19 +358,17 @@
       let
         val ctxt' = declare_elem prep_vars raw_elem ctxt;
         val elems' = elems @ [parse_elem parse_typ parse_prop ctxt' raw_elem];
-        (* FIXME term bindings *)
         val (_, _, _, ctxt'') = check_autofix insts elems' [] ctxt';
       in (insts, elems', ctxt') end;
 
     fun prep_concl raw_concl (insts, elems, ctxt) =
       let
-        val concl = (map o map) (fn (t, ps) =>
-          (parse_prop ctxt t, map (parse_prop ctxt) ps)) raw_concl;
+        val concl = parse_concl parse_prop ctxt raw_concl;
       in check_autofix insts elems concl ctxt end;
 
     val fors = prep_vars fixed context |> fst;
     val ctxt = context |> ProofContext.add_fixes_i fors |> snd;
-    val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], NewLocale.clear_local_idents ctxt);
+    val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], ctxt);
     val (_, elems'', ctxt'') = fold prep_elem raw_elems (insts', [], ctxt');
     val (insts, elems, concl, ctxt) = prep_concl raw_concl (insts', elems'', ctxt'');
 
@@ -421,28 +379,9 @@
     val parms = xs ~~ Ts;  (* params from expression and elements *)
 
     val Fixes fors' = finish_primitive parms I (Fixes fors);
-    val (deps, elems', text) = finish ctxt' parms do_close insts elems ((([], []), ([], [])), ([], [], []));
-    (* text has the following structure:
-           (((exts, exts'), (ints, ints')), (xs, env, defs))
-       where
-         exts: external assumptions (terms in assumes elements)
-         exts': dito, normalised wrt. env
-         ints: internal assumptions (terms in assumptions from insts)
-         ints': dito, normalised wrt. env
-         xs: the free variables in exts' and ints' and rhss of definitions,
-           this includes parameters except defined parameters
-         env: list of term pairs encoding substitutions, where the first term
-           is a free variable; substitutions represent defines elements and
-           the rhs is normalised wrt. the previous env
-         defs: the equations from the defines elements
-       elems is an updated version of raw_elems:
-         - type info added to Fixes and modified in Constrains
-         - axiom and definition statement replaced by corresponding one
-           from proppss in Assumes and Defines
-         - Facts unchanged
-       *)
+    val (deps, elems') = finish ctxt' parms do_close insts elems;
 
-  in ((fors', deps, elems', concl), (parms, text)) end
+  in ((fors', deps, elems', concl), (parms, ctxt')) end
 
 in
 
@@ -479,14 +418,13 @@
 
 fun prep_declaration prep activate raw_import raw_elems context =
   let
-    val ((fixed, deps, elems, _), (parms, (spec, (_, _, defs)))) =
-      prep false true context raw_import raw_elems [];
+    val ((fixed, deps, elems, _), (parms, ctxt')) = prep false true context raw_import raw_elems [];
     (* Declare parameters and imported facts *)
     val context' = context |>
       ProofContext.add_fixes_i fixed |> snd |>
-      NewLocale.clear_local_idents |> fold NewLocale.activate_local_facts deps;
-    val ((elems', _), _) = activate elems (ProofContext.set_stmt true context');
-  in ((fixed, deps, elems'), (parms, spec, defs)) end;
+      fold NewLocale.activate_local_facts deps;
+    val (elems', _) = activate elems (ProofContext.set_stmt true context');
+  in ((fixed, deps, elems'), (parms, ctxt')) end;
 
 in
 
@@ -521,7 +459,7 @@
 
     val export = Variable.export_morphism goal_ctxt context;
     val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export;
-    val exp_term = Drule.term_rule thy (singleton exp_fact);
+    val exp_term = TermSubst.zero_var_indexes o Morphism.term export;
     val exp_typ = Logic.type_map exp_term;
     val export' =
       Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
@@ -537,6 +475,81 @@
 
 (*** Locale declarations ***)
 
+(* extract specification text *)
+
+val norm_term = Envir.beta_norm oo Term.subst_atomic;
+
+fun bind_def ctxt eq (xs, env, eqs) =
+  let
+    val _ = LocalDefs.cert_def ctxt eq;
+    val ((y, T), b) = LocalDefs.abs_def eq;
+    val b' = norm_term env b;
+    fun err msg = error (msg ^ ": " ^ quote y);
+  in
+    case filter (fn (Free (y', _), _) => y = y' | _ => false) env of
+      [] => (Term.add_frees b' xs, (Free (y, T), b') :: env, eq :: eqs) |
+      dups => if forall (fn (_, b'') => b' aconv b'') dups
+        then (xs, env, eqs)
+        else err "Attempt to redefine variable"
+  end;
+
+(* text has the following structure:
+       (((exts, exts'), (ints, ints')), (xs, env, defs))
+   where
+     exts: external assumptions (terms in assumes elements)
+     exts': dito, normalised wrt. env
+     ints: internal assumptions (terms in assumptions from insts)
+     ints': dito, normalised wrt. env
+     xs: the free variables in exts' and ints' and rhss of definitions,
+       this includes parameters except defined parameters
+     env: list of term pairs encoding substitutions, where the first term
+       is a free variable; substitutions represent defines elements and
+       the rhs is normalised wrt. the previous env
+     defs: the equations from the defines elements
+   *)
+
+fun eval_text _ _ (Fixes _) text = text
+  | eval_text _ _ (Constrains _) text = text
+  | eval_text _ is_ext (Assumes asms)
+        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
+      let
+        val ts = maps (map #1 o #2) asms;
+        val ts' = map (norm_term env) ts;
+        val spec' =
+          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
+          else ((exts, exts'), (ints @ ts, ints' @ ts'));
+      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
+  | eval_text ctxt _ (Defines defs) (spec, binds) =
+      (spec, fold (bind_def ctxt o #1 o #2) defs binds)
+  | eval_text _ _ (Notes _) text = text;
+
+fun eval_inst ctxt (loc, morph) text =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val (asm, defs) = NewLocale.specification_of thy loc;
+    val asm' = Option.map (Morphism.term morph) asm;
+    val defs' = map (Morphism.term morph) defs;
+    val text' = text |>
+      (if is_some asm
+        then eval_text ctxt false (Assumes [(Attrib.empty_binding, [(the asm', [])])])
+        else I) |>
+      (if not (null defs)
+        then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
+        else I)
+(* FIXME clone from new_locale.ML *)
+  in text' end;
+
+fun eval_elem ctxt elem text =
+  let
+    val text' = eval_text ctxt true elem text;
+  in text' end;
+
+fun eval ctxt deps elems =
+  let
+    val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [], []));
+    val ((spec, (_, _, defs))) = fold (eval_elem ctxt) elems text';
+  in (spec, defs) end;
+
 (* axiomsN: name of theorem set with destruct rules for locale predicates,
      also name suffix of delta predicates and assumptions. *)
 
@@ -581,10 +594,11 @@
     val name = Sign.full_bname thy bname;
 
     val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
-    val env = Term.add_term_free_names (body, []);
+    val env = Term.add_free_names body [];
     val xs = filter (member (op =) env o #1) parms;
     val Ts = map #2 xs;
-    val extraTs = (Term.term_tfrees body \\ fold Term.add_tfreesT Ts [])
+    val extraTs =
+      (Term.add_tfrees body [] \\ fold Term.add_tfreesT Ts [])
       |> Library.sort_wrt #1 |> map TFree;
     val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
 
@@ -622,7 +636,7 @@
 
 (* CB: main predicate definition function *)
 
-fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) thy =
+fun define_preds pname parms (((exts, exts'), (ints, ints')), defs) thy =
   let
     val defs' = map (cterm_of thy #> Assumption.assume #> Drule.gen_all #> Drule.abs_def) defs;
 
@@ -675,7 +689,8 @@
 
 fun defines_to_notes thy (Defines defs) =
       Notes (Thm.definitionK, map (fn (a, (def, _)) =>
-        (a, [([Assumption.assume (cterm_of thy def)], [])])) defs)
+        (a, [([Assumption.assume (cterm_of thy def)],
+          [(Attrib.internal o K) NewLocale.witness_attrib])])) defs)
   | defines_to_notes _ e = e;
 
 fun gen_add_locale prep_decl
@@ -685,10 +700,12 @@
     val _ = NewLocale.test_locale thy name andalso
       error ("Duplicate definition of locale " ^ quote name);
 
-    val ((fixed, deps, body_elems), text as (parms, ((_, exts'), _), defs)) =
+    val ((fixed, deps, body_elems), (parms, ctxt')) =
       prep_decl raw_imprt raw_body (ProofContext.init thy);
+    val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems;
+
     val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') =
-      define_preds predicate_name text thy;
+      define_preds predicate_name parms text thy;
 
     val extraTs = fold Term.add_tfrees exts' [] \\ fold Term.add_tfreesT (map snd parms) [];
     val _ = if null extraTs then ()
@@ -742,8 +759,8 @@
 
 fun prep_propp propss = propss |> map (map (rpair [] o Element.mark_witness));
 
-fun prep_result propps thmss =
-  ListPair.map (fn (props, thms) => map2 Element.make_witness props thms) (propps, thmss);
+val prep_result = map2 (fn props => fn thms =>
+  map2 Element.make_witness props (map Thm.close_derivation thms));
 
 
 (** Interpretation between locales: declaring sublocale relationships **)
@@ -786,33 +803,63 @@
 
 local
 
-fun gen_interpretation prep_expr
-    expression thy =
+datatype goal = Reg of string * Morphism.morphism | Eqns of Attrib.binding list;
+
+fun gen_interpretation prep_expr parse_prop prep_attr
+    expression equations thy =
   let
     val ctxt = ProofContext.init thy;
 
-    val ((propss, regs, export), goal_ctxt) = prep_expr expression ctxt;
+    val ((propss, regs, export), expr_ctxt) = prep_expr expression ctxt;
     
-    fun store_reg ((name, morph), thms) =
-      let
-        val morph' = morph $> Element.satisfy_morphism thms $> export;
-      in
-        NewLocale.add_global_registration (name, morph') #>
-        NewLocale.activate_global_facts (name, morph')
-      end;
+    val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt;
+    val eq_attns = map ((apsnd o map) (prep_attr thy) o fst) equations;
+    val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
+    val export' = Variable.export_morphism goal_ctxt expr_ctxt;
+
+    fun store (Reg (name, morph), thms) (regs, thy) =
+        let
+          val thms' = map (Element.morph_witness export') thms;
+          val morph' = morph $> Element.satisfy_morphism thms';
+          val add = NewLocale.add_global_registration (name, (morph', export));
+        in ((name, morph') :: regs, add thy) end
+      | store (Eqns [], []) (regs, thy) =
+        let val add = fold_rev (fn (name, morph) =>
+              NewLocale.activate_global_facts (name, morph $> export)) regs;
+        in (regs, add thy) end
+      | store (Eqns attns, thms) (regs, thy) =
+        let
+          val thms' = thms |> map (Element.conclude_witness #>
+            Morphism.thm (export' $> export) #>
+            LocalDefs.meta_rewrite_rule (ProofContext.init thy) #>
+            Drule.abs_def);
+          val eq_morph =
+            Morphism.term_morphism (MetaSimplifier.rewrite_term thy thms' []) $>
+            Morphism.thm_morphism (MetaSimplifier.rewrite_rule thms');
+          val attns' = map ((apsnd o map) (Attrib.attribute_i thy)) attns;
+          val add =
+            fold_rev (fn (name, morph) =>
+              NewLocale.amend_global_registration eq_morph (name, morph) #>
+              NewLocale.activate_global_facts (name, morph $> eq_morph $> export)) regs #>
+            PureThy.note_thmss Thm.lemmaK (attns' ~~ map (fn th => [([th], [])]) thms') #>
+            snd
+        in (regs, add thy) end;
 
     fun after_qed results =
-      ProofContext.theory (fold store_reg (regs ~~ prep_result propss results));
+      ProofContext.theory (fn thy =>
+        fold store (map Reg regs @ [Eqns eq_attns] ~~
+          prep_result (propss @ [eqns]) results) ([], thy) |> snd);
   in
     goal_ctxt |>
-      Proof.theorem_i NONE after_qed (prep_propp propss) |>
+      Proof.theorem_i NONE after_qed (prep_propp (propss @ [eqns])) |>
       Element.refine_witness |> Seq.hd
   end;
 
 in
 
-fun interpretation_cmd x = gen_interpretation read_goal_expression x;
-fun interpretation x = gen_interpretation cert_goal_expression x;
+fun interpretation_cmd x = gen_interpretation read_goal_expression
+  Syntax.parse_prop Attrib.intern_src x;
+fun interpretation x = gen_interpretation cert_goal_expression (K I) (K I) x;
 
 end;
 
--- a/src/Pure/Isar/find_theorems.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/find_theorems.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/find_theorems.ML
-    ID:         $Id$
     Author:     Rafal Kolanski and Gerwin Klein, NICTA
 
 Retrieve theorems from proof context.
@@ -166,7 +165,7 @@
 fun filter_simp ctxt t (_, thm) =
   let
     val (_, {mk_rews = {mk, ...}, ...}) =
-      MetaSimplifier.rep_ss (Simplifier.local_simpset_of ctxt);
+      Simplifier.rep_ss (Simplifier.local_simpset_of ctxt);
     val extract_simp =
       (map Thm.full_prop_of o mk, #1 o Logic.dest_equals o Logic.strip_imp_concl);
     val ss = is_matching_thm extract_simp ctxt false t thm
@@ -177,9 +176,9 @@
 
 (* filter_pattern *)
 
-fun get_names (_, thm) = let
-    val t = Thm.full_prop_of thm;
-  in (term_consts t) union (add_term_free_names (t, [])) end;
+fun get_names (_, thm) =
+  fold_aterms (fn Const (c, _) => insert (op =) c | Free (x, _) => insert (op =) x | _ => I)
+    (Thm.full_prop_of thm) [];
 
 fun add_pat_names (t, cs) =
       case strip_comb t of
@@ -287,7 +286,7 @@
 
 fun rem_thm_dups xs =
   xs ~~ (1 upto length xs)
-  |> sort (Term.fast_term_ord o pairself (Thm.prop_of o #2 o #1))
+  |> sort (TermOrd.fast_term_ord o pairself (Thm.prop_of o #2 o #1))
   |> rem_cdups
   |> sort (int_ord o pairself #2)
   |> map #1;
--- a/src/Pure/Isar/isar.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/isar.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/isar.ML
-    ID:         $Id$
     Author:     Makarius
 
 The global Isabelle/Isar state and main read-eval-print loop.
@@ -7,15 +6,12 @@
 
 signature ISAR =
 sig
-  type id = string
-  val no_id: id
-  val create_command: Toplevel.transition -> id
-  val init_point: unit -> unit
+  val init: unit -> unit
+  val exn: unit -> (exn * string) option
   val state: unit -> Toplevel.state
   val context: unit -> Proof.context
   val goal: unit -> thm
   val print: unit -> unit
-  val exn: unit -> (exn * string) option
   val >> : Toplevel.transition -> bool
   val >>> : Toplevel.transition list -> unit
   val linear_undo: int -> unit
@@ -26,6 +22,10 @@
   val toplevel_loop: {init: bool, welcome: bool, sync: bool, secure: bool} -> unit
   val loop: unit -> unit
   val main: unit -> unit
+
+  type id = string
+  val no_id: id
+  val create_command: Toplevel.transition -> id
   val insert_command: id -> id -> unit
   val remove_command: id -> unit
 end;
@@ -34,6 +34,130 @@
 struct
 
 
+(** TTY model -- SINGLE-THREADED! **)
+
+(* the global state *)
+
+type history = (Toplevel.state * Toplevel.transition) list;
+  (*previous state, state transition -- regular commands only*)
+
+local
+  val global_history = ref ([]: history);
+  val global_state = ref Toplevel.toplevel;
+  val global_exn = ref (NONE: (exn * string) option);
+in
+
+fun edit_history count f = NAMED_CRITICAL "Isar" (fn () =>
+  let
+    fun edit 0 (st, hist) = (global_history := hist; global_state := st; global_exn := NONE)
+      | edit n (st, hist) = edit (n - 1) (f st hist);
+  in edit count (! global_state, ! global_history) end);
+
+fun state () = NAMED_CRITICAL "Isar" (fn () => ! global_state);
+fun set_state state = NAMED_CRITICAL "Isar" (fn () => global_state := state);
+
+fun exn () = NAMED_CRITICAL "Isar" (fn () => ! global_exn);
+fun set_exn exn =  NAMED_CRITICAL "Isar" (fn () => global_exn := exn);
+
+end;
+
+
+fun init () = edit_history 1 (K (K (Toplevel.toplevel, [])));
+
+fun context () = Toplevel.context_of (state ())
+  handle Toplevel.UNDEF => error "Unknown context";
+
+fun goal () = #2 (#2 (Proof.get_goal (Toplevel.proof_of (state ()))))
+  handle Toplevel.UNDEF => error "No goal present";
+
+fun print () = Toplevel.print_state false (state ());
+
+
+(* history navigation *)
+
+local
+
+fun find_and_undo _ [] = error "Undo history exhausted"
+  | find_and_undo which ((prev, tr) :: hist) =
+      ((case Toplevel.init_of tr of SOME name => ThyInfo.kill_thy name | NONE => ());
+        if which (Toplevel.name_of tr) then (prev, hist) else find_and_undo which hist);
+
+in
+
+fun linear_undo n = edit_history n (K (find_and_undo (K true)));
+
+fun undo n = edit_history n (fn st => fn hist =>
+  find_and_undo (if Toplevel.is_proof st then K true else OuterKeyword.is_theory) hist);
+
+fun kill () = edit_history 1 (fn st => fn hist =>
+  find_and_undo
+    (if Toplevel.is_proof st then OuterKeyword.is_theory else OuterKeyword.is_theory_begin) hist);
+
+fun kill_proof () = edit_history 1 (fn st => fn hist =>
+  if Toplevel.is_proof st then find_and_undo OuterKeyword.is_theory hist
+  else raise Toplevel.UNDEF);
+
+end;
+
+
+(* interactive state transformations *)
+
+fun op >> tr =
+  (case Toplevel.transition true tr (state ()) of
+    NONE => false
+  | SOME (_, SOME err) => (set_exn (SOME err); Toplevel.error_msg tr err; true)
+  | SOME (st', NONE) =>
+      let
+        val name = Toplevel.name_of tr;
+        val _ = if OuterKeyword.is_theory_begin name then init () else ();
+        val _ =
+          if OuterKeyword.is_regular name
+          then edit_history 1 (fn st => fn hist => (st', (st, tr) :: hist)) else ();
+      in true end);
+
+fun op >>> [] = ()
+  | op >>> (tr :: trs) = if op >> tr then op >>> trs else ();
+
+
+(* toplevel loop *)
+
+val crashes = ref ([]: exn list);
+
+local
+
+fun raw_loop secure src =
+  let
+    fun check_secure () =
+      (if secure then warning "Secure loop -- cannot exit to ML" else (); secure);
+  in
+    (case Source.get_single (Source.set_prompt Source.default_prompt src) of
+      NONE => if secure then quit () else ()
+    | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ())
+    handle exn => (Output.error_msg (Toplevel.exn_message exn)
+    handle crash =>
+      (CRITICAL (fn () => change crashes (cons crash));
+        warning "Recovering after Isar toplevel crash -- see also Isar.crashes");
+      raw_loop secure src)
+  end;
+
+in
+
+fun toplevel_loop {init = do_init, welcome, sync, secure} =
+ (Context.set_thread_data NONE;
+  if do_init then init () else ();  (* FIXME init editor model *)
+  if welcome then writeln (Session.welcome ()) else ();
+  uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
+
+end;
+
+fun loop () =
+  toplevel_loop {init = false, welcome = false, sync = false, secure = Secure.is_secure ()};
+
+fun main () =
+  toplevel_loop {init = true, welcome = true, sync = false, secure = Secure.is_secure ()};
+
+
+
 (** individual toplevel commands **)
 
 (* unique identification *)
@@ -41,15 +165,6 @@
 type id = string;
 val no_id : id = "";
 
-fun identify tr =
-  (case Toplevel.get_id tr of
-    SOME id => (id, tr)
-  | NONE =>
-      let val id =
-        if ! Toplevel.debug then "isabelle:" ^ Toplevel.name_of tr ^ serial_string ()
-        else "isabelle:" ^ serial_string ()
-      in (id, Toplevel.put_id id tr) end);
-
 
 (* command category *)
 
@@ -168,7 +283,15 @@
 
 fun create_command raw_tr =
   let
-    val (id, tr) = identify raw_tr;
+    val (id, tr) =
+      (case Toplevel.get_id raw_tr of
+        SOME id => (id, raw_tr)
+      | NONE =>
+          let val id =
+            if ! Toplevel.debug then "isabelle:" ^ Toplevel.name_of raw_tr ^ serial_string ()
+            else "isabelle:" ^ serial_string ()
+          in (id, Toplevel.put_id id raw_tr) end);
+
     val cmd = make_command (category_of tr, tr, Unprocessed);
     val _ = change_commands (Graph.new_node (id, cmd));
   in id end;
@@ -190,153 +313,6 @@
 
 
 
-(** TTY model -- single-threaded **)
-
-(* global point *)
-
-local val global_point = ref no_id in
-
-fun change_point f = NAMED_CRITICAL "Isar" (fn () => change global_point f);
-fun point () = NAMED_CRITICAL "Isar" (fn () => ! global_point);
-
-end;
-
-
-fun set_point id = change_point (K id);
-fun init_point () = set_point no_id;
-
-fun point_state () = NAMED_CRITICAL "Isar" (fn () =>
-  let val id = point () in (id, the_state id) end);
-
-fun state () = #2 (point_state ());
-
-fun context () =
-  Toplevel.context_of (state ())
-    handle Toplevel.UNDEF => error "Unknown context";
-
-fun goal () =
-  #2 (#2 (Proof.get_goal (Toplevel.proof_of (state ()))))
-    handle Toplevel.UNDEF => error "No goal present";
-
-fun print () = Toplevel.print_state false (state ());
-
-
-(* global failure status *)
-
-local val global_exn = ref (NONE: (exn * string) option) in
-
-fun set_exn err = global_exn := err;
-fun exn () = ! global_exn;
-
-end;
-
-
-(* interactive state transformations *)
-
-fun op >> raw_tr =
-  let
-    val id = create_command raw_tr;
-    val {category, transition = tr, ...} = the_command id;
-    val (prev, prev_state) = point_state ();
-    val _ =
-      if is_regular category
-      then (dispose_commands (next_commands prev); change_commands (add_edge (prev, id)))
-      else ();
-  in
-    (case run true tr prev_state of
-      NONE => false
-    | SOME (status as Failed err) => (update_status status id; set_exn (SOME err); true)
-    | SOME status =>
-       (update_status status id; set_exn NONE;
-        if is_regular category then set_point id else ();
-        true))
-  end;
-
-fun op >>> [] = ()
-  | op >>> (tr :: trs) = if op >> tr then op >>> trs else ();
-
-
-(* implicit navigation wrt. proper commands *)
-
-local
-
-fun err_undo () = error "Undo history exhausted";
-
-fun find_category which id =
-  (case #category (the_command id) of
-    Empty => err_undo ()
-  | category => if which category then id else find_category which (prev_command id));
-
-fun find_begin_theory id =
-  if id = no_id then err_undo ()
-  else if is_some (Toplevel.init_of (#transition (the_command id))) then id
-  else find_begin_theory (prev_command id);
-
-fun undo_command id =
-  (case Toplevel.init_of (#transition (the_command id)) of
-    SOME name => prev_command id before ThyInfo.kill_thy name
-  | NONE => prev_command id);
-
-in
-
-fun linear_undo n = change_point (funpow n (fn id => undo_command (find_category is_proper id)));
-
-fun undo n = change_point (funpow n (fn id => undo_command
-  (find_category (if Toplevel.is_proof (the_state id) then is_proper else is_theory) id)));
-
-fun kill () = change_point (fn id => undo_command
-  (if Toplevel.is_proof (the_state id) then find_category is_theory id else find_begin_theory id));
-
-fun kill_proof () = change_point (fn id =>
-  if Toplevel.is_proof (the_state id) then undo_command (find_category is_theory id)
-  else raise Toplevel.UNDEF);
-
-end;
-
-
-(* toplevel loop *)
-
-val crashes = ref ([]: exn list);
-
-local
-
-fun raw_loop secure src =
-  let
-    fun check_secure () =
-      (if secure then warning "Secure loop -- cannot exit to ML" else (); secure);
-    val prev = point ();
-    val prev_name = Toplevel.name_of (#transition (the_command prev));
-    val prompt_markup =
-      prev <> no_id ? Markup.markup
-        (Markup.properties [(Markup.idN, prev), (Markup.nameN, prev_name)] Markup.prompt);
-  in
-    (case Source.get_single (Source.set_prompt (prompt_markup Source.default_prompt) src) of
-      NONE => if secure then quit () else ()
-    | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ())
-    handle exn => (Output.error_msg (Toplevel.exn_message exn) handle crash =>
-      (CRITICAL (fn () => change crashes (cons crash));
-        warning "Recovering after Isar toplevel crash -- see also Isar.crashes");
-      raw_loop secure src)
-  end;
-
-in
-
-fun toplevel_loop {init, welcome, sync, secure} =
- (Context.set_thread_data NONE;
-  if init then (init_point (); init_commands ()) else ();
-  if welcome then writeln (Session.welcome ()) else ();
-  uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
-
-end;
-
-fun loop () =
-  toplevel_loop {init = false, welcome = false, sync = false, secure = Secure.is_secure ()};
-
-fun main () =
-  toplevel_loop {init = true, welcome = true, sync = false, secure = Secure.is_secure ()};
-
-
-
 (** editor model **)
 
 (* run commands *)
@@ -377,26 +353,25 @@
 
 local
 
-structure P = OuterParse and K = OuterKeyword;
+structure P = OuterParse;
 val op >> = Scan.>>;
 
 in
 
 val _ =
-  OuterSyntax.improper_command "Isar.command" "define command (Isar editor model)" K.control
-    (P.props_text >> (fn (pos, str) =>
-      Toplevel.no_timing o Toplevel.imperative (fn () =>
-        ignore (create_command (OuterSyntax.prepare_command pos str)))));
+  OuterSyntax.internal_command "Isar.command"
+    (P.string -- P.string >> (fn (id, text) =>
+      Toplevel.imperative (fn () =>
+        ignore (create_command (OuterSyntax.prepare_command (Position.id id) text)))));
 
 val _ =
-  OuterSyntax.improper_command "Isar.insert" "insert command (Isar editor model)" K.control
+  OuterSyntax.internal_command "Isar.insert"
     (P.string -- P.string >> (fn (prev, id) =>
-      Toplevel.no_timing o Toplevel.imperative (fn () => insert_command prev id)));
+      Toplevel.imperative (fn () => insert_command prev id)));
 
 val _ =
-  OuterSyntax.improper_command "Isar.remove" "remove command (Isar editor model)" K.control
-    (P.string >> (fn id =>
-      Toplevel.no_timing o Toplevel.imperative (fn () => remove_command id)));
+  OuterSyntax.internal_command "Isar.remove"
+    (P.string >> (fn id => Toplevel.imperative (fn () => remove_command id)));
 
 end;
 
--- a/src/Pure/Isar/isar.scala	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/isar.scala	Mon Jan 05 07:54:16 2009 -0800
@@ -1,22 +1,19 @@
 /*  Title:      Pure/Isar/isar.scala
     Author:     Makarius
 
-Isar toplevel editor model.
+Isar document model.
 */
 
 package isabelle
 
-import java.util.Properties
-
 
 class Isar(isabelle_system: IsabelleSystem, results: EventBus[IsabelleProcess.Result], args: String*)
   extends IsabelleProcess(isabelle_system, results, args: _*)
 {
-
   /* basic editor commands */
 
-  def create_command(props: Properties, text: String) =
-    output_sync("Isar.command " + IsabelleSyntax.encode_properties(props) + " " +
+  def create_command(id: String, text: String) =
+    output_sync("Isar.command " + IsabelleSyntax.encode_string(id) + " " +
       IsabelleSyntax.encode_string(text))
 
   def insert_command(prev: String, id: String) =
@@ -25,5 +22,4 @@
 
   def remove_command(id: String) =
     output_sync("Isar.remove " + IsabelleSyntax.encode_string(id))
-
 }
--- a/src/Pure/Isar/isar_cmd.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/isar_cmd.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -53,8 +53,7 @@
   val print_configs: Toplevel.transition -> Toplevel.transition
   val print_theorems: Toplevel.transition -> Toplevel.transition
   val print_locales: Toplevel.transition -> Toplevel.transition
-  val print_locale: bool * (Locale.expr * Element.context list)
-    -> Toplevel.transition -> Toplevel.transition
+  val print_locale: bool * xstring -> Toplevel.transition -> Toplevel.transition
   val print_registrations: bool -> string -> Toplevel.transition -> Toplevel.transition
   val print_attributes: Toplevel.transition -> Toplevel.transition
   val print_simpset: Toplevel.transition -> Toplevel.transition
@@ -355,11 +354,11 @@
 val print_theorems = Toplevel.unknown_context o print_theorems_theory o print_theorems_proof;
 
 val print_locales = Toplevel.unknown_theory o
-  Toplevel.keep (Locale.print_locales o Toplevel.theory_of);
+  Toplevel.keep (NewLocale.print_locales o Toplevel.theory_of);
 
-fun print_locale (show_facts, (imports, body)) = Toplevel.unknown_theory o
+fun print_locale (show_facts, name) = Toplevel.unknown_theory o
   Toplevel.keep (fn state =>
-    Locale.print_locale (Toplevel.theory_of state) show_facts imports body);
+    NewLocale.print_locale (Toplevel.theory_of state) show_facts name);
 
 fun print_registrations show_wits name = Toplevel.unknown_context o
   Toplevel.keep (Toplevel.node_case
--- a/src/Pure/Isar/isar_syn.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/isar_syn.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -385,18 +385,18 @@
 (* locales *)
 
 val locale_val =
-  SpecParse.locale_expr --
+  SpecParse.locale_expression --
     Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||
-  Scan.repeat1 SpecParse.context_element >> pair Locale.empty;
+  Scan.repeat1 SpecParse.context_element >> pair ([], []);
 
 val _ =
   OuterSyntax.command "locale" "define named proof context" K.thy_decl
-    (P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (Locale.empty, []) -- P.opt_begin
+    (P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (([], []), []) -- P.opt_begin
       >> (fn ((name, (expr, elems)), begin) =>
           (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
-            (Locale.add_locale_cmd name expr elems #-> TheoryTarget.begin)));
-
-val opt_prefix = Scan.optional (P.binding --| P.$$$ ":") Binding.empty;
+            (Expression.add_locale_cmd name name expr elems #>
+              (fn ((target, notes), ctxt) => TheoryTarget.begin target ctxt |>
+                fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes))));
 
 val _ =
   OuterSyntax.command "sublocale"
@@ -407,6 +407,40 @@
 
 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)));
+
+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)));
+
+local
+
+val opt_prefix = Scan.optional (P.binding --| P.$$$ ":") Binding.empty;
+
+in
+
+val locale_val =
+  SpecParse.locale_expr --
+    Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||
+  Scan.repeat1 SpecParse.context_element >> pair Locale.empty;
+
+val _ =
+  OuterSyntax.command "class_locale" "define named proof context based on classes" K.thy_decl
+    (P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (Locale.empty, []) -- P.opt_begin
+      >> (fn ((name, (expr, elems)), begin) =>
+          (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
+            (Locale.add_locale_cmd name expr elems #-> TheoryTarget.begin)));
+
+val _ =
+  OuterSyntax.command "class_interpretation"
     "prove and register interpretation of locale expression in theory or locale" K.thy_goal
     (P.xname --| (P.$$$ "\\<subseteq>" || P.$$$ "<") -- P.!!! SpecParse.locale_expr
       >> (Toplevel.print oo (Toplevel.theory_to_proof o Locale.interpretation_in_locale I)) ||
@@ -416,7 +450,7 @@
               (Locale.interpretation_cmd (Binding.base_name name) expr insts)));
 
 val _ =
-  OuterSyntax.command "interpret"
+  OuterSyntax.command "class_interpret"
     "prove and register interpretation of locale expression in proof context"
     (K.tag_proof K.prf_goal)
     (opt_prefix -- SpecParse.locale_expr -- SpecParse.locale_insts
@@ -424,6 +458,8 @@
           Toplevel.proof'
             (fn int => Locale.interpret_cmd (Binding.base_name name) expr insts int)));
 
+end;
+
 
 (* classes *)
 
@@ -721,9 +757,13 @@
 
 (* nested commands *)
 
+val props_text =
+  Scan.optional ValueParse.properties [] -- P.position P.string >> (fn (props, (str, pos)) =>
+    (Position.of_properties (Position.default_properties pos props), str));
+
 val _ =
   OuterSyntax.improper_command "Isabelle.command" "nested Isabelle command" K.control
-    (P.props_text :|-- (fn (pos, str) =>
+    (props_text :|-- (fn (pos, str) =>
       (case OuterSyntax.parse pos str of
         [tr] => Scan.succeed (K tr)
       | _ => Scan.fail_with (K "exactly one command expected"))
@@ -734,7 +774,7 @@
 
 val _ =
   OuterSyntax.improper_command "init_toplevel" "init toplevel point-of-interest" K.control
-    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Isar.init_point));
+    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Isar.init));
 
 val _ =
   OuterSyntax.improper_command "linear_undo" "undo commands" K.control
@@ -817,7 +857,7 @@
 
 val _ =
   OuterSyntax.improper_command "print_locale" "print locale expression in this theory" K.diag
-    (opt_bang -- locale_val >> (Toplevel.no_timing oo IsarCmd.print_locale));
+    (opt_bang -- P.xname >> (Toplevel.no_timing oo IsarCmd.print_locale));
 
 val _ =
   OuterSyntax.improper_command "print_interps"
--- a/src/Pure/Isar/locale.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/locale.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,6 @@
 (*  Title:      Pure/Isar/locale.ML
-    ID:         $Id$
-    Author:     Clemens Ballarin, TU Muenchen; Markus Wenzel, LMU/TU Muenchen
+    Author:     Clemens Ballarin, TU Muenchen
+    Author:     Markus Wenzel, LMU/TU Muenchen
 
 Locales -- Isar proof contexts as meta-level predicates, with local
 syntax and implicit structures.
@@ -714,7 +714,7 @@
     val unifier' = fold Vartab.update_new (frozen_tvars ctxt (map #2 parms')) unifier;
 
     fun inst_parms (i, ps) =
-      List.foldr Term.add_typ_tfrees [] (map_filter snd ps)
+      List.foldr OldTerm.add_typ_tfrees [] (map_filter snd ps)
       |> map_filter (fn (a, S) =>
           let val T = Envir.norm_type unifier' (TVar ((a, i), S))
           in if T = TFree (a, S) then NONE else SOME (a, T) end)
@@ -1297,7 +1297,7 @@
 
 fun defs_ord (defs1, defs2) =
     list_ord (fn ((_, (d1, _)), (_, (d2, _))) =>
-      Term.fast_term_ord (d1, d2)) (defs1, defs2);
+      TermOrd.fast_term_ord (d1, d2)) (defs1, defs2);
 structure Defstab =
     TableFun(type key = (Attrib.binding * (term * term list)) list val ord = defs_ord);
 
@@ -1791,10 +1791,11 @@
     val name = Sign.full_bname thy bname;
 
     val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
-    val env = Term.add_term_free_names (body, []);
+    val env = Term.add_free_names body [];
     val xs = filter (member (op =) env o #1) parms;
     val Ts = map #2 xs;
-    val extraTs = (Term.term_tfrees body \\ List.foldr Term.add_typ_tfrees [] Ts)
+    val extraTs =
+      (Term.add_tfrees body [] \\ fold Term.add_tfreesT Ts [])
       |> Library.sort_wrt #1 |> map TFree;
     val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
 
@@ -1949,8 +1950,8 @@
     val elemss = import_elemss @ body_elemss |>
       map_filter (fn ((id, Assumed axs), elems) => SOME (id, elems) | _ => NONE);
 
-    val extraTs = List.foldr Term.add_term_tfrees [] exts' \\
-      List.foldr Term.add_typ_tfrees [] (map snd parms);
+    val extraTs = List.foldr OldTerm.add_term_tfrees [] exts' \\
+      List.foldr OldTerm.add_typ_tfrees [] (map snd parms);
     val _ = if null extraTs then ()
       else warning ("Additional type variable(s) in locale specification " ^ quote bname);
 
--- a/src/Pure/Isar/method.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/method.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/method.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Isar proof methods.
@@ -223,15 +222,9 @@
   if cond (Logic.strip_assums_concl prop)
   then Tactic.rtac rule i else no_tac);
 
-fun legacy_tac st =
-  (legacy_feature
-      ("Implicit use of prems in assumption proof" ^ Position.str_of (Position.thread_data ()));
-    all_tac st);
-
 fun assm_tac ctxt =
   assume_tac APPEND'
   Goal.assume_rule_tac ctxt APPEND'
-  (CHANGED o solve_tac (Assumption.prems_of ctxt) THEN' K legacy_tac) APPEND'
   cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND'
   cond_rtac (can Logic.dest_term) Drule.termI;
 
--- a/src/Pure/Isar/new_locale.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/new_locale.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -47,9 +47,11 @@
   val unfold_attrib: attribute
   val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
 
-  (* Identifiers and registrations *)
-  val clear_local_idents: Proof.context -> Proof.context
-  val add_global_registration: (string * Morphism.morphism) -> theory -> theory
+  (* Registrations *)
+  val add_global_registration: (string * (Morphism.morphism * Morphism.morphism)) ->
+    theory -> theory
+  val amend_global_registration: Morphism.morphism -> (string * Morphism.morphism) ->
+    theory -> theory
   val get_global_registrations: theory -> (string * Morphism.morphism) list
 
   (* Diagnostic *)
@@ -113,17 +115,17 @@
   val extend = I;
 
   fun join_locales _
-    (Loc {parameters, spec, decls = (decls1, decls2), notes, dependencies},
-      Loc {decls = (decls1', decls2'), notes = notes',
-        dependencies = dependencies', ...}) =
-        let fun s_merge x = merge (eq_snd (op =)) x in
-          Loc {parameters = parameters,
-            spec = spec,
-            decls = (s_merge (decls1, decls1'), s_merge (decls2, decls2')),
-            notes = s_merge (notes, notes'),
-            dependencies = s_merge (dependencies, dependencies')
-          }          
-        end;
+   (Loc {parameters, spec, decls = (decls1, decls2), notes, dependencies},
+    Loc {decls = (decls1', decls2'), notes = notes', dependencies = dependencies', ...}) =
+      Loc {
+        parameters = parameters,
+        spec = spec,
+        decls =
+         (merge (eq_snd op =) (decls1, decls1'),
+          merge (eq_snd op =) (decls2, decls2')),
+        notes = merge (eq_snd op =) (notes, notes'),
+        dependencies = merge (eq_snd op =) (dependencies, dependencies')};
+
   fun merge _ = NameSpace.join_tables join_locales;
 );
 
@@ -195,7 +197,7 @@
 
 val empty = ([] : identifiers);
 
-fun ident_eq thy ((n, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);
+fun ident_eq thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);
 
 local
 
@@ -219,16 +221,14 @@
     Ready _ => NONE |
     ids => SOME (Context.theory_map (IdentifiersData.put (Ready (finish thy ids))) thy))
   )));
-  
+
 fun get_global_idents thy =
   let val (Ready ids) = (IdentifiersData.get o Context.Theory) thy in ids end;
 val put_global_idents = Context.theory_map o IdentifiersData.put o Ready;
-val clear_global_idents = put_global_idents empty;
 
 fun get_local_idents ctxt =
   let val (Ready ids) = (IdentifiersData.get o Context.Proof) ctxt in ids end;
 val put_local_idents = Context.proof_map o IdentifiersData.put o Ready;
-val clear_local_idents = put_local_idents empty;
 
 end;
 
@@ -287,18 +287,18 @@
       fold_rev (fn (decl, _) => Context.proof_map (decl morph)) term_decls
   end;
 
-fun activate_notes activ_elem thy (name, morph) input =
+fun activate_notes activ_elem transfer thy (name, morph) input =
   let
     val Loc {notes, ...} = the_locale thy name;
     fun activate ((kind, facts), _) input =
       let
-        val facts' = facts |> Element.facts_map (Element.morph_ctxt morph)
+        val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph))
       in activ_elem (Notes (kind, facts')) input end;
   in
     fold_rev activate notes input
   end;
 
-fun activate_all name thy activ_elem (marked, input) =
+fun activate_all name thy activ_elem transfer (marked, input) =
   let
     val Loc {parameters = (_, params), spec = (asm, defs), ...} =
       the_locale thy name;
@@ -310,7 +310,7 @@
       (if not (null defs)
         then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))
         else I) |>
-      pair marked |> roundup thy (activate_notes activ_elem) (name, Morphism.identity)
+      pair marked |> roundup thy (activate_notes activ_elem transfer) (name, Morphism.identity)
   end;
 
 
@@ -331,7 +331,7 @@
       in
         ctxt |> fold Variable.auto_fixes (maps (map fst o snd) assms') |>
           ProofContext.add_assms_i Assumption.assume_export assms' |> snd
-     end 
+     end
   | init_local_elem (Defines defs) ctxt =
       let
         val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs
@@ -354,15 +354,19 @@
   roundup thy activate_decls dep (get_local_idents ctxt, ctxt) |> uncurry put_local_idents;
 
 fun activate_global_facts dep thy =
-  roundup thy (activate_notes init_global_elem) dep (get_global_idents thy, thy) |>
+  roundup thy (activate_notes init_global_elem Element.transfer_morphism)
+    dep (get_global_idents thy, thy) |>
   uncurry put_global_idents;
 
 fun activate_local_facts dep ctxt =
-  roundup (ProofContext.theory_of ctxt) (activate_notes init_local_elem) dep
+  roundup (ProofContext.theory_of ctxt)
+  (activate_notes init_local_elem (Element.transfer_morphism o ProofContext.theory_of)) dep
     (get_local_idents ctxt, ctxt) |>
   uncurry put_local_idents;
 
-fun init name thy = activate_all name thy init_local_elem (empty, ProofContext.init thy) |>
+fun init name thy =
+  activate_all name thy init_local_elem (Element.transfer_morphism o ProofContext.theory_of)
+    (empty, ProofContext.init thy) |>
   uncurry put_local_idents;
 
 fun print_locale thy show_facts name =
@@ -371,25 +375,74 @@
     val ctxt = init name' thy
   in
     Pretty.big_list "locale elements:"
-      (activate_all name' thy (cons_elem show_facts) (empty, []) |> snd |> rev |> 
+      (activate_all name' thy (cons_elem show_facts) (K (Element.transfer_morphism thy))
+        (empty, []) |> snd |> rev |>
         map (Element.pretty_ctxt ctxt) |> map Pretty.chunks) |> Pretty.writeln
   end
 
 end;
 
 
+(*** Registrations: interpretations in theories ***)
+
+(* FIXME only global variant needed *)
+structure RegistrationsData = GenericDataFun
+(
+  type T = ((string * (Morphism.morphism * Morphism.morphism)) * stamp) list;
+(* FIXME mixins need to be stamped *)
+    (* registrations, in reverse order of declaration *)
+  val empty = [];
+  val extend = I;
+  fun merge _ data : T = Library.merge (eq_snd op =) data;
+    (* FIXME consolidate with dependencies, consider one data slot only *)
+);
+
+val get_global_registrations =
+  Context.Theory #> RegistrationsData.get #> map fst #> map (apsnd op $>);
+
+fun add_global reg =
+  (Context.theory_map o RegistrationsData.map) (cons (reg, stamp ()));
+
+fun add_global_registration (name, (base_morph, export)) thy =
+  roundup thy (fn _ => fn (name', morph') => fn thy => add_global (name', (morph', export)) thy)
+    (name, base_morph) (get_global_idents thy, thy) |>
+    snd (* FIXME ?? uncurry put_global_idents *);
+
+fun amend_global_registration morph (name, base_morph) thy =
+  let
+    val regs = (Context.Theory #> RegistrationsData.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');
+    val i = find_index match (rev regs);
+    val _ = if i = ~1 then error ("No interpretation of locale " ^
+        quote (extern thy name) ^ " and parameter instantiation " ^
+        space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available.")
+      else ();
+  in
+    (Context.theory_map o RegistrationsData.map) (nth_map (length regs - 1 - i)
+      (fn ((name, (base, export)), stamp) => ((name, (base $> morph, export)), stamp))) thy
+  end;
+
+
 (*** Storing results ***)
 
 (* Theorems *)
 
 fun add_thmss loc kind args ctxt =
   let
-    val (([Notes args'], _), ctxt') = Element.activate_i [Notes (kind, args)] ctxt;
-    val ctxt'' = ctxt' |> ProofContext.theory
-      (change_locale loc
+    val ([Notes args'], ctxt') = Element.activate_i [Notes (kind, args)] ctxt;
+    val ctxt'' = ctxt' |> ProofContext.theory (
+      change_locale loc
         (fn (parameters, spec, decls, notes, dependencies) =>
-          (parameters, spec, decls, (args', stamp ()) :: notes, dependencies)))
-      (* FIXME registrations *)
+          (parameters, spec, decls, (args', stamp ()) :: notes, dependencies)) #>
+      (* Registrations *)
+      (fn thy => fold_rev (fn (name, morph) =>
+            let
+              val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>
+                Attrib.map_facts (Attrib.attribute_i thy)
+            in Locale.global_note_qualified kind args'' #> snd end)
+        (get_global_registrations thy |> filter (fn (name, _) => name = loc)) thy))
   in ctxt'' end;
 
 
@@ -451,23 +504,5 @@
         Locale.intro_locales_tac true ctxt)),
       "back-chain all introduction rules of locales")]));
 
-
-(*** Registrations: interpretations in theories ***)
-
-(* FIXME only global variant needed *)
-structure RegistrationsData = GenericDataFun
-(
-  type T = ((string * Morphism.morphism) * stamp) list;
-    (* registrations, in reverse order of declaration *)
-  val empty = [];
-  val extend = I;
-  fun merge _ = Library.merge (eq_snd (op =));
-    (* FIXME consolidate with dependencies, consider one data slot only *)
-);
-
-val get_global_registrations = map fst o RegistrationsData.get o Context.Theory;
-fun add_global_registration reg =
-  (Context.theory_map o RegistrationsData.map) (cons (reg, stamp ()));
-
 end;
 
--- a/src/Pure/Isar/outer_keyword.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/outer_keyword.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/outer_keyword.ML
-    ID:         $Id$
     Author:     Makarius
 
 Isar command keyword classification and global keyword tables.
@@ -47,9 +46,12 @@
   val report: unit -> unit
   val keyword: string -> unit
   val command: string -> T -> unit
+  val is_diag: string -> bool
+  val is_control: string -> bool
+  val is_regular: string -> bool
+  val is_theory_begin: string -> bool
   val is_theory: string -> bool
   val is_proof: string -> bool
-  val is_diag: string -> bool
   val is_theory_goal: string -> bool
   val is_proof_goal: string -> bool
   val is_qed: string -> bool
@@ -174,6 +176,12 @@
     NONE => false
   | SOME k => member (op = o pairself kind_of) ks k);
 
+val is_diag = command_category [diag];
+val is_control = command_category [control];
+fun is_regular name = not (is_diag name orelse is_control name);
+
+val is_theory_begin = command_category [thy_begin];
+
 val is_theory = command_category
   [thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal];
 
@@ -181,8 +189,6 @@
   [qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close,
     prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script];
 
-val is_diag = command_category [diag];
-
 val is_theory_goal = command_category [thy_goal];
 val is_proof_goal = command_category [prf_goal, prf_asm_goal];
 val is_qed = command_category [qed, qed_block];
--- a/src/Pure/Isar/outer_parse.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/outer_parse.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/outer_parse.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Generic parsers for Isabelle/Isar outer syntax.
@@ -8,48 +7,49 @@
 signature OUTER_PARSE =
 sig
   type token = OuterLex.token
+  type 'a parser = token list -> 'a * token list
   val group: string -> (token list -> 'a) -> token list -> 'a
   val !!! : (token list -> 'a) -> token list -> 'a
   val !!!! : (token list -> 'a) -> token list -> 'a
   val triple1: ('a * 'b) * 'c -> 'a * 'b * 'c
   val triple2: 'a * ('b * 'c) -> 'a * 'b * 'c
   val triple_swap: ('a * 'b) * 'c -> ('a * 'c) * 'b
-  val not_eof: token list -> token * token list
+  val not_eof: token parser
   val position: (token list -> 'a * 'b) -> token list -> ('a * Position.T) * 'b
-  val command: token list -> string * token list
-  val keyword: token list -> string * token list
-  val short_ident: token list -> string * token list
-  val long_ident: token list -> string * token list
-  val sym_ident: token list -> string * token list
-  val minus: token list -> string * token list
-  val term_var: token list -> string * token list
-  val type_ident: token list -> string * token list
-  val type_var: token list -> string * token list
-  val number: token list -> string * token list
-  val string: token list -> string * token list
-  val alt_string: token list -> string * token list
-  val verbatim: token list -> string * token list
-  val sync: token list -> string * token list
-  val eof: token list -> string * token list
-  val keyword_with: (string -> bool) -> token list -> string * token list
-  val keyword_ident_or_symbolic: token list -> string * token list
-  val $$$ : string -> token list -> string * token list
-  val reserved: string -> token list -> string * token list
-  val semicolon: token list -> string * token list
-  val underscore: token list -> string * token list
-  val maybe: (token list -> 'a * token list) -> token list -> 'a option * token list
-  val tag_name: token list -> string * token list
-  val tags: token list -> string list * token list
-  val opt_unit: token list -> unit * token list
-  val opt_keyword: string -> token list -> bool * token list
-  val begin: token list -> string * token list
-  val opt_begin: token list -> bool * token list
-  val nat: token list -> int * token list
-  val int: token list -> int * token list
-  val enum: string -> (token list -> 'a * token list) -> token list -> 'a list * token list
-  val enum1: string -> (token list -> 'a * token list) -> token list -> 'a list * token list
-  val and_list: (token list -> 'a * token list) -> token list -> 'a list * token list
-  val and_list1: (token list -> 'a * token list) -> token list -> 'a list * token list
+  val command: string parser
+  val keyword: string parser
+  val short_ident: string parser
+  val long_ident: string parser
+  val sym_ident: string parser
+  val minus: string parser
+  val term_var: string parser
+  val type_ident: string parser
+  val type_var: string parser
+  val number: string parser
+  val string: string parser
+  val alt_string: string parser
+  val verbatim: string parser
+  val sync: string parser
+  val eof: string parser
+  val keyword_with: (string -> bool) -> string parser
+  val keyword_ident_or_symbolic: string parser
+  val $$$ : string -> string parser
+  val reserved: string -> string parser
+  val semicolon: string parser
+  val underscore: string parser
+  val maybe: 'a parser -> 'a option parser
+  val tag_name: string parser
+  val tags: string list parser
+  val opt_unit: unit parser
+  val opt_keyword: string -> bool parser
+  val begin: string parser
+  val opt_begin: bool parser
+  val nat: int parser
+  val int: int parser
+  val enum: string -> 'a parser -> 'a list parser
+  val enum1: string -> 'a parser -> 'a list parser
+  val and_list: 'a parser -> 'a list parser
+  val and_list1: 'a parser -> 'a list parser
   val enum': string -> ('a * token list -> 'b * ('a * token list)) ->
     'a * token list -> 'b list * ('a * token list)
   val enum1': string -> ('a * token list -> 'b * ('a * token list)) ->
@@ -58,46 +58,44 @@
     'a * token list -> 'b list * ('a * token list)
   val and_list1': ('a * token list -> 'b * ('a * token list)) ->
     'a * token list -> 'b list * ('a * token list)
-  val list: (token list -> 'a * token list) -> token list -> 'a list * token list
-  val list1: (token list -> 'a * token list) -> token list -> 'a list * token list
-  val name: token list -> bstring * token list
-  val binding: token list -> Binding.T * token list
-  val xname: token list -> xstring * token list
-  val text: token list -> string * token list
-  val path: token list -> Path.T * token list
-  val parname: token list -> string * token list
-  val parbinding: token list -> Binding.T * token list
-  val sort: token list -> string * token list
-  val arity: token list -> (string * string list * string) * token list
-  val multi_arity: token list -> (string list * string list * string) * token list
-  val type_args: token list -> string list * token list
-  val typ_group: token list -> string * token list
-  val typ: token list -> string * token list
-  val mixfix: token list -> mixfix * token list
-  val mixfix': token list -> mixfix * token list
-  val opt_infix: token list -> mixfix * token list
-  val opt_infix': token list -> mixfix * token list
-  val opt_mixfix: token list -> mixfix * token list
-  val opt_mixfix': token list -> mixfix * token list
-  val where_: token list -> string * token list
-  val const: token list -> (string * string * mixfix) * token list
-  val params: token list -> (Binding.T * string option) list * token list
-  val simple_fixes: token list -> (Binding.T * string option) list * token list
-  val fixes: token list -> (Binding.T * string option * mixfix) list * token list
-  val for_fixes: token list -> (Binding.T * string option * mixfix) list * token list
-  val for_simple_fixes: token list -> (Binding.T * string option) list * token list
-  val ML_source: token list -> (SymbolPos.text * Position.T) * token list
-  val doc_source: token list -> (SymbolPos.text * Position.T) * token list
-  val properties: token list -> Properties.T * token list
-  val props_text: token list -> (Position.T * string) * token list
-  val term_group: token list -> string * token list
-  val prop_group: token list -> string * token list
-  val term: token list -> string * token list
-  val prop: token list -> string * token list
-  val propp: token list -> (string * string list) * token list
-  val termp: token list -> (string * string list) * token list
-  val target: token list -> xstring * token list
-  val opt_target: token list -> xstring option * token list
+  val list: 'a parser -> 'a list parser
+  val list1: 'a parser -> 'a list parser
+  val name: bstring parser
+  val binding: Binding.T parser
+  val xname: xstring parser
+  val text: string parser
+  val path: Path.T parser
+  val parname: string parser
+  val parbinding: Binding.T parser
+  val sort: string parser
+  val arity: (string * string list * string) parser
+  val multi_arity: (string list * string list * string) parser
+  val type_args: string list parser
+  val typ_group: string parser
+  val typ: string parser
+  val mixfix: mixfix parser
+  val mixfix': mixfix parser
+  val opt_infix: mixfix parser
+  val opt_infix': mixfix parser
+  val opt_mixfix: mixfix parser
+  val opt_mixfix': mixfix parser
+  val where_: string parser
+  val const: (string * string * mixfix) parser
+  val params: (Binding.T * string option) list parser
+  val simple_fixes: (Binding.T * string option) list parser
+  val fixes: (Binding.T * string option * mixfix) list parser
+  val for_fixes: (Binding.T * string option * mixfix) list parser
+  val for_simple_fixes: (Binding.T * string option) list parser
+  val ML_source: (SymbolPos.text * Position.T) parser
+  val doc_source: (SymbolPos.text * Position.T) parser
+  val term_group: string parser
+  val prop_group: string parser
+  val term: string parser
+  val prop: string parser
+  val propp: (string * string list) parser
+  val termp: (string * string list) parser
+  val target: xstring parser
+  val opt_target: xstring option parser
 end;
 
 structure OuterParse: OUTER_PARSE =
@@ -106,6 +104,8 @@
 structure T = OuterLex;
 type token = T.token;
 
+type 'a parser = token list -> 'a * token list;
+
 
 (** error handling **)
 
@@ -310,12 +310,6 @@
 val ML_source = source_position (group "ML source" text);
 val doc_source = source_position (group "document source" text);
 
-val properties = $$$ "(" |-- !!! (list1 (string -- ($$$ "=" |-- string)) --| $$$ ")");
-
-val props_text =
-  Scan.optional properties [] -- position string >> (fn (props, (str, pos)) =>
-    (Position.of_properties (Position.default_properties pos props), str));
-
 
 (* terms *)
 
--- a/src/Pure/Isar/outer_syntax.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/outer_syntax.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/outer_syntax.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 The global Isabelle/Isar outer syntax.
@@ -10,17 +9,20 @@
 
 signature OUTER_SYNTAX =
 sig
-  type parser_fn = OuterLex.token list ->
-    (Toplevel.transition -> Toplevel.transition) * OuterLex.token list
-  val command: string -> string -> OuterKeyword.T -> parser_fn -> unit
-  val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T -> parser_fn -> unit
-  val improper_command: string -> string -> OuterKeyword.T -> parser_fn -> unit
+  type 'a parser = 'a OuterParse.parser
+  val command: string -> string -> OuterKeyword.T ->
+    (Toplevel.transition -> Toplevel.transition) parser -> unit
+  val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T ->
+    (Toplevel.transition -> Toplevel.transition) parser -> unit
+  val improper_command: string -> string -> OuterKeyword.T ->
+    (Toplevel.transition -> Toplevel.transition) parser -> unit
+  val internal_command: string -> (Toplevel.transition -> Toplevel.transition) parser -> unit
   val local_theory: string -> string -> OuterKeyword.T ->
-    (OuterParse.token list -> (local_theory -> local_theory) * OuterLex.token list) -> unit
+    (local_theory -> local_theory) parser -> unit
   val local_theory_to_proof': string -> string -> OuterKeyword.T ->
-    (OuterParse.token list -> (bool -> local_theory -> Proof.state) * OuterLex.token list) -> unit
+    (bool -> local_theory -> Proof.state) parser -> unit
   val local_theory_to_proof: string -> string -> OuterKeyword.T ->
-    (OuterParse.token list -> (local_theory -> Proof.state) * OuterLex.token list) -> unit
+    (local_theory -> Proof.state) parser -> unit
   val print_outer_syntax: unit -> unit
   val scan: Position.T -> string -> OuterLex.token list
   val parse: Position.T -> string -> Toplevel.transition list
@@ -36,22 +38,22 @@
 
 structure T = OuterLex;
 structure P = OuterParse;
+type 'a parser = 'a P.parser;
+
 
 
 (** outer syntax **)
 
-(* parsers *)
+(* command parsers *)
 
-type parser_fn = T.token list -> (Toplevel.transition -> Toplevel.transition) * T.token list;
-
-datatype parser = Parser of
+datatype command = Command of
  {comment: string,
   markup: ThyOutput.markup option,
   int_only: bool,
-  parse: parser_fn};
+  parse: (Toplevel.transition -> Toplevel.transition) parser};
 
-fun make_parser comment markup int_only parse =
-  Parser {comment = comment, markup = markup, int_only = int_only, parse = parse};
+fun make_command comment markup int_only parse =
+  Command {comment = comment, markup = markup, int_only = int_only, parse = parse};
 
 
 (* parse command *)
@@ -63,7 +65,7 @@
 
 fun body cmd (name, _) =
   (case cmd name of
-    SOME (Parser {int_only, parse, ...}) =>
+    SOME (Command {int_only, parse, ...}) =>
       P.!!! (Scan.prompt (name ^ "# ") (P.tags |-- parse >> pair int_only))
   | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
 
@@ -85,47 +87,50 @@
 
 local
 
-val global_parsers = ref (Symtab.empty: parser Symtab.table);
+val global_commands = ref (Symtab.empty: command Symtab.table);
 val global_markups = ref ([]: (string * ThyOutput.markup) list);
 
-fun change_parsers f = CRITICAL (fn () =>
- (change global_parsers f;
+fun change_commands f = CRITICAL (fn () =>
+ (change global_commands f;
   global_markups :=
-    Symtab.fold (fn (name, Parser {markup = SOME m, ...}) => cons (name, m) | _ => I)
-      (! global_parsers) []));
+    Symtab.fold (fn (name, Command {markup = SOME m, ...}) => cons (name, m) | _ => I)
+      (! global_commands) []));
 
 in
 
 (* access current syntax *)
 
-fun get_parsers () = CRITICAL (fn () => ! global_parsers);
+fun get_commands () = CRITICAL (fn () => ! global_commands);
 fun get_markups () = CRITICAL (fn () => ! global_markups);
 
-fun get_parser () = Symtab.lookup (get_parsers ());
-fun get_syntax () = CRITICAL (fn () => (OuterKeyword.get_lexicons (), get_parser ()));
+fun get_command () = Symtab.lookup (get_commands ());
+fun get_syntax () = CRITICAL (fn () => (OuterKeyword.get_lexicons (), get_command ()));
 
 fun is_markup kind name = AList.lookup (op =) (get_markups ()) name = SOME kind;
 
 
 (* augment syntax *)
 
-fun add_parser name kind parser = CRITICAL (fn () =>
+fun add_command name kind cmd = CRITICAL (fn () =>
  (OuterKeyword.command name kind;
-  if not (Symtab.defined (get_parsers ()) name) then ()
+  if not (Symtab.defined (get_commands ()) name) then ()
   else warning ("Redefining outer syntax command " ^ quote name);
-  change_parsers (Symtab.update (name, parser))));
+  change_commands (Symtab.update (name, cmd))));
 
 fun command name comment kind parse =
-  add_parser name kind (make_parser comment NONE false parse);
+  add_command name kind (make_command comment NONE false parse);
 
 fun markup_command markup name comment kind parse =
-  add_parser name kind (make_parser comment (SOME markup) false parse);
+  add_command name kind (make_command comment (SOME markup) false parse);
 
 fun improper_command name comment kind parse =
-  add_parser name kind (make_parser comment NONE true parse);
+  add_command name kind (make_command comment NONE true parse);
 
 end;
 
+fun internal_command name parse =
+  command name "(internal)" OuterKeyword.control (parse >> (fn tr => Toplevel.no_timing o tr));
+
 
 (* local_theory commands *)
 
@@ -133,22 +138,22 @@
   command name comment kind (P.opt_target -- parse
     >> (fn (loc, f) => (if do_print then Toplevel.print else I) o trans loc f));
 
-val local_theory           = local_theory_command false Toplevel.local_theory;
+val local_theory = local_theory_command false Toplevel.local_theory;
 val local_theory_to_proof' = local_theory_command true Toplevel.local_theory_to_proof';
-val local_theory_to_proof  = local_theory_command true Toplevel.local_theory_to_proof;
+val local_theory_to_proof = local_theory_command true Toplevel.local_theory_to_proof;
 
 
 (* inspect syntax *)
 
-fun dest_parsers () =
-  get_parsers () |> Symtab.dest |> sort_wrt #1
-  |> map (fn (name, Parser {comment, int_only, ...}) => (name, comment, int_only));
+fun dest_commands () =
+  get_commands () |> Symtab.dest |> sort_wrt #1
+  |> map (fn (name, Command {comment, int_only, ...}) => (name, comment, int_only));
 
 fun print_outer_syntax () =
   let
     fun pretty_cmd (name, comment, _) =
       Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
-    val (int_cmds, cmds) = List.partition #3 (dest_parsers ());
+    val (int_cmds, cmds) = List.partition #3 (dest_commands ());
   in
     [Pretty.strs ("syntax keywords:" :: map quote (OuterKeyword.dest_keywords ())),
       Pretty.big_list "commands:" (map pretty_cmd cmds),
@@ -194,7 +199,7 @@
   Source.of_string str
   |> Symbol.source {do_recover = false}
   |> T.source {do_recover = SOME false} OuterKeyword.get_lexicons pos
-  |> toplevel_source false NONE get_parser
+  |> toplevel_source false NONE get_command
   |> Source.exhaust;
 
 
@@ -225,39 +230,39 @@
   Source.tty
   |> Symbol.source {do_recover = true}
   |> T.source {do_recover = SOME true} OuterKeyword.get_lexicons Position.none
-  |> toplevel_source term (SOME true) get_parser;
+  |> toplevel_source term (SOME true) get_command;
 
 
 (* prepare toplevel commands -- fail-safe *)
 
 val not_singleton = "Exactly one command expected";
 
-fun prepare_span parser span =
+fun prepare_span commands span =
   let
-    val range_pos = Position.encode_range (ThyEdit.span_range span);
-    val toks = ThyEdit.span_content span;
-    val _ = List.app ThyEdit.report_token toks;
+    val range_pos = Position.encode_range (ThySyntax.span_range span);
+    val toks = ThySyntax.span_content span;
+    val _ = List.app ThySyntax.report_token toks;
   in
-    (case Source.exhaust (toplevel_source false NONE (K parser) (Source.of_list toks)) of
+    (case Source.exhaust (toplevel_source false NONE (K commands) (Source.of_list toks)) of
       [tr] => (tr, true)
     | [] => (Toplevel.ignored range_pos, false)
     | _ => (Toplevel.malformed range_pos not_singleton, true))
     handle ERROR msg => (Toplevel.malformed range_pos msg, true)
   end;
 
-fun prepare_unit parser (cmd, proof, proper_proof) =
+fun prepare_unit commands (cmd, proof, proper_proof) =
   let
-    val (tr, proper_cmd) = prepare_span parser cmd;
-    val proof_trs = map (prepare_span parser) proof |> filter #2 |> map #1;
+    val (tr, proper_cmd) = prepare_span commands cmd;
+    val proof_trs = map (prepare_span commands) proof |> filter #2 |> map #1;
   in
     if proper_cmd andalso proper_proof then [(tr, proof_trs)]
     else map (rpair []) (if proper_cmd then tr :: proof_trs else proof_trs)
   end;
 
 fun prepare_command pos str =
-  let val (lexs, parser) = get_syntax () in
-    (case ThyEdit.parse_spans lexs pos str of
-      [span] => #1 (prepare_span parser span)
+  let val (lexs, commands) = get_syntax () in
+    (case ThySyntax.parse_spans lexs pos str of
+      [span] => #1 (prepare_span commands span)
     | _ => Toplevel.malformed pos not_singleton)
   end;
 
@@ -266,18 +271,18 @@
 
 fun load_thy name pos text time =
   let
-    val (lexs, parser) = get_syntax ();
+    val (lexs, commands) = get_syntax ();
 
     val _ = Present.init_theory name;
 
-    val toks = Source.exhausted (ThyEdit.token_source lexs pos (Source.of_list text));
-    val spans = Source.exhaust (ThyEdit.span_source toks);
-    val _ = List.app ThyEdit.report_span spans;
-    val units = Source.exhaust (ThyEdit.unit_source (Source.of_list spans))
-      |> maps (prepare_unit parser);
+    val toks = Source.exhausted (ThySyntax.token_source lexs pos (Source.of_list text));
+    val spans = Source.exhaust (ThySyntax.span_source toks);
+    val _ = List.app ThySyntax.report_span spans;
+    val units = Source.exhaust (ThySyntax.unit_source (Source.of_list spans))
+      |> maps (prepare_unit commands);
 
     val _ = Present.theory_source name
-      (fn () => HTML.html_mode (implode o map ThyEdit.present_span) spans);
+      (fn () => HTML.html_mode (implode o map ThySyntax.present_span) spans);
 
     val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
     val _ = cond_timeit time "" (fn () =>
--- a/src/Pure/Isar/proof.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/proof.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -115,6 +115,7 @@
     ((Binding.T * attribute list) * (term * term list) list) list -> bool -> state -> state
   val is_relevant: state -> bool
   val future_proof: (state -> ('a * context) future) -> state -> 'a future * context
+  val future_terminal_proof: Method.text * Method.text option -> state -> context
 end;
 
 structure Proof: PROOF =
@@ -140,7 +141,7 @@
     goal: goal option}
 and goal =
   Goal of
-   {statement: (string * Position.T) * term list list * cterm,
+   {statement: (string * Position.T) * term list list * term,
       (*goal kind and statement (starting with vars), initial proposition*)
     messages: (unit -> Pretty.T) list,    (*persistent messages (hints etc.)*)
     using: thm list,                      (*goal facts*)
@@ -313,8 +314,8 @@
 
 fun schematic_goal state =
   let val (_, (_, {statement = (_, _, prop), ...})) = find_goal state in
-    Term.exists_subterm Term.is_Var (Thm.term_of prop) orelse
-    Term.exists_type (Term.exists_subtype Term.is_TVar) (Thm.term_of prop)
+    Term.exists_subterm Term.is_Var prop orelse
+    Term.exists_type (Term.exists_subtype Term.is_TVar) prop
   end;
 
 
@@ -806,9 +807,10 @@
 
     val propss' = map (Logic.mk_term o Var) vars :: propss;
     val goal_propss = filter_out null propss';
-    val goal = cert (Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss))
+    val goal =
+      cert (Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss))
       |> Thm.weaken_sorts (Variable.sorts_of (context_of goal_state));
-    val statement = ((kind, pos), propss', goal);
+    val statement = ((kind, pos), propss', Thm.term_of goal);
     val after_qed' = after_qed |>> (fn after_local =>
       fn results => map_context after_ctxt #> after_local results);
   in
@@ -979,40 +981,45 @@
 
 (* fork global proofs *)
 
-fun is_initial state =
+local
+
+fun is_original state =
   (case try find_goal state of
     NONE => false
-  | SOME (_, (_, {statement = (_, _, prop), goal, ...})) => Thm.eq_thm_prop (Goal.init prop, goal));
+  | SOME (_, (_, {statement = (_, _, prop), goal, ...})) =>
+      Logic.protect prop aconv Thm.concl_of goal);
+
+in
 
 fun is_relevant state =
   can (assert_bottom false) state orelse
   can assert_forward state orelse
-  not (is_initial state) orelse
+  not (is_original state) orelse
   schematic_goal state;
 
-fun future_proof proof state =
+fun future_proof fork_proof state =
   let
     val _ = is_relevant state andalso error "Cannot fork relevant proof";
 
     val (goal_ctxt, (_, goal)) = find_goal state;
-    val {statement as (kind, propss, cprop), messages, using, goal, before_qed, after_qed} = goal;
-    val prop = Thm.term_of cprop;
-    val cprop_protected = #2 (Thm.dest_implies (Thm.cprop_of goal));
+    val {statement as (kind, propss, prop), messages, using, goal, before_qed, after_qed} = goal;
 
-    val statement' = (kind, [[], [Thm.term_of cprop_protected]], cprop_protected);
-    val goal' = Thm.adjust_maxidx_thm ~1 (Drule.protectI RS Goal.init cprop_protected);
-    fun after_qed' [[th]] = ProofContext.put_thms false (AutoBind.thisN, SOME [Goal.conclude th]);
-    val this_name = ProofContext.full_bname (ProofContext.reset_naming goal_ctxt) AutoBind.thisN;
+    val prop' = Logic.protect prop;
+    val statement' = (kind, [[], [prop']], prop');
+    val goal' = Thm.adjust_maxidx_thm (Thm.maxidx_of goal)
+      (Drule.comp_no_flatten (goal, Thm.nprems_of goal) 1 Drule.protectI);
+    fun after_qed' [[th]] = ProofContext.put_thms false (AutoBind.thisN, SOME [th]);
+    val this_name = ProofContext.full_bname goal_ctxt AutoBind.thisN;
 
     val result_ctxt =
       state
       |> map_contexts (Variable.auto_fixes prop)
       |> map_goal I (K (statement', messages, using, goal', before_qed, (#1 after_qed, after_qed')))
-      |> proof;
+      |> fork_proof;
 
-    val thm = result_ctxt |> Future.map (fn (_, ctxt) =>
+    val future_thm = result_ctxt |> Future.map (fn (_, ctxt) =>
       ProofContext.get_fact_single ctxt (Facts.named this_name));
-    val finished_goal = Goal.protect (Goal.future_result goal_ctxt thm prop);
+    val finished_goal = Goal.future_result goal_ctxt future_thm prop';
     val state' =
       state
       |> map_goal I (K (statement, messages, using, finished_goal, NONE, after_qed))
@@ -1020,3 +1027,11 @@
   in (Future.map #1 result_ctxt, state') end;
 
 end;
+
+fun future_terminal_proof meths state =
+  if is_relevant state then global_terminal_proof meths state
+  else #2 (state |> future_proof (fn state' =>
+    Future.fork_pri ~1 (fn () => ((), global_terminal_proof meths state'))));
+
+end;
+
--- a/src/Pure/Isar/proof_context.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/proof_context.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -408,9 +408,8 @@
    ("free", free_or_skolem),
    ("bound", plain_markup Markup.bound),
    ("var", var_or_skolem),
-   ("num", plain_markup Markup.num),
-   ("xnum", plain_markup Markup.xnum),
-   ("xstr", plain_markup Markup.xstr)];
+   ("numeral", plain_markup Markup.numeral),
+   ("inner_string", plain_markup Markup.inner_string)];
 
 in val _ = Context.>> (Context.map_theory (Sign.add_tokentrfuns token_trans)) end;
 
--- a/src/Pure/Isar/rule_cases.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/rule_cases.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/rule_cases.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Annotations and local contexts of rules.
@@ -319,7 +318,7 @@
 local
 
 fun equal_cterms ts us =
-  is_equal (list_ord (Term.fast_term_ord o pairself Thm.term_of) (ts, us));
+  is_equal (list_ord (TermOrd.fast_term_ord o pairself Thm.term_of) (ts, us));
 
 fun prep_rule n th =
   let
--- a/src/Pure/Isar/spec_parse.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/spec_parse.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/spec_parse.ML
-    ID:         $Id$
     Author:     Makarius
 
 Parsers for complex specifications.
@@ -7,35 +6,33 @@
 
 signature SPEC_PARSE =
 sig
-  type token
-  val attrib: OuterLex.token list -> Attrib.src * token list
-  val attribs: token list -> Attrib.src list * token list
-  val opt_attribs: token list -> Attrib.src list * token list
-  val thm_name: string -> token list -> Attrib.binding * token list
-  val opt_thm_name: string -> token list -> Attrib.binding * token list
-  val spec: token list -> (Attrib.binding * string list) * token list
-  val named_spec: token list -> (Attrib.binding * string list) * token list
-  val spec_name: token list -> ((Binding.T * string) * Attrib.src list) * token list
-  val spec_opt_name: token list -> ((Binding.T * string) * Attrib.src list) * token list
-  val xthm: token list -> (Facts.ref * Attrib.src list) * token list
-  val xthms1: token list -> (Facts.ref * Attrib.src list) list * token list
-  val name_facts: token list ->
-    (Attrib.binding * (Facts.ref * Attrib.src list) list) list * token list
-  val locale_mixfix: token list -> mixfix * token list
-  val locale_fixes: token list -> (Binding.T * string option * mixfix) list * token list
-  val locale_insts: token list -> (string option list * (Attrib.binding * string) list) * token list
-  val class_expr: token list -> string list * token list
-  val locale_expr: token list -> Locale.expr * token list
-  val locale_expression: OuterParse.token list -> Expression.expression * OuterParse.token list
-  val locale_keyword: token list -> string * token list
-  val context_element: token list -> Element.context * token list
-  val statement: token list -> (Attrib.binding * (string * string list) list) list * token list
-  val general_statement: token list ->
-    (Element.context list * Element.statement) * OuterLex.token list
-  val statement_keyword: token list -> string * token list
-  val specification: token list ->
-    (Binding.T *
-      ((Attrib.binding * string list) list * (Binding.T * string option) list)) list * token list
+  type token = OuterParse.token
+  type 'a parser = 'a OuterParse.parser
+  val attrib: Attrib.src parser
+  val attribs: Attrib.src list parser
+  val opt_attribs: Attrib.src list parser
+  val thm_name: string -> Attrib.binding parser
+  val opt_thm_name: string -> Attrib.binding parser
+  val spec: (Attrib.binding * string list) parser
+  val named_spec: (Attrib.binding * string list) parser
+  val spec_name: ((Binding.T * string) * Attrib.src list) parser
+  val spec_opt_name: ((Binding.T * string) * Attrib.src list) parser
+  val xthm: (Facts.ref * Attrib.src list) parser
+  val xthms1: (Facts.ref * Attrib.src list) list parser
+  val name_facts: (Attrib.binding * (Facts.ref * Attrib.src list) list) list parser
+  val locale_mixfix: mixfix parser
+  val locale_fixes: (Binding.T * string option * mixfix) list parser
+  val locale_insts: (string option list * (Attrib.binding * string) list) parser
+  val class_expr: string list parser
+  val locale_expr: Locale.expr parser
+  val locale_expression: Expression.expression parser
+  val locale_keyword: string parser
+  val context_element: Element.context parser
+  val statement: (Attrib.binding * (string * string list) list) list parser
+  val general_statement: (Element.context list * Element.statement) parser
+  val statement_keyword: string parser
+  val specification:
+    (Binding.T * ((Attrib.binding * string list) list * (Binding.T * string option) list)) list parser
 end;
 
 structure SpecParse: SPEC_PARSE =
@@ -43,6 +40,7 @@
 
 structure P = OuterParse;
 type token = P.token;
+type 'a parser = 'a P.parser;
 
 
 (* theorem specifications *)
@@ -104,7 +102,7 @@
 
 val rename = P.name -- Scan.option P.mixfix;
 
-val prefix = P.name --| 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 ||
@@ -127,7 +125,7 @@
 val locale_expression =
   let
     fun expr2 x = P.xname x;
-    fun expr1 x = (Scan.optional prefix "" -- expr2 --
+    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;
   in expr0 -- Scan.optional (P.$$$ "for" |-- P.!!! locale_fixes) [] end;
--- a/src/Pure/Isar/specification.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/specification.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -59,18 +59,6 @@
 structure Specification: SPECIFICATION =
 struct
 
-(* new locales *)
-
-fun cert_stmt body concl ctxt =
-  let val (_, _, ctxt', concl') = Locale.cert_context_statement NONE body concl ctxt
-  in (concl', ctxt') end;
-fun read_stmt body concl ctxt =
-  let val (_, _, ctxt', concl') = Locale.read_context_statement NONE body concl ctxt
-  in (concl', ctxt') end;
-  
-fun cert_context_statement x = if !new_locales then Expression.cert_statement x else cert_stmt x;
-fun read_context_statement x = if !new_locales then Expression.read_statement x else read_stmt x;
-
 (* diagnostics *)
 
 fun print_consts _ _ [] = ()
@@ -370,8 +358,8 @@
 
 in
 
-val theorem = gen_theorem (K I) cert_context_statement;
-val theorem_cmd = gen_theorem Attrib.intern_src read_context_statement;
+val theorem = gen_theorem (K I) Expression.cert_statement;
+val theorem_cmd = gen_theorem Attrib.intern_src Expression.read_statement;
 
 fun add_theorem_hook f = TheoremHook.map (cons (f, stamp ()));
 
--- a/src/Pure/Isar/subclass.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/subclass.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/subclass.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 User interface for proving subclass relationship between type classes.
@@ -22,7 +21,7 @@
     val thy = ProofContext.theory_of lthy;
     val sup = prep_class thy raw_sup;
     val sub = case TheoryTarget.peek lthy
-     of {is_class = false, ...} => error "No class context"
+     of {is_class = false, ...} => error "Not a class context"
       | {target, ...} => target;
     val _ = if Sign.subsort thy ([sup], [sub])
       then error ("Class " ^ Syntax.string_of_sort lthy [sup]
--- a/src/Pure/Isar/theory_target.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/theory_target.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -6,7 +6,7 @@
 
 signature THEORY_TARGET =
 sig
-  val peek: local_theory -> {target: string, is_locale: bool,
+  val peek: local_theory -> {target: string, new_locale: bool, is_locale: bool,
     is_class: bool, instantiation: string list * (string * sort) list * sort,
     overloading: (string * (string * typ) * bool) list}
   val init: string option -> theory -> local_theory
@@ -22,25 +22,32 @@
 
 (* new locales *)
 
-fun locale_extern x = if !new_locales then NewLocale.extern x else Locale.extern x;
-fun locale_add_type_syntax x = if !new_locales then NewLocale.add_type_syntax x else Locale.add_type_syntax x;
-fun locale_add_term_syntax x = if !new_locales then NewLocale.add_term_syntax x else Locale.add_term_syntax x;
-fun locale_add_declaration x = if !new_locales then NewLocale.add_declaration x else Locale.add_declaration x;
-fun locale_add_thmss x = if !new_locales then NewLocale.add_thmss x else Locale.add_thmss x;
-fun locale_init x = if !new_locales then NewLocale.init x else Locale.init x;
-fun locale_intern x = if !new_locales then NewLocale.intern x else Locale.intern x;
+fun locale_extern new_locale x = 
+  if new_locale then NewLocale.extern x else Locale.extern x;
+fun locale_add_type_syntax new_locale x =
+  if new_locale then NewLocale.add_type_syntax x else Locale.add_type_syntax x;
+fun locale_add_term_syntax new_locale x =
+  if new_locale then NewLocale.add_term_syntax x else Locale.add_term_syntax x;
+fun locale_add_declaration new_locale x =
+  if new_locale then NewLocale.add_declaration x else Locale.add_declaration x;
+fun locale_add_thmss new_locale x =
+  if new_locale then NewLocale.add_thmss x else Locale.add_thmss x;
+fun locale_init new_locale x =
+  if new_locale then NewLocale.init x else Locale.init x;
+fun locale_intern new_locale x =
+  if new_locale then NewLocale.intern x else Locale.intern x;
 
 (* context data *)
 
-datatype target = Target of {target: string, is_locale: bool,
+datatype target = Target of {target: string, new_locale: bool, is_locale: bool,
   is_class: bool, instantiation: string list * (string * sort) list * sort,
   overloading: (string * (string * typ) * bool) list};
 
-fun make_target target is_locale is_class instantiation overloading =
-  Target {target = target, is_locale = is_locale,
+fun make_target target new_locale is_locale is_class instantiation overloading =
+  Target {target = target, new_locale = new_locale, is_locale = is_locale,
     is_class = is_class, instantiation = instantiation, overloading = overloading};
 
-val global_target = make_target "" false false ([], [], []) [];
+val global_target = make_target "" false false false ([], [], []) [];
 
 structure Data = ProofDataFun
 (
@@ -56,7 +63,7 @@
 fun pretty_thy ctxt target is_locale is_class =
   let
     val thy = ProofContext.theory_of ctxt;
-    val target_name = (if is_class then "class " else "locale ") ^ locale_extern thy target;
+    val target_name = (if is_class then "class " else "locale ") ^ locale_extern is_class thy target;
     val fixes = map (fn (x, T) => (Binding.name x, SOME T, NoSyn))
       (#1 (ProofContext.inferred_fixes ctxt));
     val assumes = map (fn A => (Attrib.empty_binding, [(Thm.term_of A, [])]))
@@ -71,7 +78,7 @@
       (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)]
   end;
 
-fun pretty (Target {target, is_locale, is_class, instantiation, overloading}) ctxt =
+fun pretty (Target {target, is_locale, is_class, instantiation, overloading, ...}) ctxt =
   Pretty.block [Pretty.str "theory", Pretty.brk 1,
       Pretty.str (Context.theory_name (ProofContext.theory_of ctxt))] ::
     (if not (null overloading) then [Overloading.pretty ctxt]
@@ -81,7 +88,7 @@
 
 (* target declarations *)
 
-fun target_decl add (Target {target, is_class, ...}) d lthy =
+fun target_decl add (Target {target, new_locale, ...}) d lthy =
   let
     val d' = Morphism.transform (LocalTheory.target_morphism lthy) d;
     val d0 = Morphism.form d';
@@ -92,7 +99,7 @@
       |> LocalTheory.target (Context.proof_map d0)
     else
       lthy
-      |> LocalTheory.target (add target d')
+      |> LocalTheory.target (add new_locale target d')
   end;
 
 val type_syntax = target_decl locale_add_type_syntax;
@@ -158,7 +165,7 @@
   |> ProofContext.note_thmss_i kind facts
   ||> ProofContext.restore_naming ctxt;
 
-fun notes (Target {target, is_locale, is_class, ...}) kind facts lthy =
+fun notes (Target {target, is_locale, new_locale, ...}) kind facts lthy =
   let
     val thy = ProofContext.theory_of lthy;
     val facts' = facts
@@ -177,7 +184,7 @@
         #> PureThy.note_thmss_grouped kind (LocalTheory.group_of lthy) global_facts #> snd
         #> Sign.restore_naming thy)
     |> not is_locale ? LocalTheory.target (note_local kind global_facts #> snd)
-    |> is_locale ? LocalTheory.target (locale_add_thmss target kind target_facts)
+    |> is_locale ? LocalTheory.target (locale_add_thmss new_locale target kind target_facts)
     |> note_local kind local_facts
   end;
 
@@ -326,13 +333,14 @@
 
 fun init_target _ NONE = global_target
   | init_target thy (SOME target) =
-      make_target target true (Class.is_class thy target) ([], [], []) [];
+      make_target target (NewLocale.test_locale thy (NewLocale.intern thy target))
+      true (Class.is_class thy target) ([], [], []) [];
 
-fun init_ctxt (Target {target, is_locale, is_class, instantiation, overloading}) =
+fun init_ctxt (Target {target, new_locale, is_locale, is_class, instantiation, overloading}) =
   if not (null (#1 instantiation)) then Class.init_instantiation instantiation
   else if not (null overloading) then Overloading.init overloading
   else if not is_locale then ProofContext.init
-  else if not is_class then locale_init target
+  else if not is_class then locale_init new_locale target
   else Class.init target;
 
 fun init_lthy (ta as Target {target, instantiation, overloading, ...}) =
@@ -357,7 +365,7 @@
     val ctxt = ProofContext.init thy;
     val ops = raw_ops |> map (fn (name, const, checked) =>
       (name, Term.dest_Const (prep_const ctxt const), checked));
-  in thy |> init_lthy_ctxt (make_target "" false false ([], [], []) ops) end;
+  in thy |> init_lthy_ctxt (make_target "" false false false ([], [], []) ops) end;
 
 in
 
@@ -365,9 +373,10 @@
 fun begin target ctxt = init_lthy (init_target (ProofContext.theory_of ctxt) (SOME target)) ctxt;
 
 fun context "-" thy = init NONE thy
-  | context target thy = init (SOME (locale_intern thy target)) thy;
+  | context target thy = init (SOME (locale_intern
+      (NewLocale.test_locale thy (NewLocale.intern thy target)) thy target)) thy;
 
-fun instantiation arities = init_lthy_ctxt (make_target "" false false arities []);
+fun instantiation arities = init_lthy_ctxt (make_target "" false false false arities []);
 
 val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const);
 val overloading_cmd = gen_overloading Syntax.read_term;
--- a/src/Pure/Isar/toplevel.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Isar/toplevel.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -718,7 +718,7 @@
 
         val future_proof = Proof.future_proof
           (fn prf =>
-            Future.fork_pri 1 (fn () =>
+            Future.fork_pri ~1 (fn () =>
               let val (states, State (result_node, _)) =
                 (case st' of State (SOME (Proof (_, (_, orig_gthy)), exit), prev)
                   => State (SOME (Proof (ProofNode.init prf, (finish, orig_gthy)), exit), prev))
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/value_parse.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,45 @@
+(*  Title:      Pure/Isar/value_parse.ML
+    Author:     Makarius
+
+Outer syntax parsers for basic ML values.
+*)
+
+signature VALUE_PARSE =
+sig
+  type 'a parser = 'a OuterParse.parser
+  val comma: 'a parser -> 'a parser
+  val equal: 'a parser -> 'a parser
+  val parens: 'a parser -> 'a parser
+  val pair: 'a parser -> 'b parser -> ('a * 'b) parser
+  val triple: 'a parser -> 'b parser -> 'c parser -> ('a * 'b * 'c) parser
+  val properties: Properties.T parser
+end;
+
+structure ValueParse: VALUE_PARSE =
+struct
+
+structure P = OuterParse;
+type 'a parser = 'a P.parser;
+
+
+(* syntax utilities *)
+
+fun comma p = P.$$$ "," |-- P.!!! p;
+fun equal p = P.$$$ "=" |-- P.!!! p;
+fun parens p = P.$$$ "(" |-- P.!!! (p --| P.$$$ ")");
+
+
+(* tuples *)
+
+val unit = parens (Scan.succeed ());
+fun pair p1 p2 = parens (p1 -- comma p2);
+fun triple p1 p2 p3 = parens (p1 -- comma p2 -- comma p3) >> P.triple1;
+
+
+(* lists *)
+
+fun list p = parens (P.enum "," p);
+val properties = list (P.string -- equal P.string);
+
+end;
+
--- a/src/Pure/ML-Systems/mosml.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/ML-Systems/mosml.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,14 +1,13 @@
 (*  Title:      Pure/ML-Systems/mosml.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1999  University of Cambridge
+    Author:     Makarius
 
 Compatibility file for Moscow ML 2.01
 
 NOTE: saving images does not work; may run it interactively as follows:
 
 $ cd Isabelle/src/Pure
-$ isabelle RAW_ML_SYSTEM
+$ isabelle-process RAW_ML_SYSTEM
 > val ml_system = "mosml";
 > use "ML-Systems/mosml.ML";
 > use "ROOT.ML";
@@ -27,6 +26,7 @@
 fun forget_structure _ = ();
 
 load "Obj";
+load "Word";
 load "Bool";
 load "Int";
 load "Real";
@@ -35,6 +35,7 @@
 load "Process";
 load "FileSys";
 load "IO";
+load "CharVector";
 
 exception Interrupt;
 
@@ -54,6 +55,14 @@
 structure IntInf =
 struct
   fun divMod (x, y) = (x div y, x mod y);
+
+  local
+    fun log 1 a = a
+      | log n a = log (n div 2) (a + 1);
+  in
+    fun log2 n = if n > 0 then log n 0 else raise Domain;
+  end;
+
   open Int;
 end;
 
--- a/src/Pure/Proof/extraction.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Proof/extraction.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/extraction.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Extraction of programs from proofs.
@@ -549,7 +548,7 @@
           let
             val prf = force_proof body;
             val (vs', tye) = find_inst prop Ts ts vs;
-            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye;
+            val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye;
             val T = etype_of thy' vs' [] prop;
             val defs' = if T = nullT then defs
               else fst (extr d defs vs ts Ts hs prf0)
@@ -570,7 +569,7 @@
                     val corr_prf' = List.foldr forall_intr_prf
                       (proof_combt
                          (PThm (serial (),
-                          ((corr_name name vs', corr_prop, SOME (map TVar (term_tvars corr_prop))),
+                          ((corr_name name vs', corr_prop, SOME (map TVar (OldTerm.term_tvars corr_prop))),
                             Lazy.value (make_proof_body corr_prf))), vfs_of corr_prop))
                       (map (get_var_type corr_prop) (vfs_of prop))
                   in
@@ -588,7 +587,7 @@
       | corr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) _ _ =
           let
             val (vs', tye) = find_inst prop Ts ts vs;
-            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
+            val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye
           in
             if etype_of thy' vs' [] prop = nullT andalso
               realizes_null vs' prop aconv prop then (defs, prf0)
@@ -639,7 +638,7 @@
           let
             val prf = force_proof body;
             val (vs', tye) = find_inst prop Ts ts vs;
-            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
+            val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye
           in
             case Symtab.lookup realizers s of
               NONE => (case find vs' (find' s defs) of
@@ -676,12 +675,12 @@
                          (chtype [propT, T] combination_axm %> f %> f %> c %> t' %%
                            (chtype [T --> propT] reflexive_axm %> f) %%
                            PAxm (cname ^ "_def", eqn,
-                             SOME (map TVar (term_tvars eqn))))) %% corr_prf;
+                             SOME (map TVar (OldTerm.term_tvars eqn))))) %% corr_prf;
                     val corr_prop = Reconstruct.prop_of corr_prf';
                     val corr_prf'' = List.foldr forall_intr_prf
                       (proof_combt
                         (PThm (serial (),
-                         ((corr_name s vs', corr_prop, SOME (map TVar (term_tvars corr_prop))),
+                         ((corr_name s vs', corr_prop, SOME (map TVar (OldTerm.term_tvars corr_prop))),
                            Lazy.value (make_proof_body corr_prf'))), vfs_of corr_prop))
                       (map (get_var_type corr_prop) (vfs_of prop));
                   in
@@ -699,7 +698,7 @@
       | extr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) =
           let
             val (vs', tye) = find_inst prop Ts ts vs;
-            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
+            val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye
           in
             case find vs' (Symtab.lookup_list realizers s) of
               SOME (t, _) => (defs, subst_TVars tye' t)
@@ -739,7 +738,7 @@
            in
              thy'
              |> PureThy.store_thm (corr_name s vs,
-                  Thm.varifyT (funpow (length (term_vars corr_prop))
+                  Thm.varifyT (funpow (length (OldTerm.term_vars corr_prop))
                     (Thm.forall_elim_var 0) (forall_intr_frees
                       (ProofChecker.thm_of_proof thy'
                        (fst (Proofterm.freeze_thaw_prf prf))))))
--- a/src/Pure/Proof/proof_rewrite_rules.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Proof/proof_rewrite_rules.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/proof_rewrite_rules.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Simplification functions for proof terms involving meta level rules.
@@ -196,7 +195,8 @@
   let
     fun rew_term Ts t =
       let
-        val frees = map Free (Name.invent_list (add_term_names (t, [])) "xa" (length Ts) ~~ Ts);
+        val frees =
+          map Free (Name.invent_list (OldTerm.add_term_names (t, [])) "xa" (length Ts) ~~ Ts);
         val t' = r (subst_bounds (frees, t));
         fun strip [] t = t
           | strip (_ :: xs) (Abs (_, _, t)) = strip xs t;
@@ -228,7 +228,7 @@
           if member (op =) defs s then
             let
               val vs = vars_of prop;
-              val tvars = term_tvars prop;
+              val tvars = OldTerm.term_tvars prop;
               val (_, rhs) = Logic.dest_equals prop;
               val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts)
                 (fold_rev (fn x => fn b => Abs ("", dummyT, abstract_over (x, b))) vs rhs),
@@ -249,7 +249,9 @@
         val cnames = map (fst o dest_Const o fst) defs';
         val thms = fold_proof_atoms true
           (fn PThm (_, ((name, prop, _), _)) =>
-              if member (op =) defnames name orelse null (term_consts prop inter cnames) then I
+              if member (op =) defnames name orelse
+                not (exists_Const (member (op =) cnames o #1) prop)
+              then I
               else cons (name, SOME prop)
             | _ => I) [prf] [];
       in Reconstruct.expand_proof thy thms end;
--- a/src/Pure/Proof/proofchecker.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Proof/proofchecker.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/proofchecker.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Simple proof checker based only on the core inference rules
@@ -31,13 +30,12 @@
 
 fun thm_of_proof thy prf =
   let
-    val prf_names = Proofterm.fold_proof_terms
-      (fold_aterms (fn Free (x, _) => Name.declare x | _ => I)) (K I) prf Name.context;
+    val prf_names = Proofterm.fold_proof_terms Term.declare_term_frees (K I) prf Name.context;
     val lookup = lookup_thm thy;
 
     fun thm_of_atom thm Ts =
       let
-        val tvars = term_tvars (Thm.full_prop_of thm);
+        val tvars = OldTerm.term_tvars (Thm.full_prop_of thm);
         val (fmap, thm') = Thm.varifyT' [] thm;
         val ctye = map (pairself (Thm.ctyp_of thy))
           (map TVar tvars @ map (fn ((_, S), ixn) => TVar (ixn, S)) fmap ~~ Ts)
--- a/src/Pure/Proof/reconstruct.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Proof/reconstruct.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/reconstruct.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Reconstruction of partial proof terms.
@@ -139,8 +138,8 @@
 
     fun mk_cnstrts_atom env vTs prop opTs prf =
           let
-            val tvars = term_tvars prop;
-            val tfrees = term_tfrees prop;
+            val tvars = OldTerm.term_tvars prop;
+            val tfrees = OldTerm.term_tfrees prop;
             val (env', Ts) = (case opTs of
                 NONE => foldl_map mk_tvar (env, map snd tvars @ map snd tfrees)
               | SOME Ts => (env, Ts));
@@ -232,9 +231,6 @@
 
 (**** update list of free variables of constraints ****)
 
-val add_term_ixns = fold_aterms (fn Var (xi, _) => insert (op =) xi | _ => I);
-val add_typ_ixns = fold_atyps (fn TVar (ai, _) => insert (op =) ai | _ => I);
-
 fun upd_constrs env cs =
   let
     val Envir.Envir {asol, iTs, ...} = env;
@@ -242,8 +238,8 @@
       |> Vartab.fold (cons o #1) asol
       |> Vartab.fold (cons o #1) iTs;
     val vran = []
-      |> Vartab.fold (add_term_ixns o #2 o #2) asol
-      |> Vartab.fold (add_typ_ixns o #2 o #2) iTs;
+      |> Vartab.fold (Term.add_var_names o #2 o #2) asol
+      |> Vartab.fold (Term.add_tvar_namesT o #2 o #2) iTs;
     fun check_cs [] = []
       | check_cs ((u, p, vs)::ps) =
           let val vs' = subtract (op =) dom vs;
@@ -294,7 +290,7 @@
     val (t, prf, cs, env, _) = make_constraints_cprf thy
       (Envir.empty (maxidx_proof cprf ~1)) cprf';
     val cs' = map (fn p => (true, p, op union
-        (pairself (map (fst o dest_Var) o term_vars) p)))
+        (pairself (map (fst o dest_Var) o OldTerm.term_vars) p)))
       (map (pairself (Envir.norm_term env)) ((t, prop')::cs));
     val _ = message ("Solving remaining constraints (" ^ string_of_int (length cs') ^ ") ...");
     val env' = solve thy cs' env
@@ -303,7 +299,7 @@
   end;
 
 fun prop_of_atom prop Ts = subst_atomic_types
-  (map TVar (term_tvars prop) @ map TFree (term_tfrees prop) ~~ Ts)
+  (map TVar (OldTerm.term_tvars prop) @ map TFree (OldTerm.term_tfrees prop) ~~ Ts)
   (forall_intr_vfs prop);
 
 val head_norm = Envir.head_norm (Envir.empty 0);
@@ -370,9 +366,9 @@
                   end
               | SOME (maxidx', prf) => (maxidx' + maxidx + 1,
                   inc (maxidx + 1) prf, prfs));
-            val tfrees = term_tfrees prop;
+            val tfrees = OldTerm.term_tfrees prop;
             val tye = map (fn ((s, j), _) => (s, maxidx + 1 + j))
-              (term_tvars prop) @ map (rpair ~1 o fst) tfrees ~~ Ts;
+              (OldTerm.term_tvars prop) @ map (rpair ~1 o fst) tfrees ~~ Ts;
             val varify = map_type_tfree (fn p as (a, S) =>
               if member (op =) tfrees p then TVar ((a, ~1), S) else TFree p)
           in
--- a/src/Pure/ProofGeneral/pgip_parser.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/ProofGeneral/pgip_parser.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -92,18 +92,18 @@
 
 fun parse_span span =
   let
-    val kind = ThyEdit.span_kind span;
-    val toks = ThyEdit.span_content span;
-    val text = implode (map (PrintMode.setmp [] ThyEdit.present_token) toks);
+    val kind = ThySyntax.span_kind span;
+    val toks = ThySyntax.span_content span;
+    val text = implode (map (PrintMode.setmp [] ThySyntax.present_token) toks);
   in
     (case kind of
-      ThyEdit.Command name => parse_command name text
-    | ThyEdit.Ignored => [D.Whitespace {text = text}]
-    | ThyEdit.Malformed => [D.Unparseable {text = text}])
+      ThySyntax.Command name => parse_command name text
+    | ThySyntax.Ignored => [D.Whitespace {text = text}]
+    | ThySyntax.Malformed => [D.Unparseable {text = text}])
   end;
 
 
 fun pgip_parser pos str =
-  maps parse_span (ThyEdit.parse_spans (OuterKeyword.get_lexicons ()) pos str);
+  maps parse_span (ThySyntax.parse_spans (OuterKeyword.get_lexicons ()) pos str);
 
 end;
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/proof_general_emacs.ML
-    ID:         $Id$
     Author:     David Aspinall and Markus Wenzel
 
 Isabelle/Isar configuration for Emacs Proof General.
@@ -21,81 +20,77 @@
 (* print modes *)
 
 val proof_generalN = "ProofGeneralEmacs";  (*token markup (colouring vars, etc.)*)
-val pgasciiN = "PGASCII";                  (*plain 7-bit ASCII communication*)
 val thm_depsN = "thm_deps";                (*meta-information about theorem deps*)
 val test_markupN = "test_markup";          (*XML markup for everything*)
 
-val _ = Markup.add_mode test_markupN (fn (name, props) =>
-  if name = Markup.promptN then ("", "") else XML.output_markup (name, props));
-
-fun special oct =
-  if print_mode_active pgasciiN then Symbol.SOH ^ chr (ord (oct_char oct) - 167)
-  else oct_char oct;
+fun special ch = Symbol.SOH ^ ch;
 
 
-(* text output: print modes for xsymbol *)
+(* render markup *)
 
 local
 
-fun xsym_output "\\" = "\\\\"
-  | xsym_output s = if Symbol.is_raw s then Symbol.decode_raw s else s;
-
-fun xsymbols_output s =
-  if print_mode_active Symbol.xsymbolsN andalso exists_string (fn c => c = "\\") s then
-    let val syms = Symbol.explode s
-    in (implode (map xsym_output syms), Symbol.length syms) end
-  else Output.default_output s;
+fun render_trees ts = fold render_tree ts
+and render_tree (XML.Text s) = Buffer.add s
+  | render_tree (XML.Elem (name, props, ts)) =
+      let
+        val (bg1, en1) =
+          if name <> Markup.promptN andalso print_mode_active test_markupN
+          then XML.output_markup (name, props)
+          else Markup.no_output;
+        val (bg2, en2) =
+          if (case ts of [XML.Text _] => false | _ => true) then Markup.no_output
+          else if name = Markup.stateN then (special "O" ^ "\n", "\n" ^ special "P")
+          else if name = Markup.sendbackN then (special "W", special "X")
+          else if name = Markup.hiliteN then (special "0", special "1")
+          else if name = Markup.tclassN then (special "B", special "A")
+          else if name = Markup.tfreeN then (special "C", special "A")
+          else if name = Markup.tvarN then (special "D", special "A")
+          else if name = Markup.freeN then (special "E", special "A")
+          else if name = Markup.boundN then (special "F", special "A")
+          else if name = Markup.varN then (special "G", special "A")
+          else if name = Markup.skolemN then (special "H", special "A")
+          else Markup.no_output;
+      in
+        Buffer.add bg1 #>
+        Buffer.add bg2 #>
+        render_trees ts #>
+        Buffer.add en2 #>
+        Buffer.add en1
+      end;
 
 in
 
-fun setup_xsymbols_output () =
-  Output.add_mode Symbol.xsymbolsN xsymbols_output Symbol.encode_raw;
+fun render text =
+  Buffer.content (render_trees (YXML.parse_body text) Buffer.empty);
 
 end;
 
 
-(* common markup *)
+(* messages *)
+
+fun message bg en prfx text =
+  (case render text of
+    "" => ()
+  | s => Output.writeln_default (enclose bg en (prefix_lines prfx s)));
 
-val _ = Markup.add_mode proof_generalN (fn (name, props) =>
-  let
-    val (bg1, en1) =
-      if name = Markup.stateN then (special "366" ^ "\n", "\n" ^ special "367")
-      else if name = Markup.sendbackN then (special "376", special "377")
-      else if name = Markup.hiliteN then (special "327", special "330")
-      else if name = Markup.tclassN then (special "351", special "350")
-      else if name = Markup.tfreeN then (special "352", special "350")
-      else if name = Markup.tvarN then (special "353", special "350")
-      else if name = Markup.freeN then (special "354", special "350")
-      else if name = Markup.boundN then (special "355", special "350")
-      else if name = Markup.varN then (special "356", special "350")
-      else if name = Markup.skolemN then (special "357", special "350")
-      else ("", "");
-    val (bg2, en2) =
-      if name <> Markup.promptN andalso print_mode_active test_markupN
-      then XML.output_markup (name, props)
-      else ("", "");
-  in (bg1 ^ bg2, en2 ^ en1) end);
+fun setup_messages () =
+ (Output.writeln_fn := message "" "" "";
+  Output.status_fn := (fn s => ! Output.priority_fn s);
+  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") "+++ ";
+  Output.warning_fn := message (special "K") (special "L") "### ";
+  Output.error_fn := message (special "M") (special "N") "*** ";
+  Output.prompt_fn := (fn s => Output.std_output (render s ^ special "S")));
+
+fun panic s =
+  (message (special "M") (special "N") "!!! " ("## SYSTEM EXIT ##\n" ^ s); exit 1);
 
 
-(* messages and notification *)
-
-fun decorate bg en prfx =
-  Output.writeln_default o enclose bg en o prefix_lines prfx;
+(* notification *)
 
-fun setup_messages () =
- (Output.writeln_fn := Output.writeln_default;
-  Output.status_fn := (fn "" => () | s => ! Output.priority_fn s);
-  Output.priority_fn := (fn s => decorate (special "360") (special "361") "" s);
-  Output.tracing_fn := (fn s => decorate (special "360" ^ special "375") (special "361") "" s);
-  Output.debug_fn := (fn s => decorate (special "362") (special "363") "+++ " s);
-  Output.warning_fn := (fn s => decorate (special "362") (special "363") "### " s);
-  Output.error_fn := (fn s => decorate (special "364") (special "365") "*** " s);
-  Output.prompt_fn := (fn s => Output.std_output (s ^ special "372")));
-
-fun panic s =
-  (decorate (special "364") (special "365") "!!! " ("## SYSTEM EXIT ##\n" ^ s); exit 1);
-
-fun emacs_notify s = decorate (special "360") (special "361") "" s;
+val emacs_notify = message (special "I") (special "J") "";
 
 fun tell_clear_goals () =
   emacs_notify "Proof General, please clear the goals buffer.";
@@ -149,7 +144,7 @@
             (warning (cat_lines ["Failed to register theory: " ^ quote name, msg]);
               tell_file_retracted (ThyLoad.thy_path name))
       else ();
-    val _ = Isar.init_point ();
+    val _ = Isar.init ();
   in () end;
 
 
@@ -161,7 +156,7 @@
  (sync_thy_loader ();
   tell_clear_goals ();
   tell_clear_response ();
-  Isar.init_point ();
+  Isar.init ();
   welcome ());
 
 
@@ -237,7 +232,9 @@
   | init true =
       (! initialized orelse
         (Output.no_warnings init_outer_syntax ();
-          setup_xsymbols_output ();
+          Output.add_mode Symbol.xsymbolsN Symbol.output Output.default_escape;
+          Output.add_mode proof_generalN Output.default_output Output.default_escape;
+          Markup.add_mode proof_generalN YXML.output_markup;
           setup_messages ();
           ProofGeneralPgip.init_pgip_channel (! Output.priority_fn);
           setup_thy_loader ();
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -242,7 +242,7 @@
     (sync_thy_loader ();
      tell_clear_goals ();
      tell_clear_response ();
-     Isar.init_point ();
+     Isar.init ();
      welcome ());
 
 
--- a/src/Pure/ROOT.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/ROOT.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -25,7 +25,9 @@
 (*fundamental structures*)
 use "name.ML";
 use "term.ML";
+use "term_ord.ML";
 use "term_subst.ML";
+use "old_term.ML";
 use "logic.ML";
 use "General/pretty.ML";
 use "context.ML";
--- a/src/Pure/Syntax/lexicon.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Syntax/lexicon.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -169,10 +169,10 @@
   | VarSy       => Markup.var
   | TFreeSy     => Markup.tfree
   | TVarSy      => Markup.tvar
-  | NumSy       => Markup.num
-  | FloatSy     => Markup.float
-  | XNumSy      => Markup.xnum
-  | StrSy       => Markup.xstr
+  | NumSy       => Markup.numeral
+  | FloatSy     => Markup.numeral
+  | XNumSy      => Markup.numeral
+  | StrSy       => Markup.inner_string
   | Space       => Markup.none
   | Comment     => Markup.inner_comment
   | EOF         => Markup.none;
--- a/src/Pure/Syntax/syn_ext.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Syntax/syn_ext.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -379,7 +379,7 @@
 
 fun stamp_trfun s (c, f) = (c, (f, s));
 fun mk_trfun tr = stamp_trfun (stamp ()) tr;
-fun eq_trfun ((_, s1:stamp), (_, s2)) = s1 = s2;
+fun eq_trfun ((_, s1: stamp), (_, s2)) = s1 = s2;
 
 
 (* token translations *)
@@ -387,7 +387,7 @@
 fun tokentrans_mode m trs = map (fn (s, f) => (m, s, f)) trs;
 
 val standard_token_classes =
-  ["class", "tfree", "tvar", "free", "bound", "var", "num", "float", "xnum", "xstr"];
+  ["class", "tfree", "tvar", "free", "bound", "var", "numeral", "inner_string"];
 
 val standard_token_markers = map (fn s => "_" ^ s) standard_token_classes;
 
--- a/src/Pure/Syntax/syn_trans.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Syntax/syn_trans.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/syn_trans.ML
-    ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 
 Syntax translation functions.
@@ -264,7 +263,7 @@
   let
     val vars = vars_of tm;
     val body = body_of tm;
-    val rev_new_vars = rename_wrt_term body vars;
+    val rev_new_vars = Term.rename_wrt_term body vars;
     fun subst (x, T) b =
       if can Name.dest_internal x andalso not (Term.loose_bvar1 (b, 0))
       then (Const ("_idtdummy", T), incr_boundvars ~1 b)
@@ -302,7 +301,7 @@
     (strip_abss strip_abs_vars strip_abs_body (eta_contr tm));
 
 fun atomic_abs_tr' (x, T, t) =
-  let val [xT] = rename_wrt_term t [(x, T)]
+  let val [xT] = Term.rename_wrt_term t [(x, T)]
   in (mark_boundT xT, subst_bound (mark_bound (fst xT), t)) end;
 
 fun abs_ast_tr' (*"_abs"*) asts =
--- a/src/Pure/Thy/latex.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Thy/latex.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -174,7 +174,7 @@
 fun latex_markup (s, _) =
   if s = Markup.keywordN then ("\\isakeyword{", "}")
   else if s = Markup.commandN then ("\\isacommand{", "}")
-  else ("", "");
+  else Markup.no_output;
 
 fun latex_indent "" _ = ""
   | latex_indent s _ = enclose "\\isaindent{" "}" s;
--- a/src/Pure/Thy/thy_edit.ML	Mon Dec 29 11:04:27 2008 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,194 +0,0 @@
-(*  Title:      Pure/Thy/thy_edit.ML
-    ID:         $Id$
-    Author:     Makarius
-
-Basic editor operations on theory sources.
-*)
-
-signature THY_EDIT =
-sig
-  val token_source: Scan.lexicon * Scan.lexicon -> Position.T -> (string, 'a) Source.source ->
-    (OuterLex.token, (SymbolPos.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
-  val report_token: OuterLex.token -> unit
-  datatype span_kind = Command of string | Ignored | Malformed
-  type span
-  val span_kind: span -> span_kind
-  val span_content: span -> OuterLex.token list
-  val span_range: span -> Position.range
-  val span_source: (OuterLex.token, 'a) Source.source ->
-    (span, (OuterLex.token, 'a) Source.source) Source.source
-  val parse_spans: Scan.lexicon * Scan.lexicon -> Position.T -> string -> span list
-  val present_span: span -> output
-  val report_span: span -> unit
-  val unit_source: (span, 'a) Source.source ->
-    (span * span list * bool, (span, 'a) Source.source) Source.source
-end;
-
-structure ThyEdit: THY_EDIT =
-struct
-
-structure K = OuterKeyword;
-structure T = OuterLex;
-structure P = OuterParse;
-
-
-(** tokens **)
-
-(* parse *)
-
-fun token_source lexs pos src =
-  Symbol.source {do_recover = true} src
-  |> T.source {do_recover = SOME false} (K lexs) pos;
-
-fun parse_tokens lexs pos str =
-  Source.of_string str
-  |> token_source lexs pos
-  |> Source.exhaust;
-
-
-(* present *)
-
-local
-
-val token_kind_markup =
- fn T.Command       => (Markup.commandN, [])
-  | T.Keyword       => (Markup.keywordN, [])
-  | T.Ident         => Markup.ident
-  | T.LongIdent     => Markup.ident
-  | T.SymIdent      => Markup.ident
-  | T.Var           => Markup.ident
-  | T.TypeIdent     => Markup.ident
-  | T.TypeVar       => Markup.ident
-  | T.Nat           => Markup.ident
-  | T.String        => Markup.string
-  | T.AltString     => Markup.altstring
-  | T.Verbatim      => Markup.verbatim
-  | T.Space         => Markup.none
-  | T.Comment       => Markup.comment
-  | T.InternalValue => Markup.none
-  | T.Malformed     => Markup.malformed
-  | T.Error _       => Markup.malformed
-  | T.Sync          => Markup.control
-  | T.EOF           => Markup.control;
-
-in
-
-fun present_token tok =
-  Markup.enclose (token_kind_markup (T.kind_of tok)) (Output.output (T.unparse tok));
-
-fun report_token tok =
-  Position.report (token_kind_markup (T.kind_of tok)) (T.position_of tok);
-
-end;
-
-
-
-(** spans **)
-
-(* type span *)
-
-datatype span_kind = Command of string | Ignored | Malformed;
-datatype span = Span of span_kind * OuterLex.token list;
-
-fun span_kind (Span (k, _)) = k;
-fun span_content (Span (_, toks)) = toks;
-
-fun span_range span =
-  (case span_content span of
-    [] => (Position.none, Position.none)
-  | toks =>
-      let
-        val start_pos = T.position_of (hd toks);
-        val end_pos = T.end_position_of (List.last toks);
-      in (start_pos, end_pos) end);
-
-
-(* parse *)
-
-local
-
-val is_whitespace = T.is_kind T.Space orf T.is_kind T.Comment;
-
-val body = Scan.unless (Scan.many is_whitespace -- Scan.ahead (P.command || P.eof)) P.not_eof;
-
-val span =
-  Scan.ahead P.command -- P.not_eof -- Scan.repeat body
-    >> (fn ((name, c), bs) => Span (Command name, c :: bs)) ||
-  Scan.many1 is_whitespace >> (fn toks => Span (Ignored, toks)) ||
-  Scan.repeat1 body >> (fn toks => Span (Malformed, toks));
-
-in
-
-fun span_source src = Source.source T.stopper (Scan.bulk span) NONE src;
-
-end;
-
-fun parse_spans lexs pos str =
-  Source.of_string str
-  |> token_source lexs pos
-  |> span_source
-  |> Source.exhaust;
-
-
-(* present *)
-
-local
-
-fun kind_markup (Command name) = Markup.command_span name
-  | kind_markup Ignored = Markup.ignored_span
-  | kind_markup Malformed = Markup.malformed_span;
-
-in
-
-fun present_span span =
-  Markup.enclose (kind_markup (span_kind span)) (implode (map present_token (span_content span)));
-
-fun report_span span =
-  Position.report (kind_markup (span_kind span)) (Position.encode_range (span_range span));
-
-end;
-
-
-
-(** units: commands with proof **)
-
-(* scanning spans *)
-
-val eof = Span (Command "", []);
-
-fun is_eof (Span (Command "", _)) = true
-  | is_eof _ = false;
-
-val not_eof = not o is_eof;
-
-val stopper = Scan.stopper (K eof) is_eof;
-
-
-(* unit_source *)
-
-local
-
-fun command_with pred = Scan.one (fn (Span (Command name, _)) => pred name | _ => false);
-
-val proof = Scan.pass 1 (Scan.repeat (Scan.depend (fn d =>
-  if d <= 0 then Scan.fail
-  else
-    command_with K.is_qed_global >> pair ~1 ||
-    command_with K.is_proof_goal >> pair (d + 1) ||
-    (if d = 0 then Scan.fail else command_with K.is_qed >> pair (d - 1)) ||
-    Scan.unless (command_with K.is_theory) (Scan.one not_eof) >> pair d)) -- Scan.state);
-
-val unit =
-  command_with K.is_theory_goal -- proof >> (fn (a, (bs, d)) => (a, bs, d >= 0)) ||
-  Scan.one not_eof >> (fn a => (a, [], true));
-
-in
-
-fun unit_source src = Source.source stopper (Scan.bulk unit) NONE src;
-
-end;
-
-end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Thy/thy_syntax.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,193 @@
+(*  Title:      Pure/Thy/thy_syntax.ML
+    Author:     Makarius
+
+Superficial theory syntax: tokens and spans.
+*)
+
+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)
+      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
+  val report_token: OuterLex.token -> unit
+  datatype span_kind = Command of string | Ignored | Malformed
+  type span
+  val span_kind: span -> span_kind
+  val span_content: span -> OuterLex.token list
+  val span_range: span -> Position.range
+  val span_source: (OuterLex.token, 'a) Source.source ->
+    (span, (OuterLex.token, 'a) Source.source) Source.source
+  val parse_spans: Scan.lexicon * Scan.lexicon -> Position.T -> string -> span list
+  val present_span: span -> output
+  val report_span: span -> unit
+  val unit_source: (span, 'a) Source.source ->
+    (span * span list * bool, (span, 'a) Source.source) Source.source
+end;
+
+structure ThySyntax: THY_SYNTAX =
+struct
+
+structure K = OuterKeyword;
+structure T = OuterLex;
+structure P = OuterParse;
+
+
+(** tokens **)
+
+(* parse *)
+
+fun token_source lexs pos src =
+  Symbol.source {do_recover = true} src
+  |> T.source {do_recover = SOME false} (K lexs) pos;
+
+fun parse_tokens lexs pos str =
+  Source.of_string str
+  |> token_source lexs pos
+  |> Source.exhaust;
+
+
+(* present *)
+
+local
+
+val token_kind_markup =
+ fn T.Command       => (Markup.commandN, [])
+  | T.Keyword       => (Markup.keywordN, [])
+  | T.Ident         => Markup.ident
+  | T.LongIdent     => Markup.ident
+  | T.SymIdent      => Markup.ident
+  | T.Var           => Markup.var
+  | T.TypeIdent     => Markup.tfree
+  | T.TypeVar       => Markup.tvar
+  | T.Nat           => Markup.ident
+  | T.String        => Markup.string
+  | T.AltString     => Markup.altstring
+  | T.Verbatim      => Markup.verbatim
+  | T.Space         => Markup.none
+  | T.Comment       => Markup.comment
+  | T.InternalValue => Markup.none
+  | T.Malformed     => Markup.malformed
+  | T.Error _       => Markup.malformed
+  | T.Sync          => Markup.control
+  | T.EOF           => Markup.control;
+
+in
+
+fun present_token tok =
+  Markup.enclose (token_kind_markup (T.kind_of tok)) (Output.output (T.unparse tok));
+
+fun report_token tok =
+  Position.report (token_kind_markup (T.kind_of tok)) (T.position_of tok);
+
+end;
+
+
+
+(** spans **)
+
+(* type span *)
+
+datatype span_kind = Command of string | Ignored | Malformed;
+datatype span = Span of span_kind * OuterLex.token list;
+
+fun span_kind (Span (k, _)) = k;
+fun span_content (Span (_, toks)) = toks;
+
+fun span_range span =
+  (case span_content span of
+    [] => (Position.none, Position.none)
+  | toks =>
+      let
+        val start_pos = T.position_of (hd toks);
+        val end_pos = T.end_position_of (List.last toks);
+      in (start_pos, end_pos) end);
+
+
+(* parse *)
+
+local
+
+val is_whitespace = T.is_kind T.Space orf T.is_kind T.Comment;
+
+val body = Scan.unless (Scan.many is_whitespace -- Scan.ahead (P.command || P.eof)) P.not_eof;
+
+val span =
+  Scan.ahead P.command -- P.not_eof -- Scan.repeat body
+    >> (fn ((name, c), bs) => Span (Command name, c :: bs)) ||
+  Scan.many1 is_whitespace >> (fn toks => Span (Ignored, toks)) ||
+  Scan.repeat1 body >> (fn toks => Span (Malformed, toks));
+
+in
+
+fun span_source src = Source.source T.stopper (Scan.bulk span) NONE src;
+
+end;
+
+fun parse_spans lexs pos str =
+  Source.of_string str
+  |> token_source lexs pos
+  |> span_source
+  |> Source.exhaust;
+
+
+(* present *)
+
+local
+
+fun kind_markup (Command name) = Markup.command_span name
+  | kind_markup Ignored = Markup.ignored_span
+  | kind_markup Malformed = Markup.malformed_span;
+
+in
+
+fun present_span span =
+  Markup.enclose (kind_markup (span_kind span)) (implode (map present_token (span_content span)));
+
+fun report_span span =
+  Position.report (kind_markup (span_kind span)) (Position.encode_range (span_range span));
+
+end;
+
+
+
+(** units: commands with proof **)
+
+(* scanning spans *)
+
+val eof = Span (Command "", []);
+
+fun is_eof (Span (Command "", _)) = true
+  | is_eof _ = false;
+
+val not_eof = not o is_eof;
+
+val stopper = Scan.stopper (K eof) is_eof;
+
+
+(* unit_source *)
+
+local
+
+fun command_with pred = Scan.one (fn (Span (Command name, _)) => pred name | _ => false);
+
+val proof = Scan.pass 1 (Scan.repeat (Scan.depend (fn d =>
+  if d <= 0 then Scan.fail
+  else
+    command_with K.is_qed_global >> pair ~1 ||
+    command_with K.is_proof_goal >> pair (d + 1) ||
+    (if d = 0 then Scan.fail else command_with K.is_qed >> pair (d - 1)) ||
+    Scan.unless (command_with K.is_theory) (Scan.one not_eof) >> pair d)) -- Scan.state);
+
+val unit =
+  command_with K.is_theory_goal -- proof >> (fn (a, (bs, d)) => (a, bs, d >= 0)) ||
+  Scan.one not_eof >> (fn a => (a, [], true));
+
+in
+
+fun unit_source src = Source.source stopper (Scan.bulk unit) NONE src;
+
+end;
+
+end;
--- a/src/Pure/Tools/isabelle_process.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Tools/isabelle_process.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,15 +1,14 @@
 (*  Title:      Pure/Tools/isabelle_process.ML
-    ID:         $Id$
     Author:     Makarius
 
 Isabelle process wrapper -- interaction via external program.
 
 General format of process output:
 
-  (a) unmarked stdout/stderr, no line structure (output should be
+  (1) unmarked stdout/stderr, no line structure (output should be
   processed immediately as it arrives);
 
-  (b) properly marked-up messages, e.g. for writeln channel
+  (2) properly marked-up messages, e.g. for writeln channel
 
   "\002A" ^ props ^ "\002,\n" ^ text ^ "\002.\n"
 
@@ -17,13 +16,14 @@
   each, and the remaining text is any number of lines (output is
   supposed to be processed in one piece);
 
-  (c) special init message holds "pid" and "session" property.
+  (3) special init message holds "pid" and "session" property;
+
+  (4) message content is encoded in YXML format.
 *)
 
 signature ISABELLE_PROCESS =
 sig
   val isabelle_processN: string
-  val xmlN: string
   val init: string -> unit
 end;
 
@@ -33,7 +33,6 @@
 (* print modes *)
 
 val isabelle_processN = "isabelle_process";
-val xmlN = "XML";
 
 val _ = Output.add_mode isabelle_processN Output.default_output Output.default_escape;
 val _ = Markup.add_mode isabelle_processN YXML.output_markup;
@@ -56,19 +55,14 @@
   let val clean = clean_string [Symbol.STX, "\n", "\r"]
   in implode (map (fn (x, y) => clean x ^ "=" ^ clean y ^ special_sep ^ "\n") props) end;
 
-fun message_text class ts =
-  let
-    val doc = XML.Elem (Markup.messageN, [(Markup.classN, class)], ts);
-    val msg =
-      if print_mode_active xmlN then XML.header ^ XML.string_of doc
-      else YXML.string_of doc;
-  in clean_string [Symbol.STX] msg end;
+fun message_text class body =
+  YXML.element Markup.messageN [(Markup.classN, class)] [clean_string [Symbol.STX] body];
 
-fun message_pos ts = get_first get_pos ts
-and get_pos (elem as XML.Elem (name, atts, ts)) =
-      if name = Markup.positionN then SOME (Position.of_properties atts)
-      else message_pos ts
-  | get_pos _ = NONE;
+fun message_pos trees = trees |> get_first
+  (fn XML.Elem (name, atts, ts) =>
+        if name = Markup.positionN then SOME (Position.of_properties atts)
+        else message_pos ts
+    | _ => NONE);
 
 fun output out_stream s = NAMED_CRITICAL "IO" (fn () =>
   (TextIO.output (out_stream, s); TextIO.output (out_stream, "\n")));
@@ -78,19 +72,18 @@
 fun message _ _ _ "" = ()
   | message out_stream ch class body =
       let
-        val (txt, pos) =
-          let val ts = YXML.parse_body body
-          in (message_text class ts, the_default Position.none (message_pos ts)) end;
+        val pos = the_default Position.none (message_pos (YXML.parse_body body));
         val props =
           Position.properties_of (Position.thread_data ())
           |> Position.default_properties pos;
+        val txt = message_text class body;
       in output out_stream (special ch ^ message_props props ^ txt ^ special_end) end;
 
 fun init_message out_stream =
   let
     val pid = (Markup.pidN, process_id ());
     val session = (Markup.sessionN, List.last (Session.id ()) handle List.Empty => "unknown");
-    val text = message_text Markup.initN [XML.Text (Session.welcome ())];
+    val text = message_text Markup.initN (Session.welcome ());
   in output out_stream (special "A" ^ message_props [pid, session] ^ text ^ special_end) end;
 
 end;
--- a/src/Pure/Tools/isabelle_process.scala	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/Tools/isabelle_process.scala	Mon Jan 05 07:54:16 2009 -0800
@@ -88,7 +88,7 @@
   /* demo constructor */
 
   def this(args: String*) =
-    this(new IsabelleSystem, new EventBus[IsabelleProcess.Result] + (Console.println(_)), args: _*)
+    this(new IsabelleSystem, new EventBus[IsabelleProcess.Result] + Console.println, args: _*)
 
 
   /* process information */
--- a/src/Pure/codegen.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/codegen.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/codegen.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Generic code generator.
@@ -278,7 +277,7 @@
 
 fun preprocess_term thy t =
   let
-    val x = Free (Name.variant (add_term_names (t, [])) "x", fastype_of t);
+    val x = Free (Name.variant (OldTerm.add_term_names (t, [])) "x", fastype_of t);
     (* fake definition *)
     val eq = setmp quick_and_dirty true (SkipProof.make_thm thy)
       (Logic.mk_equals (x, t));
@@ -460,7 +459,7 @@
 
 fun rename_terms ts =
   let
-    val names = List.foldr add_term_names
+    val names = List.foldr OldTerm.add_term_names
       (map (fst o fst) (rev (fold Term.add_vars ts []))) ts;
     val reserved = filter ML_Syntax.is_reserved names;
     val (illegal, alt_names) = split_list (map_filter (fn s =>
@@ -485,7 +484,7 @@
 (**** retrieve definition of constant ****)
 
 fun is_instance T1 T2 =
-  Type.raw_instance (T1, if null (typ_tfrees T2) then T2 else Logic.varifyT T2);
+  Type.raw_instance (T1, if null (OldTerm.typ_tfrees T2) then T2 else Logic.varifyT T2);
 
 fun get_assoc_code thy (s, T) = Option.map snd (find_first (fn ((s', T'), _) =>
   s = s' andalso is_instance T T') (#consts (CodegenData.get thy)));
@@ -598,8 +597,8 @@
      else Pretty.block (separate (Pretty.brk 1) (p :: ps));
 
 fun new_names t xs = Name.variant_list
-  (map (fst o fst o dest_Var) (term_vars t) union
-    add_term_names (t, ML_Syntax.reserved_names)) (map mk_id xs);
+  (map (fst o fst o dest_Var) (OldTerm.term_vars t) union
+    OldTerm.add_term_names (t, ML_Syntax.reserved_names)) (map mk_id xs);
 
 fun new_name t x = hd (new_names t [x]);
 
@@ -917,9 +916,9 @@
     val ctxt = ProofContext.init thy;
     val e =
       let
-        val _ = (null (term_tvars t) andalso null (term_tfrees t)) orelse
+        val _ = (null (Term.add_tvars t []) andalso null (Term.add_tfrees t [])) orelse
           error "Term to be evaluated contains type variables";
-        val _ = (null (term_vars t) andalso null (term_frees t)) orelse
+        val _ = (null (Term.add_vars t []) andalso null (Term.add_frees t [])) orelse
           error "Term to be evaluated contains variables";
         val (code, gr) = setmp mode ["term_of"]
           (generate_code_i thy [] "Generated")
--- a/src/Pure/consts.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/consts.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -275,8 +275,8 @@
     val T = Term.fastype_of rhs;
     val lhs = Const (NameSpace.full_name naming b, T);
 
-    fun err msg = error (msg ^ " on rhs of abbreviation:\n" ^
-      Pretty.string_of_term pp (Logic.mk_equals (lhs, rhs)));
+    fun err msg = (warning (* FIXME should be error *) (msg ^ " on rhs of abbreviation:\n" ^
+      Pretty.string_of_term pp (Logic.mk_equals (lhs, rhs))); true);
     val _ = Term.exists_subterm Term.is_Var rhs andalso err "Illegal schematic variables";
     val _ = null (Term.hidden_polymorphism rhs) orelse err "Extra type variables";
   in
--- a/src/Pure/drule.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/drule.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      Pure/drule.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1993  University of Cambridge
 
 Derived rules and other operations on theorems.
 *)
@@ -112,6 +110,7 @@
   val sort_triv: theory -> typ * sort -> thm list
   val unconstrainTs: thm -> thm
   val with_subgoal: int -> (thm -> thm) -> thm -> thm
+  val comp_no_flatten: thm * int -> int -> thm -> thm
   val rename_bvars: (string * string) list -> thm -> thm
   val rename_bvars': string option list -> thm -> thm
   val incr_type_indexes: int -> thm -> thm
@@ -340,7 +339,7 @@
      val thy = Thm.theory_of_thm fth
      val {prop, tpairs, ...} = rep_thm fth
  in
-   case List.foldr add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of
+   case List.foldr OldTerm.add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of
        [] => (fth, fn i => fn x => x)   (*No vars: nothing to do!*)
      | vars =>
          let fun newName (Var(ix,_)) = (ix, gensym (string_of_indexname ix))
@@ -363,13 +362,13 @@
      val thy = Thm.theory_of_thm fth
      val {prop, tpairs, ...} = rep_thm fth
  in
-   case List.foldr add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of
+   case List.foldr OldTerm.add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of
        [] => (fth, fn x => x)
      | vars =>
          let fun newName (Var(ix,_), (pairs,used)) =
                    let val v = Name.variant used (string_of_indexname ix)
                    in  ((ix,v)::pairs, v::used)  end;
-             val (alist, _) = List.foldr newName ([], Library.foldr add_term_names
+             val (alist, _) = List.foldr newName ([], Library.foldr OldTerm.add_term_names
                (prop :: Thm.terms_of_tpairs tpairs, [])) vars
              fun mk_inst (Var(v,T)) =
                  (cterm_of thy (Var(v,T)),
@@ -807,8 +806,8 @@
 
 (*Increment the indexes of only the type variables*)
 fun incr_type_indexes inc th =
-  let val tvs = term_tvars (prop_of th)
-      and thy = theory_of_thm th
+  let val tvs = OldTerm.term_tvars (prop_of th)
+      and thy = Thm.theory_of_thm th
       fun inc_tvar ((a,i),s) = pairself (ctyp_of thy) (TVar ((a,i),s), TVar ((a,i+inc),s))
   in Thm.instantiate (map inc_tvar tvs, []) th end;
 
@@ -820,6 +819,14 @@
 fun th1 INCR_COMP th2 = incr_indexes th2 th1 COMP th2;
 fun th1 COMP_INCR th2 = th1 COMP incr_indexes th1 th2;
 
+fun comp_no_flatten (th, n) i rule =
+  (case distinct Thm.eq_thm (Seq.list_of
+      (Thm.compose_no_flatten false (th, n) i (incr_indexes th rule))) of
+    [th'] => th'
+  | [] => raise THM ("comp_no_flatten", i, [th, rule])
+  | _ => raise THM ("comp_no_flatten: unique result expected", i, [th, rule]));
+
+
 
 (*** Instantiate theorem th, reading instantiations in theory thy ****)
 
--- a/src/Pure/envir.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/envir.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      Pure/envir.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1988  University of Cambridge
 
 Environments.  The type of a term variable / sort of a type variable is
 part of its name. The lookup function must apply type substitutions,
@@ -118,7 +116,7 @@
 fun vupdate ((aU as (a, U), t), env as Envir {iTs, ...}) = case t of
       Var (nT as (name', T)) =>
         if a = name' then env     (*cycle!*)
-        else if Term.indexname_ord (a, name') = LESS then
+        else if TermOrd.indexname_ord (a, name') = LESS then
            (case lookup (env, nT) of  (*if already assigned, chase*)
                 NONE => update ((nT, Var (a, T)), env)
               | SOME u => vupdate ((aU, u), env))
--- a/src/Pure/facts.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/facts.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/facts.ML
-    ID:         $Id$
     Author:     Makarius
 
 Environment of named facts, optionally indexed by proposition.
@@ -166,7 +165,7 @@
 
 (* indexed props *)
 
-val prop_ord = Term.term_ord o pairself Thm.full_prop_of;
+val prop_ord = TermOrd.term_ord o pairself Thm.full_prop_of;
 
 fun props (Facts {props, ...}) = sort_distinct prop_ord (Net.content props);
 fun could_unify (Facts {props, ...}) = Net.unify_term props;
--- a/src/Pure/goal.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/goal.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*  Title:      Pure/goal.ML
-    ID:         $Id$
-    Author:     Makarius and Lawrence C Paulson
+    Author:     Makarius
 
 Goals in tactical theorem proving.
 *)
@@ -58,18 +57,14 @@
   --- (protect)
   #C
 *)
-fun protect th = th COMP_INCR Drule.protectI;
+fun protect th = Drule.comp_no_flatten (th, 0) 1 Drule.protectI;
 
 (*
   A ==> ... ==> #C
   ---------------- (conclude)
   A ==> ... ==> C
 *)
-fun conclude th =
-  (case SINGLE (Thm.compose_no_flatten false (th, Thm.nprems_of th) 1)
-      (Drule.incr_indexes th Drule.protectD) of
-    SOME th' => th'
-  | NONE => raise THM ("Failed to conclude goal", 0, [th]));
+fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD;
 
 (*
   #C
@@ -179,7 +174,7 @@
     val res =
       if immediate orelse #maxidx (Thm.rep_cterm stmt) >= 0 orelse not (Future.enabled ())
       then result ()
-      else future_result ctxt' (Future.fork_pri 1 result) (Thm.term_of stmt);
+      else future_result ctxt' (Future.fork_pri ~1 result) (Thm.term_of stmt);
   in
     Conjunction.elim_balanced (length props) res
     |> map (Assumption.export false ctxt' ctxt)
--- a/src/Pure/library.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/library.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/library.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Author:     Markus Wenzel, TU Muenchen
 
@@ -491,7 +490,7 @@
   | split_last [x] = ([], x)
   | split_last (x :: xs) = apfst (cons x) (split_last xs);
 
-(*find the position of an element in a list*)
+(*find position of first element satisfying a predicate*)
 fun find_index pred =
   let fun find (_: int) [] = ~1
         | find n (x :: xs) = if pred x then n else find (n + 1) xs;
--- a/src/Pure/logic.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/logic.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,6 @@
 (*  Title:      Pure/logic.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   Cambridge University 1992
+    Author:     Makarius
 
 Abstract syntax operations of the Pure meta-logic.
 *)
@@ -515,7 +514,7 @@
 (*reverses parameters for substitution*)
 fun goal_params st i =
   let val gi = get_goal st i
-      val rfrees = map Free (rename_wrt_term gi (strip_params gi))
+      val rfrees = map Free (Term.rename_wrt_term gi (strip_params gi))
   in (gi, rfrees) end;
 
 fun concl_of_goal st i =
--- a/src/Pure/meta_simplifier.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/meta_simplifier.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*  Title:      Pure/meta_simplifier.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow and Stefan Berghofer
+    Author:     Tobias Nipkow and Stefan Berghofer, TU Muenchen
 
 Meta-level Simplification.
 *)
@@ -788,7 +787,7 @@
   mk_eq_True = K NONE,
   reorient = default_reorient};
 
-val empty_ss = init_ss basic_mk_rews Term.termless (K (K no_tac)) ([], []);
+val empty_ss = init_ss basic_mk_rews TermOrd.termless (K (K no_tac)) ([], []);
 
 
 (* merge *)  (*NOTE: ignores some fields of 2nd simpset*)
--- a/src/Pure/more_thm.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/more_thm.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/more_thm.ML
-    ID:         $Id$
     Author:     Makarius
 
 Further operations on type ctyp/cterm/thm, outside the inference kernel.
@@ -128,12 +127,12 @@
     val {shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...} = Thm.rep_thm th1;
     val {shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...} = Thm.rep_thm th2;
   in
-    (case Term.fast_term_ord (prop1, prop2) of
+    (case TermOrd.fast_term_ord (prop1, prop2) of
       EQUAL =>
-        (case list_ord (prod_ord Term.fast_term_ord Term.fast_term_ord) (tpairs1, tpairs2) of
+        (case list_ord (prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord) (tpairs1, tpairs2) of
           EQUAL =>
-            (case list_ord Term.fast_term_ord (hyps1, hyps2) of
-              EQUAL => list_ord Term.sort_ord (shyps1, shyps2)
+            (case list_ord TermOrd.fast_term_ord (hyps1, hyps2) of
+              EQUAL => list_ord TermOrd.sort_ord (shyps1, shyps2)
             | ord => ord)
         | ord => ord)
     | ord => ord)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/old_term.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,94 @@
+(*  Title:      Pure/old_term.ML
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Some old-style term operations.
+*)
+
+signature OLD_TERM =
+sig
+  val it_term_types: (typ * 'a -> 'a) -> term * 'a -> 'a
+  val add_term_names: term * string list -> string list
+  val add_typ_tvars: typ * (indexname * sort) list -> (indexname * sort) list
+  val add_typ_tfree_names: typ * string list -> string list
+  val add_typ_tfrees: typ * (string * sort) list -> (string * sort) list
+  val add_term_tvars: term * (indexname * sort) list -> (indexname * sort) list
+  val add_term_tfrees: term * (string * sort) list -> (string * sort) list
+  val add_term_tfree_names: term * string list -> string list
+  val typ_tfrees: typ -> (string * sort) list
+  val typ_tvars: typ -> (indexname * sort) list
+  val term_tfrees: term -> (string * sort) list
+  val term_tvars: term -> (indexname * sort) list
+  val add_term_vars: term * term list -> term list
+  val term_vars: term -> term list
+  val add_term_frees: term * term list -> term list
+  val term_frees: term -> term list
+end;
+
+structure OldTerm: OLD_TERM =
+struct
+
+(*iterate a function over all types in a term*)
+fun it_term_types f =
+let fun iter(Const(_,T), a) = f(T,a)
+      | iter(Free(_,T), a) = f(T,a)
+      | iter(Var(_,T), a) = f(T,a)
+      | iter(Abs(_,T,t), a) = iter(t,f(T,a))
+      | iter(f$u, a) = iter(f, iter(u, a))
+      | iter(Bound _, a) = a
+in iter end
+
+(*Accumulates the names in the term, suppressing duplicates.
+  Includes Frees and Consts.  For choosing unambiguous bound var names.*)
+fun add_term_names (Const(a,_), bs) = insert (op =) (NameSpace.base a) bs
+  | add_term_names (Free(a,_), bs) = insert (op =) a bs
+  | add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs))
+  | add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs)
+  | add_term_names (_, bs) = bs;
+
+(*Accumulates the TVars in a type, suppressing duplicates.*)
+fun add_typ_tvars(Type(_,Ts),vs) = List.foldr add_typ_tvars vs Ts
+  | add_typ_tvars(TFree(_),vs) = vs
+  | add_typ_tvars(TVar(v),vs) = insert (op =) v vs;
+
+(*Accumulates the TFrees in a type, suppressing duplicates.*)
+fun add_typ_tfree_names(Type(_,Ts),fs) = List.foldr add_typ_tfree_names fs Ts
+  | add_typ_tfree_names(TFree(f,_),fs) = insert (op =) f fs
+  | add_typ_tfree_names(TVar(_),fs) = fs;
+
+fun add_typ_tfrees(Type(_,Ts),fs) = List.foldr add_typ_tfrees fs Ts
+  | add_typ_tfrees(TFree(f),fs) = insert (op =) f fs
+  | add_typ_tfrees(TVar(_),fs) = fs;
+
+(*Accumulates the TVars in a term, suppressing duplicates.*)
+val add_term_tvars = it_term_types add_typ_tvars;
+
+(*Accumulates the TFrees in a term, suppressing duplicates.*)
+val add_term_tfrees = it_term_types add_typ_tfrees;
+val add_term_tfree_names = it_term_types add_typ_tfree_names;
+
+(*Non-list versions*)
+fun typ_tfrees T = add_typ_tfrees(T,[]);
+fun typ_tvars T = add_typ_tvars(T,[]);
+fun term_tfrees t = add_term_tfrees(t,[]);
+fun term_tvars t = add_term_tvars(t,[]);
+
+
+(*Accumulates the Vars in the term, suppressing duplicates.*)
+fun add_term_vars (t, vars: term list) = case t of
+    Var   _ => OrdList.insert TermOrd.term_ord t vars
+  | Abs (_,_,body) => add_term_vars(body,vars)
+  | f$t =>  add_term_vars (f, add_term_vars(t, vars))
+  | _ => vars;
+
+fun term_vars t = add_term_vars(t,[]);
+
+(*Accumulates the Frees in the term, suppressing duplicates.*)
+fun add_term_frees (t, frees: term list) = case t of
+    Free   _ => OrdList.insert TermOrd.term_ord t frees
+  | Abs (_,_,body) => add_term_frees(body,frees)
+  | f$t =>  add_term_frees (f, add_term_frees(t, frees))
+  | _ => frees;
+
+fun term_frees t = add_term_frees(t,[]);
+
+end;
--- a/src/Pure/pattern.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/pattern.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,6 +1,5 @@
 (*  Title:      Pure/pattern.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow, Christine Heinzelmann, and Stefan Berghofer
+    Author:     Tobias Nipkow, Christine Heinzelmann, and Stefan Berghofer, TU Muenchen
 
 Unification of Higher-Order Patterns.
 
@@ -223,7 +222,7 @@
                      fun lam(is) = mkabs(binders,is,app(H,map (idx is) ks));
                  in Envir.update (((G, Gty), lam js), Envir.update (((F, Fty), lam is), env'))
                  end;
-  in if Term.indexname_ord (G,F) = LESS then ff(F,Fty,is,G,Gty,js) else ff(G,Gty,js,F,Fty,is) end
+  in if TermOrd.indexname_ord (G,F) = LESS then ff(F,Fty,is,G,Gty,js) else ff(G,Gty,js,F,Fty,is) end
 
 fun unify_types thy (T,U) (env as Envir.Envir{asol,iTs,maxidx}) =
   if T=U then env
--- a/src/Pure/proofterm.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/proofterm.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/proofterm.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 LF style proof terms.
@@ -189,7 +188,7 @@
 
 (* proof body *)
 
-val oracle_ord = prod_ord fast_string_ord Term.fast_term_ord;
+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 =
@@ -412,15 +411,15 @@
 fun del_conflicting_tvars envT T = TermSubst.instantiateT
   (map_filter (fn ixnS as (_, S) =>
      (Type.lookup envT ixnS; NONE) handle TYPE _ =>
-        SOME (ixnS, TFree ("'dummy", S))) (typ_tvars T)) T;
+        SOME (ixnS, TFree ("'dummy", S))) (OldTerm.typ_tvars T)) T;
 
 fun del_conflicting_vars env t = TermSubst.instantiate
   (map_filter (fn ixnS as (_, S) =>
      (Type.lookup (type_env env) ixnS; NONE) handle TYPE _ =>
-        SOME (ixnS, TFree ("'dummy", S))) (term_tvars t),
+        SOME (ixnS, TFree ("'dummy", S))) (OldTerm.term_tvars t),
    map_filter (fn Var (ixnT as (_, T)) =>
      (Envir.lookup (env, ixnT); NONE) handle TYPE _ =>
-        SOME (ixnT, Free ("dummy", T))) (term_vars t)) t;
+        SOME (ixnT, Free ("dummy", T))) (OldTerm.term_vars t)) t;
 
 fun norm_proof env =
   let
@@ -545,15 +544,13 @@
   let
     val (fs, Tfs, vs, Tvs) = fold_proof_terms
       (fn t => fn (fs, Tfs, vs, Tvs) =>
-         (add_term_frees (t, fs), add_term_tfree_names (t, Tfs),
-          add_term_vars (t, vs), add_term_tvar_ixns (t, Tvs)))
+         (Term.add_free_names t fs, Term.add_tfree_names t Tfs,
+          Term.add_var_names t vs, Term.add_tvar_names t Tvs))
       (fn T => fn (fs, Tfs, vs, Tvs) =>
-         (fs, add_typ_tfree_names (T, Tfs),
-          vs, add_typ_ixns (Tvs, T)))
+         (fs, Term.add_tfree_namesT T Tfs,
+          vs, Term.add_tvar_namesT T Tvs))
       prf ([], [], [], []);
-    val fs' = map (fst o dest_Free) fs;
-    val vs' = map (fst o dest_Var) vs;
-    val names = vs' ~~ Name.variant_list fs' (map fst vs');
+    val names = vs ~~ Name.variant_list fs (map fst vs);
     val names' = Tvs ~~ Name.variant_list Tfs (map fst Tvs);
     val rnames = map swap names;
     val rnames' = map swap names';
@@ -591,8 +588,9 @@
   let
     val fs = Term.fold_types (Term.fold_atyps
       (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t [];
-    val ixns = add_term_tvar_ixns (t, []);
-    val fmap = fs ~~ Name.variant_list (map #1 ixns) (map fst fs);
+    val used = Name.context
+      |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t;
+    val fmap = fs ~~ (#1 (Name.variants (map fst fs) used));
     fun thaw (f as (a, S)) =
       (case AList.lookup (op =) fmap f of
         NONE => TFree f
@@ -614,8 +612,8 @@
 
 fun freezeT t prf =
   let
-    val used = it_term_types add_typ_tfree_names (t, [])
-    and tvars = map #1 (it_term_types add_typ_tvars (t, []));
+    val used = OldTerm.it_term_types OldTerm.add_typ_tfree_names (t, [])
+    and tvars = map #1 (OldTerm.it_term_types OldTerm.add_typ_tvars (t, []));
     val (alist, _) = List.foldr new_name ([], used) tvars;
   in
     (case alist of
--- a/src/Pure/pure_setup.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/pure_setup.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -33,6 +33,7 @@
   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", "T"] (Pretty.pprint o Pretty.str o Binding.display));
 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);
--- a/src/Pure/search.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/search.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/search.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson and Norbert Voelker
 
 Search tacticals.
@@ -42,10 +41,8 @@
 (** Instantiation of heaps for best-first search **)
 
 (*total ordering on theorems, allowing duplicates to be found*)
-structure ThmHeap =
-  HeapFun (type elem = int * thm
-    val ord = Library.prod_ord Library.int_ord
-      (Term.term_ord o Library.pairself (#prop o Thm.rep_thm)));
+structure ThmHeap = HeapFun(type elem = int * thm
+  val ord = prod_ord int_ord (TermOrd.term_ord o pairself Thm.prop_of));
 
 
 structure Search : SEARCH =
--- a/src/Pure/sorts.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/sorts.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/sorts.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 The order-sorted algebra of type classes.
@@ -74,13 +73,13 @@
 
 (** ordered lists of sorts **)
 
-val make = OrdList.make Term.sort_ord;
-val op subset = OrdList.subset Term.sort_ord;
-val op union = OrdList.union Term.sort_ord;
-val subtract = OrdList.subtract Term.sort_ord;
+val make = OrdList.make TermOrd.sort_ord;
+val op subset = OrdList.subset TermOrd.sort_ord;
+val op union = OrdList.union TermOrd.sort_ord;
+val subtract = OrdList.subtract TermOrd.sort_ord;
 
-val remove_sort = OrdList.remove Term.sort_ord;
-val insert_sort = OrdList.insert Term.sort_ord;
+val remove_sort = OrdList.remove TermOrd.sort_ord;
+val insert_sort = OrdList.insert TermOrd.sort_ord;
 
 fun insert_typ (TFree (_, S)) Ss = insert_sort S Ss
   | insert_typ (TVar (_, S)) Ss = insert_sort S Ss
--- a/src/Pure/tactic.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/tactic.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,7 @@
 (*  Title:      Pure/tactic.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1991  University of Cambridge
 
-Basic tactics.
+Fundamental tactics.
 *)
 
 signature BASIC_TACTIC =
@@ -192,7 +190,7 @@
 (*Determine print names of goal parameters (reversed)*)
 fun innermost_params i st =
   let val (_, _, Bi, _) = dest_state (st, i)
-  in rename_wrt_term Bi (Logic.strip_params Bi) end;
+  in Term.rename_wrt_term Bi (Logic.strip_params Bi) end;
 
 (*params of subgoal i as they are printed*)
 fun params_of_state i st = rev (innermost_params i st);
--- a/src/Pure/tctical.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/tctical.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      Pure/tctical.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1993  University of Cambridge
 
 Tacticals.
 *)
@@ -404,19 +402,18 @@
   (*Left-to-right replacements: ctpairs = [...,(vi,ti),...].
     Instantiates distinct free variables by terms of same type.*)
   fun free_instantiate ctpairs =
-      forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs);
+    forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs);
 
-  fun free_of s ((a,i), T) =
-        Free(s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i),
-             T)
+  fun free_of s ((a, i), T) =
+    Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T)
 
-  fun mk_inst (var as Var(v,T))  = (var,  free_of "METAHYP1_" (v,T))
+  fun mk_inst v = (Var v, free_of "METAHYP1_" v)
 in
 
 (*Common code for METAHYPS and metahyps_thms*)
 fun metahyps_split_prem prem =
   let (*find all vars in the hyps -- should find tvars also!*)
-      val hyps_vars = List.foldr add_term_vars [] (Logic.strip_assums_hyp prem)
+      val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) []
       val insts = map mk_inst hyps_vars
       (*replace the hyps_vars by Frees*)
       val prem' = subst_atomic insts prem
@@ -434,20 +431,18 @@
       val cparams = map cterm fparams
       fun swap_ctpair (t,u) = (cterm u, cterm t)
       (*Subgoal variables: make Free; lift type over params*)
-      fun mk_subgoal_inst concl_vars (var as Var(v,T)) =
-          if member (op =) concl_vars var
-          then (var, true, free_of "METAHYP2_" (v,T))
-          else (var, false,
-                free_of "METAHYP2_" (v, map #2 params --->T))
+      fun mk_subgoal_inst concl_vars (v, T) =
+          if member (op =) concl_vars (v, T)
+          then ((v, T), true, free_of "METAHYP2_" (v, T))
+          else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T))
       (*Instantiate subgoal vars by Free applied to params*)
-      fun mk_ctpair (t,in_concl,u) =
-          if in_concl then (cterm t,  cterm u)
-          else (cterm t,  cterm (list_comb (u,fparams)))
+      fun mk_ctpair (v, in_concl, u) =
+          if in_concl then (cterm (Var v), cterm u)
+          else (cterm (Var v), cterm (list_comb (u, fparams)))
       (*Restore Vars with higher type and index*)
-      fun mk_subgoal_swap_ctpair
-                (t as Var((a,i),_), in_concl, u as Free(_,U)) =
-          if in_concl then (cterm u, cterm t)
-          else (cterm u, cterm(Var((a, i+maxidx), U)))
+      fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) =
+          if in_concl then (cterm u, cterm (Var ((a, i), T)))
+          else (cterm u, cterm (Var ((a, i + maxidx), U)))
       (*Embed B in the original context of params and hyps*)
       fun embed B = list_all_free (params, Logic.list_implies (hyps, B))
       (*Strip the context using elimination rules*)
@@ -456,8 +451,8 @@
       fun relift st =
         let val prop = Thm.prop_of st
             val subgoal_vars = (*Vars introduced in the subgoals*)
-                  List.foldr add_term_vars [] (Logic.strip_imp_prems prop)
-            and concl_vars = add_term_vars (Logic.strip_imp_concl prop, [])
+              fold Term.add_vars (Logic.strip_imp_prems prop) []
+            and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []
             val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars
             val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st
             val emBs = map (cterm o embed) (prems_of st')
@@ -490,9 +485,8 @@
   let
     fun typed ty = " has type: " ^ Syntax.string_of_typ_global thy ty;
     fun find_vars thy (Const (c, ty)) =
-        (case Term.typ_tvars ty
-         of [] => I
-          | _ => insert (op =) (c ^ typed ty))
+          if null (Term.add_tvarsT ty []) then I
+          else insert (op =) (c ^ typed ty)
       | find_vars thy (Var (xi, ty)) = insert (op =) (Term.string_of_vname xi ^ typed ty)
       | find_vars _ (Free _) = I
       | find_vars _ (Bound _) = I
--- a/src/Pure/term.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/term.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,12 +1,11 @@
 (*  Title:      Pure/term.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   Cambridge University 1992
+    Author:     Makarius
 
 Simply typed lambda-calculus: types, terms, and basic operations.
 *)
 
-infix 9  $;
+infix 9 $;
 infixr 5 -->;
 infixr --->;
 infix aconv;
@@ -75,12 +74,7 @@
   val fold_term_types: (term -> typ -> 'a -> 'a) -> term -> 'a -> 'a
   val fold_types: (typ -> 'a -> 'a) -> term -> 'a -> 'a
   val burrow_types: (typ list -> typ list) -> term list -> term list
-  val it_term_types: (typ * 'a -> 'a) -> term * 'a -> 'a
-  val add_term_names: term * string list -> string list
   val aconv: term * term -> bool
-  structure Vartab: TABLE
-  structure Typtab: TABLE
-  structure Termtab: TABLE
   val propT: typ
   val strip_all_body: term -> term
   val strip_all_vars: term -> (string * typ) list
@@ -94,8 +88,6 @@
   val subst_bound: term * term -> term
   val betapply: term * term -> term
   val betapplys: term * term list -> term
-  val eq_ix: indexname * indexname -> bool
-  val could_unify: term * term -> bool
   val subst_free: (term * term) list -> term -> term
   val abstract_over: term * term -> term
   val lambda: term -> term -> term
@@ -115,31 +107,10 @@
   val maxidx_of_typ: typ -> int
   val maxidx_of_typs: typ list -> int
   val maxidx_of_term: term -> int
-  val add_term_consts: term * string list -> string list
-  val term_consts: term -> string list
   val exists_subtype: (typ -> bool) -> typ -> bool
   val exists_type: (typ -> bool) -> term -> bool
   val exists_subterm: (term -> bool) -> term -> bool
   val exists_Const: (string * typ -> bool) -> term -> bool
-  val add_term_free_names: term * string list -> string list
-  val add_typ_tvars: typ * (indexname * sort) list -> (indexname * sort) list
-  val add_typ_tfree_names: typ * string list -> string list
-  val add_typ_tfrees: typ * (string * sort) list -> (string * sort) list
-  val add_typ_varnames: typ * string list -> string list
-  val add_term_tvars: term * (indexname * sort) list -> (indexname * sort) list
-  val add_term_tfrees: term * (string * sort) list -> (string * sort) list
-  val add_term_tfree_names: term * string list -> string list
-  val typ_tfrees: typ -> (string * sort) list
-  val typ_tvars: typ -> (indexname * sort) list
-  val term_tfrees: term -> (string * sort) list
-  val term_tvars: term -> (indexname * sort) list
-  val add_typ_ixns: indexname list * typ -> indexname list
-  val add_term_tvar_ixns: term * indexname list -> indexname list
-  val add_term_vars: term * term list -> term list
-  val term_vars: term -> term list
-  val add_term_frees: term * term list -> term list
-  val term_frees: term -> term list
-  val rename_wrt_term: term -> (string * 'a) list -> (string * 'a) list
   val show_question_marks: bool ref
 end;
 
@@ -151,41 +122,40 @@
   val a_itselfT: typ
   val all: typ -> term
   val argument_type_of: term -> int -> typ
-  val head_name_of: term -> string
+  val add_tvar_namesT: typ -> indexname list -> indexname list
+  val add_tvar_names: term -> indexname list -> indexname list
   val add_tvarsT: typ -> (indexname * sort) list -> (indexname * sort) list
   val add_tvars: term -> (indexname * sort) list -> (indexname * sort) list
+  val add_var_names: term -> indexname list -> indexname list
   val add_vars: term -> (indexname * typ) list -> (indexname * typ) list
-  val add_varnames: term -> indexname list -> indexname list
+  val add_tfree_namesT: typ -> string list -> string list
+  val add_tfree_names: term -> string list -> string list
   val add_tfreesT: typ -> (string * sort) list -> (string * sort) list
   val add_tfrees: term -> (string * sort) list -> (string * sort) list
+  val add_free_names: term -> string list -> string list
   val add_frees: term -> (string * typ) list -> (string * typ) list
+  val add_const_names: term -> string list -> string list
+  val add_consts: term -> (string * typ) list -> (string * typ) list
   val hidden_polymorphism: term -> (indexname * sort) list
+  val declare_typ_names: typ -> Name.context -> Name.context
+  val declare_term_names: term -> Name.context -> Name.context
+  val declare_term_frees: term -> Name.context -> Name.context
+  val variant_frees: term -> (string * 'a) list -> (string * 'a) list
+  val rename_wrt_term: term -> (string * 'a) list -> (string * 'a) list
+  val eq_ix: indexname * indexname -> bool
+  val eq_tvar: (indexname * sort) * (indexname * sort) -> bool
+  val eq_var: (indexname * typ) * (indexname * typ) -> bool
+  val could_unify: term * term -> bool
   val strip_abs_eta: int -> term -> (string * typ) list * term
-  val fast_indexname_ord: indexname * indexname -> order
-  val indexname_ord: indexname * indexname -> order
-  val sort_ord: sort * sort -> order
-  val typ_ord: typ * typ -> order
-  val fast_term_ord: term * term -> order
-  val term_ord: term * term -> order
-  val hd_ord: term * term -> order
-  val termless: term * term -> bool
-  val term_lpo: (term -> int) -> term * term -> order
   val match_bvars: (term * term) * (string * string) list -> (string * string) list
   val map_abs_vars: (string -> string) -> term -> term
   val rename_abs: term -> term -> term -> term option
-  val eq_tvar: (indexname * sort) * (indexname * sort) -> bool
-  val eq_var: (indexname * typ) * (indexname * typ) -> bool
-  val tvar_ord: (indexname * sort) * (indexname * sort) -> order
-  val var_ord: (indexname * typ) * (indexname * typ) -> order
   val close_schematic_term: term -> term
   val maxidx_typ: typ -> int -> int
   val maxidx_typs: typ list -> int -> int
   val maxidx_term: term -> int -> int
   val has_abs: term -> bool
   val dest_abs: string * typ * term -> string * term
-  val declare_typ_names: typ -> Name.context -> Name.context
-  val declare_term_names: term -> Name.context -> Name.context
-  val variant_frees: term -> (string * 'a) list -> (string * 'a) list
   val dummy_patternN: string
   val dummy_pattern: typ -> term
   val is_dummy_pattern: term -> bool
@@ -413,16 +383,6 @@
 fun head_of (f$t) = head_of f
   | head_of u = u;
 
-fun head_name_of tm =
-  (case head_of tm of
-    t as Const (c, _) =>
-      if NameSpace.is_qualified c then c
-      else raise TERM ("Malformed constant name", [t])
-  | t as Free (x, _) =>
-      if not (NameSpace.is_qualified x) then x
-      else raise TERM ("Malformed fixed variable name", [t])
-  | t => raise TERM ("No fixed head of term", [t]));
-
 (*number of atoms and abstractions in a term*)
 fun size_of_term tm =
   let
@@ -451,29 +411,16 @@
       | map_aux (t $ u) = map_aux t $ map_aux u;
   in map_aux end;
 
-(* iterate a function over all types in a term *)
-fun it_term_types f =
-let fun iter(Const(_,T), a) = f(T,a)
-      | iter(Free(_,T), a) = f(T,a)
-      | iter(Var(_,T), a) = f(T,a)
-      | iter(Abs(_,T,t), a) = iter(t,f(T,a))
-      | iter(f$u, a) = iter(f, iter(u, a))
-      | iter(Bound _, a) = a
-in iter end
-
 
 (* fold types and terms *)
 
-(*fold atoms of type*)
 fun fold_atyps f (Type (_, Ts)) = fold (fold_atyps f) Ts
   | fold_atyps f T = f T;
 
-(*fold atoms of term*)
 fun fold_aterms f (t $ u) = fold_aterms f t #> fold_aterms f u
   | fold_aterms f (Abs (_, _, t)) = fold_aterms f t
   | fold_aterms f a = f a;
 
-(*fold types of term*)
 fun fold_term_types f (t as Const (_, T)) = f t T
   | fold_term_types f (t as Free (_, T)) = f t T
   | fold_term_types f (t as Var (_, T)) = f t T
@@ -504,13 +451,20 @@
   in ts' end;
 
 (*collect variables*)
+val add_tvar_namesT = fold_atyps (fn TVar (xi, _) => insert (op =) xi | _ => I);
+val add_tvar_names = fold_types add_tvar_namesT;
 val add_tvarsT = fold_atyps (fn TVar v => insert (op =) v | _ => I);
 val add_tvars = fold_types add_tvarsT;
+val add_var_names = fold_aterms (fn Var (xi, _) => insert (op =) xi | _ => I);
 val add_vars = fold_aterms (fn Var v => insert (op =) v | _ => I);
-val add_varnames = fold_aterms (fn Var (xi, _) => insert (op =) xi | _ => I);
+val add_tfree_namesT = fold_atyps (fn TFree (xi, _) => insert (op =) xi | _ => I);
+val add_tfree_names = fold_types add_tfree_namesT;
 val add_tfreesT = fold_atyps (fn TFree v => insert (op =) v | _ => I);
 val add_tfrees = fold_types add_tfreesT;
+val add_free_names = fold_aterms (fn Free (x, _) => insert (op =) x | _ => I);
 val add_frees = fold_aterms (fn Free v => insert (op =) v | _ => I);
+val add_const_names = fold_aterms (fn Const (c, _) => insert (op =) c | _ => I);
+val add_consts = fold_aterms (fn Const c => insert (op =) c | _ => I);
 
 (*extra type variables in a term, not covered by its type*)
 fun hidden_polymorphism t =
@@ -522,10 +476,37 @@
   in extra_tvars end;
 
 
+(* renaming variables *)
 
-(** Comparing terms, types, sorts etc. **)
+val declare_typ_names = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I);
+
+fun declare_term_names tm =
+  fold_aterms
+    (fn Const (a, _) => Name.declare (NameSpace.base a)
+      | Free (a, _) => Name.declare a
+      | _ => I) tm #>
+  fold_types declare_typ_names tm;
+
+val declare_term_frees = fold_aterms (fn Free (x, _) => Name.declare x | _ => I);
+
+fun variant_frees t frees =
+  fst (Name.variants (map fst frees) (declare_term_names t Name.context)) ~~ map snd frees;
 
-(* alpha equivalence -- tuned for equalities *)
+fun rename_wrt_term t frees = rev (variant_frees t frees);  (*reversed result!*)
+
+
+
+(** Comparing terms **)
+
+(* variables *)
+
+fun eq_ix ((x, i): indexname, (y, j)) = i = j andalso x = y;
+
+fun eq_tvar ((xi, S: sort), (xi', S')) = eq_ix (xi, xi') andalso S = S';
+fun eq_var ((xi, T: typ), (xi', T')) = eq_ix (xi, xi') andalso T = T';
+
+
+(* alpha equivalence *)
 
 fun tm1 aconv tm2 =
   pointer_eq (tm1, tm2) orelse
@@ -535,184 +516,24 @@
     | (a1, a2) => a1 = a2);
 
 
-(* fast syntactic ordering -- tuned for inequalities *)
-
-fun fast_indexname_ord ((x, i), (y, j)) =
-  (case int_ord (i, j) of EQUAL => fast_string_ord (x, y) | ord => ord);
-
-fun sort_ord SS =
-  if pointer_eq SS then EQUAL
-  else dict_ord fast_string_ord SS;
-
-local
-
-fun cons_nr (TVar _) = 0
-  | cons_nr (TFree _) = 1
-  | cons_nr (Type _) = 2;
-
-in
-
-fun typ_ord TU =
-  if pointer_eq TU then EQUAL
-  else
-    (case TU of
-      (Type (a, Ts), Type (b, Us)) =>
-        (case fast_string_ord (a, b) of EQUAL => dict_ord typ_ord (Ts, Us) | ord => ord)
-    | (TFree (a, S), TFree (b, S')) =>
-        (case fast_string_ord (a, b) of EQUAL => sort_ord (S, S') | ord => ord)
-    | (TVar (xi, S), TVar (yj, S')) =>
-        (case fast_indexname_ord (xi, yj) of EQUAL => sort_ord (S, S') | ord => ord)
-    | (T, U) => int_ord (cons_nr T, cons_nr U));
-
-end;
-
-local
-
-fun cons_nr (Const _) = 0
-  | cons_nr (Free _) = 1
-  | cons_nr (Var _) = 2
-  | cons_nr (Bound _) = 3
-  | cons_nr (Abs _) = 4
-  | cons_nr (_ $ _) = 5;
-
-fun struct_ord (Abs (_, _, t), Abs (_, _, u)) = struct_ord (t, u)
-  | struct_ord (t1 $ t2, u1 $ u2) =
-      (case struct_ord (t1, u1) of EQUAL => struct_ord (t2, u2) | ord => ord)
-  | struct_ord (t, u) = int_ord (cons_nr t, cons_nr u);
-
-fun atoms_ord (Abs (_, _, t), Abs (_, _, u)) = atoms_ord (t, u)
-  | atoms_ord (t1 $ t2, u1 $ u2) =
-      (case atoms_ord (t1, u1) of EQUAL => atoms_ord (t2, u2) | ord => ord)
-  | atoms_ord (Const (a, _), Const (b, _)) = fast_string_ord (a, b)
-  | atoms_ord (Free (x, _), Free (y, _)) = fast_string_ord (x, y)
-  | atoms_ord (Var (xi, _), Var (yj, _)) = fast_indexname_ord (xi, yj)
-  | atoms_ord (Bound i, Bound j) = int_ord (i, j)
-  | atoms_ord _ = sys_error "atoms_ord";
-
-fun types_ord (Abs (_, T, t), Abs (_, U, u)) =
-      (case typ_ord (T, U) of EQUAL => types_ord (t, u) | ord => ord)
-  | types_ord (t1 $ t2, u1 $ u2) =
-      (case types_ord (t1, u1) of EQUAL => types_ord (t2, u2) | ord => ord)
-  | types_ord (Const (_, T), Const (_, U)) = typ_ord (T, U)
-  | types_ord (Free (_, T), Free (_, U)) = typ_ord (T, U)
-  | types_ord (Var (_, T), Var (_, U)) = typ_ord (T, U)
-  | types_ord (Bound _, Bound _) = EQUAL
-  | types_ord _ = sys_error "types_ord";
-
-in
-
-fun fast_term_ord tu =
-  if pointer_eq tu then EQUAL
-  else
-    (case struct_ord tu of
-      EQUAL => (case atoms_ord tu of EQUAL => types_ord tu | ord => ord)
-    | ord => ord);
-
-structure Vartab = TableFun(type key = indexname val ord = fast_indexname_ord);
-structure Typtab = TableFun(type key = typ val ord = typ_ord);
-structure Termtab = TableFun(type key = term val ord = fast_term_ord);
-
-end;
-
-
-(* term_ord *)
-
-(*a linear well-founded AC-compatible ordering for terms:
-  s < t <=> 1. size(s) < size(t) or
-            2. size(s) = size(t) and s=f(...) and t=g(...) and f<g or
-            3. size(s) = size(t) and s=f(s1..sn) and t=f(t1..tn) and
-               (s1..sn) < (t1..tn) (lexicographically)*)
+(*A fast unification filter: true unless the two terms cannot be unified.
+  Terms must be NORMAL.  Treats all Vars as distinct. *)
+fun could_unify (t, u) =
+  let
+    fun matchrands (f $ t) (g $ u) = could_unify (t, u) andalso matchrands f g
+      | matchrands _ _ = true;
+  in
+    case (head_of t, head_of u) of
+      (_, Var _) => true
+    | (Var _, _) => true
+    | (Const (a, _), Const (b, _)) => a = b andalso matchrands t u
+    | (Free (a, _), Free (b, _)) => a = b andalso matchrands t u
+    | (Bound i, Bound j) => i = j andalso matchrands t u
+    | (Abs _, _) => true   (*because of possible eta equality*)
+    | (_, Abs _) => true
+    | _ => false
+  end;
 
-fun indexname_ord ((x, i), (y, j)) =
-  (case int_ord (i, j) of EQUAL => string_ord (x, y) | ord => ord);
-
-local
-
-fun hd_depth (t $ _, n) = hd_depth (t, n + 1)
-  | hd_depth p = p;
-
-fun dest_hd (Const (a, T)) = (((a, 0), T), 0)
-  | dest_hd (Free (a, T)) = (((a, 0), T), 1)
-  | dest_hd (Var v) = (v, 2)
-  | dest_hd (Bound i) = ((("", i), dummyT), 3)
-  | dest_hd (Abs (_, T, _)) = ((("", 0), T), 4);
-
-in
-
-fun term_ord tu =
-  if pointer_eq tu then EQUAL
-  else
-    (case tu of
-      (Abs (_, T, t), Abs(_, U, u)) =>
-        (case term_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
-    | (t, u) =>
-        (case int_ord (size_of_term t, size_of_term u) of
-          EQUAL =>
-            (case prod_ord hd_ord int_ord (hd_depth (t, 0), hd_depth (u, 0)) of
-              EQUAL => args_ord (t, u) | ord => ord)
-        | ord => ord))
-and hd_ord (f, g) =
-  prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd f, dest_hd g)
-and args_ord (f $ t, g $ u) =
-      (case args_ord (f, g) of EQUAL => term_ord (t, u) | ord => ord)
-  | args_ord _ = EQUAL;
-
-fun termless tu = (term_ord tu = LESS);
-
-end;
-
-
-(** Lexicographic path order on terms **)
-
-(*
-  See Baader & Nipkow, Term rewriting, CUP 1998.
-  Without variables.  Const, Var, Bound, Free and Abs are treated all as
-  constants.
-
-  f_ord maps terms to integers and serves two purposes:
-  - Predicate on constant symbols.  Those that are not recognised by f_ord
-    must be mapped to ~1.
-  - Order on the recognised symbols.  These must be mapped to distinct
-    integers >= 0.
-  The argument of f_ord is never an application.
-*)
-
-local
-
-fun unrecognized (Const (a, T)) = ((1, ((a, 0), T)), 0)
-  | unrecognized (Free (a, T)) = ((1, ((a, 0), T)), 0)
-  | unrecognized (Var v) = ((1, v), 1)
-  | unrecognized (Bound i) = ((1, (("", i), dummyT)), 2)
-  | unrecognized (Abs (_, T, _)) = ((1, (("", 0), T)), 3);
-
-fun dest_hd f_ord t =
-      let val ord = f_ord t in
-        if ord = ~1 then unrecognized t else ((0, (("", ord), fastype_of t)), 0)
-      end
-
-fun term_lpo f_ord (s, t) =
-  let val (f, ss) = strip_comb s and (g, ts) = strip_comb t in
-    if forall (fn si => term_lpo f_ord (si, t) = LESS) ss
-    then case hd_ord f_ord (f, g) of
-        GREATER =>
-          if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
-          then GREATER else LESS
-      | EQUAL =>
-          if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
-          then list_ord (term_lpo f_ord) (ss, ts)
-          else LESS
-      | LESS => LESS
-    else GREATER
-  end
-and hd_ord f_ord (f, g) = case (f, g) of
-    (Abs (_, T, t), Abs (_, U, u)) =>
-      (case term_lpo f_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
-  | (_, _) => prod_ord (prod_ord int_ord
-                  (prod_ord indexname_ord typ_ord)) int_ord
-                (dest_hd f_ord f, dest_hd f_ord g)
-in
-val term_lpo = term_lpo
-end;
 
 
 (** Connectives of higher order logic **)
@@ -842,7 +663,7 @@
   of bound variables and implicit eta-expansion*)
 fun strip_abs_eta k t =
   let
-    val used = fold_aterms (fn Free (v, _) => Name.declare v | _ => I) t Name.context;
+    val used = fold_aterms declare_term_frees t Name.context;
     fun strip_abs t (0, used) = (([], t), (0, used))
       | strip_abs (Abs (v, T, t)) (k, used) =
           let
@@ -863,35 +684,8 @@
   in (vs1 @ vs2, t'') end;
 
 
-(* comparing variables *)
-
-fun eq_ix ((x, i): indexname, (y, j)) = i = j andalso x = y;
-
-fun eq_tvar ((xi, S: sort), (xi', S')) = eq_ix (xi, xi') andalso S = S';
-fun eq_var ((xi, T: typ), (xi', T')) = eq_ix (xi, xi') andalso T = T';
-
-val tvar_ord = prod_ord indexname_ord sort_ord;
-val var_ord = prod_ord indexname_ord typ_ord;
-
-
-(*A fast unification filter: true unless the two terms cannot be unified.
-  Terms must be NORMAL.  Treats all Vars as distinct. *)
-fun could_unify (t,u) =
-  let fun matchrands (f$t, g$u) = could_unify(t,u) andalso  matchrands(f,g)
-        | matchrands _ = true
-  in case (head_of t , head_of u) of
-        (_, Var _) => true
-      | (Var _, _) => true
-      | (Const(a,_), Const(b,_)) =>  a=b andalso matchrands(t,u)
-      | (Free(a,_), Free(b,_)) =>  a=b andalso matchrands(t,u)
-      | (Bound i, Bound j) =>  i=j andalso matchrands(t,u)
-      | (Abs _, _) =>  true   (*because of possible eta equality*)
-      | (_, Abs _) =>  true
-      | _ => false
-  end;
-
 (*Substitute new for free occurrences of old in a term*)
-fun subst_free [] = (fn t=>t)
+fun subst_free [] = I
   | subst_free pairs =
       let fun substf u =
             case AList.lookup (op aconv) pairs u of
@@ -1011,11 +805,11 @@
 (** Identifying first-order terms **)
 
 (*Differs from proofterm/is_fun in its treatment of TVar*)
-fun is_funtype (Type("fun",[_,_])) = true
+fun is_funtype (Type ("fun", [_, _])) = true
   | is_funtype _ = false;
 
 (*Argument Ts is a reverse list of binder types, needed if term t contains Bound vars*)
-fun has_not_funtype Ts t = not (is_funtype (fastype_of1 (Ts,t)));
+fun has_not_funtype Ts t = not (is_funtype (fastype_of1 (Ts, t)));
 
 (*First order means in all terms of the form f(t1,...,tn) no argument has a
   function type. The supplied quantifiers are excluded: their argument always
@@ -1057,7 +851,7 @@
 
 
 
-(**** Syntax-related declarations ****)
+(** misc syntax operations **)
 
 (* substructure *)
 
@@ -1086,109 +880,13 @@
       | _ => false);
   in ex end;
 
+fun exists_Const P = exists_subterm (fn Const c => P c | _ => false);
+
 fun has_abs (Abs _) = true
   | has_abs (t $ u) = has_abs t orelse has_abs u
   | has_abs _ = false;
 
 
-
-(** Consts etc. **)
-
-fun add_term_consts (Const (c, _), cs) = insert (op =) c cs
-  | add_term_consts (t $ u, cs) = add_term_consts (t, add_term_consts (u, cs))
-  | add_term_consts (Abs (_, _, t), cs) = add_term_consts (t, cs)
-  | add_term_consts (_, cs) = cs;
-
-fun term_consts t = add_term_consts(t,[]);
-
-fun exists_Const P = exists_subterm (fn Const c => P c | _ => false);
-
-
-(** TFrees and TVars **)
-
-(*Accumulates the names of Frees in the term, suppressing duplicates.*)
-fun add_term_free_names (Free(a,_), bs) = insert (op =) a bs
-  | add_term_free_names (f$u, bs) = add_term_free_names (f, add_term_free_names(u, bs))
-  | add_term_free_names (Abs(_,_,t), bs) = add_term_free_names(t,bs)
-  | add_term_free_names (_, bs) = bs;
-
-(*Accumulates the names in the term, suppressing duplicates.
-  Includes Frees and Consts.  For choosing unambiguous bound var names.*)
-fun add_term_names (Const(a,_), bs) = insert (op =) (NameSpace.base a) bs
-  | add_term_names (Free(a,_), bs) = insert (op =) a bs
-  | add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs))
-  | add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs)
-  | add_term_names (_, bs) = bs;
-
-(*Accumulates the TVars in a type, suppressing duplicates. *)
-fun add_typ_tvars(Type(_,Ts),vs) = List.foldr add_typ_tvars vs Ts
-  | add_typ_tvars(TFree(_),vs) = vs
-  | add_typ_tvars(TVar(v),vs) = insert (op =) v vs;
-
-(*Accumulates the TFrees in a type, suppressing duplicates. *)
-fun add_typ_tfree_names(Type(_,Ts),fs) = List.foldr add_typ_tfree_names fs Ts
-  | add_typ_tfree_names(TFree(f,_),fs) = insert (op =) f fs
-  | add_typ_tfree_names(TVar(_),fs) = fs;
-
-fun add_typ_tfrees(Type(_,Ts),fs) = List.foldr add_typ_tfrees fs Ts
-  | add_typ_tfrees(TFree(f),fs) = insert (op =) f fs
-  | add_typ_tfrees(TVar(_),fs) = fs;
-
-fun add_typ_varnames(Type(_,Ts),nms) = List.foldr add_typ_varnames nms Ts
-  | add_typ_varnames(TFree(nm,_),nms) = insert (op =) nm nms
-  | add_typ_varnames(TVar((nm,_),_),nms) = insert (op =) nm nms;
-
-(*Accumulates the TVars in a term, suppressing duplicates. *)
-val add_term_tvars = it_term_types add_typ_tvars;
-
-(*Accumulates the TFrees in a term, suppressing duplicates. *)
-val add_term_tfrees = it_term_types add_typ_tfrees;
-val add_term_tfree_names = it_term_types add_typ_tfree_names;
-
-(*Non-list versions*)
-fun typ_tfrees T = add_typ_tfrees(T,[]);
-fun typ_tvars T = add_typ_tvars(T,[]);
-fun term_tfrees t = add_term_tfrees(t,[]);
-fun term_tvars t = add_term_tvars(t,[]);
-
-(*special code to enforce left-to-right collection of TVar-indexnames*)
-
-fun add_typ_ixns(ixns,Type(_,Ts)) = Library.foldl add_typ_ixns (ixns,Ts)
-  | add_typ_ixns(ixns,TVar(ixn,_)) = if member (op =) ixns ixn then ixns
-                                     else ixns@[ixn]
-  | add_typ_ixns(ixns,TFree(_)) = ixns;
-
-fun add_term_tvar_ixns(Const(_,T),ixns) = add_typ_ixns(ixns,T)
-  | add_term_tvar_ixns(Free(_,T),ixns) = add_typ_ixns(ixns,T)
-  | add_term_tvar_ixns(Var(_,T),ixns) = add_typ_ixns(ixns,T)
-  | add_term_tvar_ixns(Bound _,ixns) = ixns
-  | add_term_tvar_ixns(Abs(_,T,t),ixns) =
-      add_term_tvar_ixns(t,add_typ_ixns(ixns,T))
-  | add_term_tvar_ixns(f$t,ixns) =
-      add_term_tvar_ixns(t,add_term_tvar_ixns(f,ixns));
-
-
-(** Frees and Vars **)
-
-(*Accumulates the Vars in the term, suppressing duplicates*)
-fun add_term_vars (t, vars: term list) = case t of
-    Var   _ => OrdList.insert term_ord t vars
-  | Abs (_,_,body) => add_term_vars(body,vars)
-  | f$t =>  add_term_vars (f, add_term_vars(t, vars))
-  | _ => vars;
-
-fun term_vars t = add_term_vars(t,[]);
-
-(*Accumulates the Frees in the term, suppressing duplicates*)
-fun add_term_frees (t, frees: term list) = case t of
-    Free   _ => OrdList.insert term_ord t frees
-  | Abs (_,_,body) => add_term_frees(body,frees)
-  | f$t =>  add_term_frees (f, add_term_frees(t, frees))
-  | _ => frees;
-
-fun term_frees t = add_term_frees(t,[]);
-
-
 (* dest abstraction *)
 
 fun dest_abs (x, T, body) =
@@ -1203,23 +901,6 @@
   end;
 
 
-(* renaming variables *)
-
-val declare_typ_names = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I);
-
-fun declare_term_names tm =
-  fold_aterms
-    (fn Const (a, _) => Name.declare (NameSpace.base a)
-      | Free (a, _) => Name.declare a
-      | _ => I) tm #>
-  fold_types declare_typ_names tm;
-
-fun variant_frees t frees =
-  fst (Name.variants (map fst frees) (declare_term_names t Name.context)) ~~ map snd frees;
-
-fun rename_wrt_term t frees = rev (variant_frees t frees);  (*reversed result!*)
-
-
 (* dummy patterns *)
 
 val dummy_patternN = "dummy_pattern";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/term_ord.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -0,0 +1,209 @@
+(*  Title:      Pure/term_ord.ML
+    Author:     Tobias Nipkow and Makarius, TU Muenchen
+
+Term orderings.
+*)
+
+signature TERM_ORD =
+sig
+  val fast_indexname_ord: indexname * indexname -> order
+  val sort_ord: sort * sort -> order
+  val typ_ord: typ * typ -> order
+  val fast_term_ord: term * term -> order
+  val indexname_ord: indexname * indexname -> order
+  val tvar_ord: (indexname * sort) * (indexname * sort) -> order
+  val var_ord: (indexname * typ) * (indexname * typ) -> order
+  val term_ord: term * term -> order
+  val hd_ord: term * term -> order
+  val termless: term * term -> bool
+  val term_lpo: (term -> int) -> term * term -> order
+end;
+
+structure TermOrd: TERM_ORD =
+struct
+
+(* fast syntactic ordering -- tuned for inequalities *)
+
+fun fast_indexname_ord ((x, i), (y, j)) =
+  (case int_ord (i, j) of EQUAL => fast_string_ord (x, y) | ord => ord);
+
+fun sort_ord SS =
+  if pointer_eq SS then EQUAL
+  else dict_ord fast_string_ord SS;
+
+local
+
+fun cons_nr (TVar _) = 0
+  | cons_nr (TFree _) = 1
+  | cons_nr (Type _) = 2;
+
+in
+
+fun typ_ord TU =
+  if pointer_eq TU then EQUAL
+  else
+    (case TU of
+      (Type (a, Ts), Type (b, Us)) =>
+        (case fast_string_ord (a, b) of EQUAL => dict_ord typ_ord (Ts, Us) | ord => ord)
+    | (TFree (a, S), TFree (b, S')) =>
+        (case fast_string_ord (a, b) of EQUAL => sort_ord (S, S') | ord => ord)
+    | (TVar (xi, S), TVar (yj, S')) =>
+        (case fast_indexname_ord (xi, yj) of EQUAL => sort_ord (S, S') | ord => ord)
+    | (T, U) => int_ord (cons_nr T, cons_nr U));
+
+end;
+
+local
+
+fun cons_nr (Const _) = 0
+  | cons_nr (Free _) = 1
+  | cons_nr (Var _) = 2
+  | cons_nr (Bound _) = 3
+  | cons_nr (Abs _) = 4
+  | cons_nr (_ $ _) = 5;
+
+fun struct_ord (Abs (_, _, t), Abs (_, _, u)) = struct_ord (t, u)
+  | struct_ord (t1 $ t2, u1 $ u2) =
+      (case struct_ord (t1, u1) of EQUAL => struct_ord (t2, u2) | ord => ord)
+  | struct_ord (t, u) = int_ord (cons_nr t, cons_nr u);
+
+fun atoms_ord (Abs (_, _, t), Abs (_, _, u)) = atoms_ord (t, u)
+  | atoms_ord (t1 $ t2, u1 $ u2) =
+      (case atoms_ord (t1, u1) of EQUAL => atoms_ord (t2, u2) | ord => ord)
+  | atoms_ord (Const (a, _), Const (b, _)) = fast_string_ord (a, b)
+  | atoms_ord (Free (x, _), Free (y, _)) = fast_string_ord (x, y)
+  | atoms_ord (Var (xi, _), Var (yj, _)) = fast_indexname_ord (xi, yj)
+  | atoms_ord (Bound i, Bound j) = int_ord (i, j)
+  | atoms_ord _ = sys_error "atoms_ord";
+
+fun types_ord (Abs (_, T, t), Abs (_, U, u)) =
+      (case typ_ord (T, U) of EQUAL => types_ord (t, u) | ord => ord)
+  | types_ord (t1 $ t2, u1 $ u2) =
+      (case types_ord (t1, u1) of EQUAL => types_ord (t2, u2) | ord => ord)
+  | types_ord (Const (_, T), Const (_, U)) = typ_ord (T, U)
+  | types_ord (Free (_, T), Free (_, U)) = typ_ord (T, U)
+  | types_ord (Var (_, T), Var (_, U)) = typ_ord (T, U)
+  | types_ord (Bound _, Bound _) = EQUAL
+  | types_ord _ = sys_error "types_ord";
+
+in
+
+fun fast_term_ord tu =
+  if pointer_eq tu then EQUAL
+  else
+    (case struct_ord tu of
+      EQUAL => (case atoms_ord tu of EQUAL => types_ord tu | ord => ord)
+    | ord => ord);
+
+end;
+
+
+(* term_ord *)
+
+(*a linear well-founded AC-compatible ordering for terms:
+  s < t <=> 1. size(s) < size(t) or
+            2. size(s) = size(t) and s=f(...) and t=g(...) and f<g or
+            3. size(s) = size(t) and s=f(s1..sn) and t=f(t1..tn) and
+               (s1..sn) < (t1..tn) (lexicographically)*)
+
+fun indexname_ord ((x, i), (y, j)) =
+  (case int_ord (i, j) of EQUAL => string_ord (x, y) | ord => ord);
+
+val tvar_ord = prod_ord indexname_ord sort_ord;
+val var_ord = prod_ord indexname_ord typ_ord;
+
+
+local
+
+fun hd_depth (t $ _, n) = hd_depth (t, n + 1)
+  | hd_depth p = p;
+
+fun dest_hd (Const (a, T)) = (((a, 0), T), 0)
+  | dest_hd (Free (a, T)) = (((a, 0), T), 1)
+  | dest_hd (Var v) = (v, 2)
+  | dest_hd (Bound i) = ((("", i), dummyT), 3)
+  | dest_hd (Abs (_, T, _)) = ((("", 0), T), 4);
+
+in
+
+fun term_ord tu =
+  if pointer_eq tu then EQUAL
+  else
+    (case tu of
+      (Abs (_, T, t), Abs(_, U, u)) =>
+        (case term_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
+    | (t, u) =>
+        (case int_ord (size_of_term t, size_of_term u) of
+          EQUAL =>
+            (case prod_ord hd_ord int_ord (hd_depth (t, 0), hd_depth (u, 0)) of
+              EQUAL => args_ord (t, u) | ord => ord)
+        | ord => ord))
+and hd_ord (f, g) =
+  prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd f, dest_hd g)
+and args_ord (f $ t, g $ u) =
+      (case args_ord (f, g) of EQUAL => term_ord (t, u) | ord => ord)
+  | args_ord _ = EQUAL;
+
+fun termless tu = (term_ord tu = LESS);
+
+end;
+
+
+(* Lexicographic path order on terms *)
+
+(*
+  See Baader & Nipkow, Term rewriting, CUP 1998.
+  Without variables.  Const, Var, Bound, Free and Abs are treated all as
+  constants.
+
+  f_ord maps terms to integers and serves two purposes:
+  - Predicate on constant symbols.  Those that are not recognised by f_ord
+    must be mapped to ~1.
+  - Order on the recognised symbols.  These must be mapped to distinct
+    integers >= 0.
+  The argument of f_ord is never an application.
+*)
+
+local
+
+fun unrecognized (Const (a, T)) = ((1, ((a, 0), T)), 0)
+  | unrecognized (Free (a, T)) = ((1, ((a, 0), T)), 0)
+  | unrecognized (Var v) = ((1, v), 1)
+  | unrecognized (Bound i) = ((1, (("", i), dummyT)), 2)
+  | unrecognized (Abs (_, T, _)) = ((1, (("", 0), T)), 3);
+
+fun dest_hd f_ord t =
+  let val ord = f_ord t
+  in if ord = ~1 then unrecognized t else ((0, (("", ord), fastype_of t)), 0) end;
+
+fun term_lpo f_ord (s, t) =
+  let val (f, ss) = strip_comb s and (g, ts) = strip_comb t in
+    if forall (fn si => term_lpo f_ord (si, t) = LESS) ss
+    then case hd_ord f_ord (f, g) of
+        GREATER =>
+          if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
+          then GREATER else LESS
+      | EQUAL =>
+          if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
+          then list_ord (term_lpo f_ord) (ss, ts)
+          else LESS
+      | LESS => LESS
+    else GREATER
+  end
+and hd_ord f_ord (f, g) = case (f, g) of
+    (Abs (_, T, t), Abs (_, U, u)) =>
+      (case term_lpo f_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
+  | (_, _) => prod_ord (prod_ord int_ord
+                  (prod_ord indexname_ord typ_ord)) int_ord
+                (dest_hd f_ord f, dest_hd f_ord g);
+
+in
+val term_lpo = term_lpo
+end;
+
+
+end;
+
+structure Vartab = TableFun(type key = indexname val ord = TermOrd.fast_indexname_ord);
+structure Typtab = TableFun(type key = typ val ord = TermOrd.typ_ord);
+structure Termtab = TableFun(type key = term val ord = TermOrd.fast_term_ord);
--- a/src/Pure/term_subst.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/term_subst.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/term_subst.ML
-    ID:         $Id$
     Author:     Makarius
 
 Efficient term substitution -- avoids copying.
@@ -163,9 +162,9 @@
 
 fun zero_var_indexes_inst ts =
   let
-    val tvars = sort Term.tvar_ord (fold Term.add_tvars ts []);
+    val tvars = sort TermOrd.tvar_ord (fold Term.add_tvars ts []);
     val instT = map (apsnd TVar) (zero_var_inst tvars);
-    val vars = sort Term.var_ord (map (apsnd (instantiateT instT)) (fold Term.add_vars ts []));
+    val vars = sort TermOrd.var_ord (map (apsnd (instantiateT instT)) (fold Term.add_vars ts []));
     val inst = map (apsnd Var) (zero_var_inst vars);
   in (instT, inst) end;
 
--- a/src/Pure/thm.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/thm.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,6 @@
 (*  Title:      Pure/thm.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1994  University of Cambridge
+    Author:     Makarius
 
 The very core of Isabelle's Meta Logic: certified types and terms,
 derivations, theorems, framework rules (including lifting and
@@ -380,9 +379,9 @@
 
 fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop;
 
-val union_hyps = OrdList.union Term.fast_term_ord;
-val insert_hyps = OrdList.insert Term.fast_term_ord;
-val remove_hyps = OrdList.remove Term.fast_term_ord;
+val union_hyps = OrdList.union TermOrd.fast_term_ord;
+val insert_hyps = OrdList.insert TermOrd.fast_term_ord;
+val remove_hyps = OrdList.remove TermOrd.fast_term_ord;
 
 
 (* merge theories of cterms/thms -- trivial absorption only *)
@@ -1155,7 +1154,7 @@
 (* Replace all TFrees not fixed or in the hyps by new TVars *)
 fun varifyT' fixed (Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
   let
-    val tfrees = List.foldr add_term_tfrees fixed hyps;
+    val tfrees = fold Term.add_tfrees hyps fixed;
     val prop1 = attach_tpairs tpairs prop;
     val (al, prop2) = Type.varify tfrees prop1;
     val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2);
@@ -1561,9 +1560,9 @@
 (*Quick test whether rule is resolvable with the subgoal with hyps Hs
   and conclusion B.  If eres_flg then checks 1st premise of rule also*)
 fun could_bires (Hs, B, eres_flg, rule) =
-    let fun could_reshyp (A1::_) = exists (fn H => could_unify (A1, H)) Hs
+    let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs
           | could_reshyp [] = false;  (*no premise -- illegal*)
-    in  could_unify(concl_of rule, B) andalso
+    in  Term.could_unify(concl_of rule, B) andalso
         (not eres_flg  orelse  could_reshyp (prems_of rule))
     end;
 
@@ -1696,7 +1695,7 @@
   val empty = NameSpace.empty_table;
   val copy = I;
   val extend = I;
-  fun merge _ oracles = NameSpace.merge_tables (op =) oracles
+  fun merge _ oracles : T = NameSpace.merge_tables (op =) oracles
     handle Symtab.DUP dup => err_dup_ora dup;
 );
 
--- a/src/Pure/type.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/type.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/type.ML
-    ID:         $Id$
     Author:     Tobias Nipkow, Lawrence C Paulson, and Markus Wenzel
 
 Type signatures and certified types, special treatment of type vars,
@@ -266,9 +265,9 @@
 (* no_tvars *)
 
 fun no_tvars T =
-  (case typ_tvars T of [] => T
+  (case Term.add_tvarsT T [] of [] => T
   | vs => raise TYPE ("Illegal schematic type variable(s): " ^
-      commas_quote (map (Term.string_of_vname o #1) vs), [T], []));
+      commas_quote (map (Term.string_of_vname o #1) (rev vs)), [T], []));
 
 
 (* varify *)
@@ -277,8 +276,9 @@
   let
     val fs = Term.fold_types (Term.fold_atyps
       (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t [];
-    val ixns = add_term_tvar_ixns (t, []);
-    val fmap = fs ~~ map (rpair 0) (Name.variant_list (map #1 ixns) (map fst fs))
+    val used = Name.context
+      |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t;
+    val fmap = fs ~~ map (rpair 0) (#1 (Name.variants (map fst fs) used));
     fun thaw (f as (a, S)) =
       (case AList.lookup (op =) fmap f of
         NONE => TFree f
@@ -307,8 +307,8 @@
 (*this sort of code could replace unvarifyT*)
 fun freeze_thaw_type T =
   let
-    val used = add_typ_tfree_names (T, [])
-    and tvars = map #1 (add_typ_tvars (T, []));
+    val used = OldTerm.add_typ_tfree_names (T, [])
+    and tvars = map #1 (OldTerm.add_typ_tvars (T, []));
     val (alist, _) = List.foldr new_name ([], used) tvars;
   in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end;
 
@@ -316,8 +316,8 @@
 
 fun freeze_thaw t =
   let
-    val used = it_term_types add_typ_tfree_names (t, [])
-    and tvars = map #1 (it_term_types add_typ_tvars (t, []));
+    val used = OldTerm.it_term_types OldTerm.add_typ_tfree_names (t, [])
+    and tvars = map #1 (OldTerm.it_term_types OldTerm.add_typ_tvars (t, []));
     val (alist, _) = List.foldr new_name ([], used) tvars;
   in
     (case alist of
@@ -400,7 +400,7 @@
     fun occ (Type (_, Ts)) = exists occ Ts
       | occ (TFree _) = false
       | occ (TVar (w, S)) =
-          eq_ix (v, w) orelse
+          Term.eq_ix (v, w) orelse
             (case lookup tye (w, S) of
               NONE => false
             | SOME U => occ U);
@@ -437,7 +437,7 @@
     fun unif (ty1, ty2) tye =
       (case (devar tye ty1, devar tye ty2) of
         (T as TVar (v, S1), U as TVar (w, S2)) =>
-          if eq_ix (v, w) then
+          if Term.eq_ix (v, w) then
             if S1 = S2 then tye else tvar_clash v S1 S2
           else if Sorts.sort_le classes (S1, S2) then
             Vartab.update_new (w, (S2, T)) tye
@@ -465,7 +465,7 @@
 fun raw_unify (ty1, ty2) tye =
   (case (devar tye ty1, devar tye ty2) of
     (T as TVar (v, S1), U as TVar (w, S2)) =>
-      if eq_ix (v, w) then
+      if Term.eq_ix (v, w) then
         if S1 = S2 then tye else tvar_clash v S1 S2
       else Vartab.update_new (w, (S2, T)) tye
   | (TVar (v, S), T) =>
--- a/src/Pure/unify.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/unify.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/unify.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   Cambridge University 1992
 
@@ -104,7 +103,7 @@
   | occur (Free _)  = false
   | occur (Var (w, T))  =
       if member (op =) (!seen) w then false
-      else if eq_ix(v,w) then true
+      else if Term.eq_ix (v, w) then true
         (*no need to lookup: v has no assignment*)
       else (seen := w:: !seen;
             case Envir.lookup (env, (w, T)) of
@@ -167,7 +166,7 @@
   | occur (Free _)  = NoOcc
   | occur (Var (w, T))  =
       if member (op =) (!seen) w then NoOcc
-      else if eq_ix(v,w) then Rigid
+      else if Term.eq_ix (v, w) then Rigid
       else (seen := w:: !seen;
             case Envir.lookup (env, (w, T)) of
           NONE    => NoOcc
@@ -176,7 +175,7 @@
   | occur (t as f$_) =  (*switch to nonrigid search?*)
      (case head_of_in (env,f) of
         Var (w,_) => (*w is not assigned*)
-    if eq_ix(v,w) then Rigid
+    if Term.eq_ix (v, w) then Rigid
     else  nonrigid t
       | Abs(_,_,body) => nonrigid t (*not in normal form*)
       | _ => occomb t)
@@ -541,10 +540,10 @@
   let val t = Envir.norm_term env t0  and  u = Envir.norm_term env u0
   in  case  (head_of t, head_of u) of
       (Var(v,T), Var(w,U)) =>  (*Check for identical variables...*)
-  if eq_ix(v,w) then     (*...occur check would falsely return true!*)
+  if Term.eq_ix (v, w) then     (*...occur check would falsely return true!*)
       if T=U then (env, add_tpair (rbinder, (t,u), tpairs))
       else raise TERM ("add_ffpair: Var name confusion", [t,u])
-  else if Term.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*)
+  else if TermOrd.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*)
        clean_ffpair thy ((rbinder, u, t), (env,tpairs))
         else clean_ffpair thy ((rbinder, t, u), (env,tpairs))
     | _ => raise TERM ("add_ffpair: Vars expected", [t,u])
--- a/src/Pure/variable.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Pure/variable.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Pure/variable.ML
-    ID:         $Id$
     Author:     Makarius
 
 Fixed type/term variables and polymorphic term abbreviations.
@@ -179,12 +178,12 @@
 (* names *)
 
 fun declare_type_names t =
-  map_names (fold_types (fold_atyps (fn TFree (a, _) => Name.declare a | _ => I)) t) #>
+  map_names (fold_types (fold_atyps Term.declare_typ_names) t) #>
   map_maxidx (fold_types Term.maxidx_typ t);
 
 fun declare_names t =
   declare_type_names t #>
-  map_names (fold_aterms (fn Free (x, _) => Name.declare x | _ => I) t) #>
+  map_names (fold_aterms Term.declare_term_frees t) #>
   map_maxidx (Term.maxidx_term t);
 
 
--- a/src/Sequents/modal.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Sequents/modal.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,8 @@
 (*  Title:      Sequents/modal.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 
-Simple modal reasoner
+Simple modal reasoner.
 *)
 
 
@@ -39,7 +38,7 @@
   Assumes each formula in seqc is surrounded by sequence variables
   -- checks that each concl formula looks like some subgoal formula.*)
 fun could_res (seqp,seqc) =
-      forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) 
+      forall (fn Qc => exists (fn Qp => Term.could_unify (Qp,Qc)) 
                               (forms_of_seq seqp))
              (forms_of_seq seqc);
 
--- a/src/Sequents/prover.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Sequents/prover.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,9 +1,8 @@
-(*  Title:      Sequents/prover
-    ID:         $Id$
+(*  Title:      Sequents/prover.ML
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 
-Simple classical reasoner for the sequent calculus, based on "theorem packs"
+Simple classical reasoner for the sequent calculus, based on "theorem packs".
 *)
 
 
@@ -65,7 +64,7 @@
   -- checks that each concl formula looks like some subgoal formula.
   It SHOULD check order as well, using recursion rather than forall/exists*)
 fun could_res (seqp,seqc) =
-      forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) 
+      forall (fn Qc => exists (fn Qp => Term.could_unify (Qp,Qc)) 
                               (forms_of_seq seqp))
              (forms_of_seq seqc);
 
@@ -78,7 +77,7 @@
 	  could_res (leftp,leftc) andalso could_res (rightp,rightc)
     | (_ $ Abs(_,_,leftp) $ rightp,
        _ $ Abs(_,_,leftc) $ rightc) =>
-	  could_res (leftp,leftc)  andalso  could_unify (rightp,rightc)
+	  could_res (leftp,leftc)  andalso  Term.could_unify (rightp,rightc)
     | _ => false;
 
 
--- a/src/Tools/Compute_Oracle/compute.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/Compute_Oracle/compute.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Tools/Compute_Oracle/compute.ML
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
@@ -168,7 +167,7 @@
   | machine_of_prog (ProgHaskell _) = HASKELL
   | machine_of_prog (ProgSML _) = SML
 
-structure Sorttab = TableFun(type key = sort val ord = Term.sort_ord)
+structure Sorttab = TableFun(type key = sort val ord = TermOrd.sort_ord)
 
 type naming = int -> string
 
--- a/src/Tools/Compute_Oracle/linker.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/Compute_Oracle/linker.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Tools/Compute_Oracle/linker.ML
-    ID:         $Id$
     Author:     Steven Obua
 
 Linker.ML solves the problem that the computing oracle does not
@@ -51,7 +50,7 @@
   | constant_of _ = raise Link "constant_of"
 
 fun bool_ord (x,y) = if x then (if y then EQUAL else GREATER) else (if y then LESS else EQUAL)
-fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) Term.typ_ord) (((x1,x2),x3), ((y1,y2),y3))
+fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) TermOrd.typ_ord) (((x1,x2),x3), ((y1,y2),y3))
 fun constant_modty_ord (Constant (x1,x2,_), Constant (y1,y2,_)) = (prod_ord bool_ord fast_string_ord) ((x1,x2), (y1,y2))
 
 
@@ -71,7 +70,7 @@
     handle Type.TYPE_MATCH => NONE
 
 fun subst_ord (A:Type.tyenv, B:Type.tyenv) =
-    (list_ord (prod_ord Term.fast_indexname_ord (prod_ord Term.sort_ord Term.typ_ord))) (Vartab.dest A, Vartab.dest B)
+    (list_ord (prod_ord TermOrd.fast_indexname_ord (prod_ord TermOrd.sort_ord TermOrd.typ_ord))) (Vartab.dest A, Vartab.dest B)
 
 structure Substtab = TableFun(type key = Type.tyenv val ord = subst_ord);
 
@@ -82,7 +81,7 @@
 
 datatype instances = Instances of unit ConsttabModTy.table * Type.tyenv Consttab.table Consttab.table * constant list list * unit Substtab.table
 
-fun is_polymorphic (Constant (_, _, ty)) = not (null (typ_tvars ty))
+fun is_polymorphic (Constant (_, _, ty)) = not (null (Term.add_tvarsT ty []))
 
 fun distinct_constants cs =
     Consttab.keys (fold (fn c => Consttab.update (c, ())) cs Consttab.empty)
@@ -91,7 +90,7 @@
     let
         val cs = distinct_constants (filter is_polymorphic cs)
         val old_cs = cs
-(*      fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (typ_tvars ty) tab
+(*      fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (OldTerm.typ_tvars ty) tab
         val tvars_count = length (Typtab.keys (fold (fn c => fn tab => collect_tvars (typ_of_constant c) tab) cs Typtab.empty))
         fun tvars_of ty = collect_tvars ty Typtab.empty
         val cs = map (fn c => (c, tvars_of (typ_of_constant c))) cs
@@ -263,10 +262,10 @@
 let
     val (left, right) = collect_consts_of_thm th
     val polycs = filter Linker.is_polymorphic left
-    val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty
+    val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (OldTerm.typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty
     fun check_const (c::cs) cs' =
         let
-            val tvars = typ_tvars (Linker.typ_of_constant c)
+            val tvars = OldTerm.typ_tvars (Linker.typ_of_constant c)
             val wrong = fold (fn n => fn wrong => wrong orelse is_none (Typtab.lookup tytab (TVar n))) tvars false
         in
             if wrong then raise PCompute "right hand side of theorem contains type variables which do not occur on the left hand side"
@@ -379,7 +378,7 @@
         ComputeThm (hyps, shyps, prop)
     end
 
-val cthm_ord' = prod_ord (prod_ord (list_ord Term.term_ord) (list_ord Term.sort_ord)) Term.term_ord
+val cthm_ord' = prod_ord (prod_ord (list_ord TermOrd.term_ord) (list_ord TermOrd.sort_ord)) TermOrd.term_ord
 
 fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) = cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2))
 
--- a/src/Tools/IsaPlanner/isand.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/IsaPlanner/isand.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Tools/IsaPlanner/isand.ML
-    ID:		$Id$
     Author:     Lucas Dixon, University of Edinburgh
 
 Natural Deduction tools.
@@ -187,8 +186,8 @@
 (* change type-vars to fresh type frees *)
 fun fix_tvars_to_tfrees_in_terms names ts = 
     let 
-      val tfree_names = map fst (List.foldr Term.add_term_tfrees [] ts);
-      val tvars = List.foldr Term.add_term_tvars [] ts;
+      val tfree_names = map fst (List.foldr OldTerm.add_term_tfrees [] ts);
+      val tvars = List.foldr OldTerm.add_term_tvars [] ts;
       val (names',renamings) = 
           List.foldr (fn (tv as ((n,i),s),(Ns,Rs)) => 
                          let val n2 = Name.variant Ns n in 
@@ -213,15 +212,15 @@
 fun unfix_tfrees ns th = 
     let 
       val varfiytfrees = map (Term.dest_TFree o Thm.typ_of) ns
-      val skiptfrees = subtract (op =) varfiytfrees (Term.add_term_tfrees (Thm.prop_of th,[]));
+      val skiptfrees = subtract (op =) varfiytfrees (OldTerm.add_term_tfrees (Thm.prop_of th,[]));
     in #2 (Thm.varifyT' skiptfrees th) end;
 
 (* change schematic/meta vars to fresh free vars, avoiding name clashes 
    with "names" *)
 fun fix_vars_to_frees_in_terms names ts = 
     let 
-      val vars = map Term.dest_Var (List.foldr Term.add_term_vars [] ts);
-      val Ns = List.foldr Term.add_term_names names ts;
+      val vars = map Term.dest_Var (List.foldr OldTerm.add_term_vars [] ts);
+      val Ns = List.foldr OldTerm.add_term_names names ts;
       val (_,renamings) = 
           Library.foldl (fn ((Ns,Rs),v as ((n,i),ty)) => 
                     let val n2 = Name.variant Ns n in
@@ -246,7 +245,7 @@
       val ctypify = Thm.ctyp_of sgn
       val tpairs = Thm.terms_of_tpairs (Thm.tpairs_of th);
       val prop = (Thm.prop_of th);
-      val tvars = List.foldr Term.add_term_tvars [] (prop :: tpairs);
+      val tvars = List.foldr OldTerm.add_term_tvars [] (prop :: tpairs);
       val ctyfixes = 
           map_filter 
             (fn (v as ((s,i),ty)) => 
@@ -259,7 +258,7 @@
       val ctermify = Thm.cterm_of sgn
       val tpairs = Thm.terms_of_tpairs (Thm.tpairs_of th);
       val prop = (Thm.prop_of th);
-      val vars = map Term.dest_Var (List.foldr Term.add_term_vars 
+      val vars = map Term.dest_Var (List.foldr OldTerm.add_term_vars 
                                                [] (prop :: tpairs));
       val cfixes = 
           map_filter 
@@ -360,7 +359,7 @@
       val sgn = Thm.theory_of_thm th;
       val ctermify = Thm.cterm_of sgn;
 
-      val allfrees = map Term.dest_Free (Term.term_frees (Thm.prop_of th))
+      val allfrees = map Term.dest_Free (OldTerm.term_frees (Thm.prop_of th))
       val cfrees = map (ctermify o Free o lookupfree allfrees) vs
 
       val sgs = prems_of th;
@@ -420,8 +419,8 @@
     let
       val t = Term.strip_all_body alledt;
       val alls = rev (Term.strip_all_vars alledt);
-      val varnames = map (fst o fst o Term.dest_Var) (Term.term_vars t)
-      val names = Term.add_term_names (t,varnames);
+      val varnames = map (fst o fst o Term.dest_Var) (OldTerm.term_vars t)
+      val names = OldTerm.add_term_names (t,varnames);
       val fvs = map Free 
                     (Name.variant_list names (map fst alls)
                        ~~ (map snd alls));
@@ -429,8 +428,8 @@
 
 fun fix_alls_term i t = 
     let 
-      val varnames = map (fst o fst o Term.dest_Var) (Term.term_vars t)
-      val names = Term.add_term_names (t,varnames);
+      val varnames = map (fst o fst o Term.dest_Var) (OldTerm.term_vars t)
+      val names = OldTerm.add_term_names (t,varnames);
       val gt = Logic.get_goal t i;
       val body = Term.strip_all_body gt;
       val alls = rev (Term.strip_all_vars gt);
--- a/src/Tools/IsaPlanner/rw_inst.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/IsaPlanner/rw_inst.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Tools/IsaPlanner/rw_inst.ML
-    ID:         $Id$
     Author:     Lucas Dixon, University of Edinburgh
 
 Rewriting using a conditional meta-equality theorem which supports
@@ -72,7 +71,7 @@
       val (tpairl,tpairr) = Library.split_list (#tpairs rep)
       val prop = #prop rep
     in
-      List.foldr Term.add_term_names [] (prop :: (tpairl @ (tpairr @ hyps)))
+      List.foldr OldTerm.add_term_names [] (prop :: (tpairl @ (tpairr @ hyps)))
     end;
 
 (* Given a list of variables that were bound, and a that has been
@@ -137,15 +136,15 @@
 fun mk_renamings tgt rule_inst = 
     let
       val rule_conds = Thm.prems_of rule_inst
-      val names = foldr Term.add_term_names [] (tgt :: rule_conds);
+      val names = foldr OldTerm.add_term_names [] (tgt :: rule_conds);
       val (conds_tyvs,cond_vs) = 
           Library.foldl (fn ((tyvs, vs), t) => 
                     (Library.union
-                       (Term.term_tvars t, tyvs),
+                       (OldTerm.term_tvars t, tyvs),
                      Library.union 
-                       (map Term.dest_Var (Term.term_vars t), vs))) 
+                       (map Term.dest_Var (OldTerm.term_vars t), vs))) 
                 (([],[]), rule_conds);
-      val termvars = map Term.dest_Var (Term.term_vars tgt); 
+      val termvars = map Term.dest_Var (OldTerm.term_vars tgt); 
       val vars_to_fix = Library.union (termvars, cond_vs);
       val (renamings, names2) = 
           foldr (fn (((n,i),ty), (vs, names')) => 
@@ -168,8 +167,8 @@
       val ignore_ixs = map fst ignore_insts;
       val (tvars, tfrees) = 
             foldr (fn (t, (varixs, tfrees)) => 
-                      (Term.add_term_tvars (t,varixs),
-                       Term.add_term_tfrees (t,tfrees)))
+                      (OldTerm.add_term_tvars (t,varixs),
+                       OldTerm.add_term_tfrees (t,tfrees)))
                   ([],[]) ts;
         val unfixed_tvars = 
             List.filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars;
--- a/src/Tools/code/code_ml.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/code/code_ml.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -954,7 +954,7 @@
 fun eval eval'' term_of reff thy ct args =
   let
     val ctxt = ProofContext.init thy;
-    val _ = if null (term_frees (term_of ct)) then () else error ("Term "
+    val _ = if null (Term.add_frees (term_of ct) []) then () else error ("Term "
       ^ quote (Syntax.string_of_term_global thy (term_of ct))
       ^ " to be evaluated contains free variables");
     fun eval' naming program ((vs, ty), t) deps =
--- a/src/Tools/code/code_thingol.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/code/code_thingol.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Tools/code/code_thingol.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Intermediate language ("Thin-gol") representing executable code.
@@ -620,7 +619,7 @@
     fun stmt_fun ((vs, raw_ty), raw_thms) =
       let
         val ty = Logic.unvarifyT raw_ty;
-        val thms = if (null o Term.typ_tfrees) ty orelse (null o fst o strip_type) ty
+        val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty
           then raw_thms
           else (map o apfst) (Code_Unit.expand_eta thy 1) raw_thms;
       in
@@ -708,8 +707,7 @@
         let
           val k = length ts;
           val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty;
-          val ctxt = (fold o fold_aterms)
-            (fn Free (v, _) => Name.declare v | _ => I) ts Name.context;
+          val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context;
           val vs = Name.names ctxt "a" tys;
         in
           fold_map (translate_typ thy algbr funcgr) tys
--- a/src/Tools/induct.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/induct.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -440,7 +440,7 @@
     val thy = ProofContext.theory_of ctxt;
     val maxidx = Thm.maxidx_of st;
     val goal = Thm.term_of (Thm.cprem_of st i);  (*exception Subscript*)
-    val params = rev (rename_wrt_term goal (Logic.strip_params goal));
+    val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal));
   in
     if not (null params) then
       (warning ("Cannot determine rule instantiation due to pending parameter(s): " ^
--- a/src/Tools/nbe.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/nbe.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      Tools/nbe.ML
-    ID:         $Id$
     Authors:    Klaus Aehlig, LMU Muenchen; Tobias Nipkow, Florian Haftmann, TU Muenchen
 
 Normalization by evaluation, based on generic code generator.
@@ -448,7 +447,7 @@
       singleton (TypeInfer.infer_types (Syntax.pp_global thy) (Sign.tsig_of thy) I
         (try (Type.strip_sorts o Sign.the_const_type thy)) (K NONE) Name.context 0)
       (TypeInfer.constrain ty t);
-    fun check_tvars t = if null (Term.term_tvars t) then t else
+    fun check_tvars t = if null (Term.add_tvars t []) then t else
       error ("Illegal schematic type variables in normalized term: "
         ^ setmp show_types true (Syntax.string_of_term_global thy) t);
     val string_of_term = setmp show_types true (Syntax.string_of_term_global thy);
--- a/src/Tools/quickcheck.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/quickcheck.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -70,11 +70,11 @@
 
 fun prep_test_term t =
   let
-    val _ = (null (term_tvars t) andalso null (term_tfrees t)) orelse
+    val _ = (null (Term.add_tvars t []) andalso null (Term.add_tfrees t [])) orelse
       error "Term to be tested contains type variables";
-    val _ = null (term_vars t) orelse
+    val _ = null (Term.add_vars t []) orelse
       error "Term to be tested contains schematic variables";
-    val frees = map dest_Free (term_frees t);
+    val frees = map dest_Free (OldTerm.term_frees t);
   in (map fst frees, list_abs_free (frees, t)) end
 
 fun test_term ctxt quiet generator_name size i t =
--- a/src/Tools/value.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/Tools/value.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -20,7 +20,7 @@
   val empty = [];
   val copy = I;
   val extend = I;
-  fun merge pp = AList.merge (op =) (K true);
+  fun merge _ data = AList.merge (op =) (K true) data;
 )
 
 val add_evaluator = Evaluator.map o AList.update (op =);
--- a/src/ZF/Constructible/L_axioms.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/Constructible/L_axioms.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/L_axioms.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -100,7 +99,7 @@
   apply (rule L_nat)
   done
 
-interpretation 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	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/Constructible/Rec_Separation.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Rec_Separation.thy
-    ID:   $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -186,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 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}*}
@@ -373,7 +372,7 @@
   apply (rule M_datatypes_axioms_L) 
   done
 
-interpretation 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}*}
@@ -436,7 +435,7 @@
   apply (rule M_eclose_axioms_L)
   done
 
-interpretation 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	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/Constructible/Separation.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -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 M_basic [L] by (rule M_basic_L)
+interpretation L: M_basic L by (rule M_basic_L)
 
 
 end
--- a/src/ZF/Tools/inductive_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/Tools/inductive_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      ZF/Tools/inductive_package.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1994  University of Cambridge
 
 Fixedpoint definition module -- for Inductive/Coinductive Definitions
 
@@ -101,7 +99,7 @@
                Syntax.string_of_term ctxt t);
 
   (*** Construct the fixedpoint definition ***)
-  val mk_variant = Name.variant (foldr add_term_names [] intr_tms);
+  val mk_variant = Name.variant (foldr OldTerm.add_term_names [] intr_tms);
 
   val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w";
 
@@ -113,7 +111,7 @@
   fun fp_part intr = (*quantify over rule's free vars except parameters*)
     let val prems = map dest_tprop (Logic.strip_imp_prems intr)
         val dummy = List.app (fn rec_hd => List.app (chk_prem rec_hd) prems) rec_hds
-        val exfrees = term_frees intr \\ rec_params
+        val exfrees = OldTerm.term_frees intr \\ rec_params
         val zeq = FOLogic.mk_eq (Free(z',iT), #1 (rule_concl intr))
     in foldr FOLogic.mk_exists
              (BalancedTree.make FOLogic.mk_conj (zeq::prems)) exfrees
@@ -304,7 +302,7 @@
 
      (*Make a premise of the induction rule.*)
      fun induct_prem ind_alist intr =
-       let val quantfrees = map dest_Free (term_frees intr \\ rec_params)
+       let val quantfrees = map dest_Free (OldTerm.term_frees intr \\ rec_params)
            val iprems = foldr (add_induct_prem ind_alist) []
                               (Logic.strip_imp_prems intr)
            val (t,X) = Ind_Syntax.rule_concl intr
--- a/src/ZF/Tools/primrec_package.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/Tools/primrec_package.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,10 +1,9 @@
 (*  Title:      ZF/Tools/primrec_package.ML
-    ID:         $Id$
-    Author:     Stefan Berghofer and Norbert Voelker
-    Copyright   1998  TU Muenchen
-    ZF version by Lawrence C Paulson (Cambridge)
+    Author:     Norbert Voelker, FernUni Hagen
+    Author:     Stefan Berghofer, TU Muenchen
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 
-Package for defining functions on datatypes by primitive recursion
+Package for defining functions on datatypes by primitive recursion.
 *)
 
 signature PRIMREC_PACKAGE =
@@ -33,7 +32,7 @@
 fun process_eqn thy (eq, rec_fn_opt) =
   let
     val (lhs, rhs) =
-        if null (term_vars eq) then
+        if null (Term.add_vars eq []) then
             dest_eqn eq handle TERM _ => raise RecError "not a proper equation"
         else raise RecError "illegal schematic variable(s)";
 
@@ -66,7 +65,7 @@
   in
     if has_duplicates (op =) lfrees then
       raise RecError "repeated variable name in pattern"
-    else if not ((map dest_Free (term_frees rhs)) subset lfrees) then
+    else if not ((map dest_Free (OldTerm.term_frees rhs)) subset lfrees) then
       raise RecError "extra variables on rhs"
     else if length middle > 1 then
       raise RecError "more than one non-variable in pattern"
--- a/src/ZF/arith_data.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/arith_data.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      ZF/arith_data.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2000  University of Cambridge
 
 Arithmetic simplification: cancellation of common terms
 *)
@@ -106,7 +104,7 @@
 
 (*Dummy version: the "coefficient" is always 1.
   In the result, the factors are sorted terms*)
-fun dest_coeff t = (1, mk_prod (sort Term.term_ord (dest_prod t)));
+fun dest_coeff t = (1, mk_prod (sort TermOrd.term_ord (dest_prod t)));
 
 (*Find first coefficient-term THAT MATCHES u*)
 fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
--- a/src/ZF/ex/Group.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/ex/Group.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -1,5 +1,4 @@
 (* Title:  ZF/ex/Group.thy
-  Id:     $Id$
 
 *)
 
@@ -40,7 +39,7 @@
   m_inv :: "i => i => i" ("inv\<index> _" [81] 80) where
   "inv\<^bsub>G\<^esub> x == (THE y. y \<in> carrier(G) & y \<cdot>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub> & x \<cdot>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub>)"
 
-locale monoid = struct G +
+locale monoid = fixes G (structure)
   assumes m_closed [intro, simp]:
          "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G)\<rbrakk> \<Longrightarrow> x \<cdot> y \<in> carrier(G)"
       and m_assoc:
@@ -242,7 +241,7 @@
 
 subsection {* Substructures *}
 
-locale subgroup = var H + struct G + 
+locale subgroup = fixes H and G (structure)
   assumes subset: "H \<subseteq> carrier(G)"
     and m_closed [intro, simp]: "\<lbrakk>x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> x \<cdot> y \<in> H"
     and  one_closed [simp]: "\<one> \<in> H"
@@ -262,7 +261,7 @@
   assumes "group(G)"
   shows "group_axioms (update_carrier(G,H))"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis by (force intro: group_axioms.intro l_inv r_inv)
 qed
 
@@ -270,7 +269,7 @@
   assumes "group(G)"
   shows "group (update_carrier(G,H))"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis
   by (rule groupI) (auto intro: m_assoc l_inv mem_carrier)
 qed
@@ -316,8 +315,8 @@
   assumes "group(G)" and "group(H)"
   shows "group (G \<Otimes> 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 (force intro!: groupI G.m_assoc H.m_assoc G.l_inv H.l_inv
           simp add: DirProdGroup_def)
 qed
@@ -397,8 +396,8 @@
   assumes "group(G)" and "group(H)"
   shows "(\<lambda><x,y> \<in> carrier(G \<Otimes> H). <y,x>) \<in> (G \<Otimes> H) \<cong> (H \<Otimes> G)"
 proof -
-  interpret group [G] by fact
-  interpret group [H] by fact
+  interpret group G by fact
+  interpret group H by fact
   show ?thesis by (auto simp add: iso_def hom_def inj_def surj_def bij_def)
 qed
 
@@ -407,16 +406,17 @@
   shows "(\<lambda><<x,y>,z> \<in> carrier((G \<Otimes> H) \<Otimes> I). <x,<y,z>>)
           \<in> ((G \<Otimes> H) \<Otimes> I) \<cong> (G \<Otimes> (H \<Otimes> I))"
 proof -
-  interpret group [G] by fact
-  interpret group [H] by fact
-  interpret group [I] by fact
+  interpret group G by fact
+  interpret group H by fact
+  interpret group I by fact
   show ?thesis
     by (auto intro: lam_type simp add: iso_def hom_def inj_def surj_def bij_def) 
 qed
 
 text{*Basis for homomorphism proofs: we assume two groups @{term G} and
   @term{H}, with a homomorphism @{term h} between them*}
-locale group_hom = group G + group H + var h +
+locale group_hom = G: group G + H: group H
+  for G (structure) and H (structure) and h +
   assumes homh: "h \<in> hom(G,H)"
   notes hom_mult [simp] = hom_mult [OF homh]
     and hom_closed [simp] = hom_closed [OF homh]
@@ -870,7 +870,7 @@
    assumes "group(G)"
    shows "equiv (carrier(G), rcong H)"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis proof (simp add: equiv_def, intro conjI)
     show "rcong H \<subseteq> carrier(G) \<times> carrier(G)"
       by (auto simp add: r_congruent_def) 
@@ -907,7 +907,7 @@
   assumes a: "a \<in> carrier(G)"
   shows "a <# H = (rcong H) `` {a}" 
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   show ?thesis
     by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a
       Collect_image_eq) 
@@ -920,7 +920,7 @@
         h \<in> H;  ha \<in> H;  hb \<in> H\<rbrakk>
       \<Longrightarrow> hb \<cdot> a \<in> (\<Union>h\<in>H. {h \<cdot> b})" (is "PROP ?P")
 proof -
-  interpret subgroup [H G] by fact
+  interpret subgroup H G by fact
   show "PROP ?P"
     apply (rule UN_I [of "hb \<cdot> ((inv ha) \<cdot> h)"], simp)
     apply (simp add: m_assoc transpose_inv)
@@ -931,7 +931,7 @@
   assumes "subgroup(H, G)"
   shows "\<lbrakk>a \<in> rcosets H; b \<in> rcosets H; a\<noteq>b\<rbrakk> \<Longrightarrow> a \<inter> b = 0" (is "PROP ?P")
 proof -
-  interpret subgroup [H G] by fact
+  interpret subgroup H G by fact
   show "PROP ?P"
     apply (simp add: RCOSETS_def r_coset_def)
     apply (blast intro: rcos_equation prems sym)
@@ -949,7 +949,7 @@
   assumes "subgroup(H, G)"
   shows "x \<in> carrier(G) \<Longrightarrow> x \<in> H #> x" (is "PROP ?P")
 proof -
-  interpret subgroup [H G] by fact
+  interpret subgroup H G by fact
   show "PROP ?P"
     apply (simp add: r_coset_def)
     apply (rule_tac x="\<one>" in bexI) apply (auto)
@@ -960,7 +960,7 @@
   assumes "subgroup(H, G)"
   shows "\<Union>(rcosets H) = carrier(G)"
 proof -
-  interpret subgroup [H G] by fact
+  interpret subgroup H G by fact
   show ?thesis
     apply (rule equalityI)
     apply (force simp add: RCOSETS_def r_coset_def)
@@ -1044,7 +1044,7 @@
   assumes "group(G)"
   shows "H \<in> rcosets H"
 proof -
-  interpret group [G] by fact
+  interpret group G by fact
   have "H #> \<one> = H"
     using _ subgroup_axioms by (rule coset_join2) simp_all
   then show ?thesis
--- a/src/ZF/ex/Ring.thy	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/ex/Ring.thy	Mon Jan 05 07:54:16 2009 -0800
@@ -45,7 +45,7 @@
   minus :: "[i,i,i] => i" (infixl "\<ominus>\<index>" 65) where
   "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"
 
-locale abelian_monoid = struct G +
+locale abelian_monoid = fixes G (structure)
   assumes a_comm_monoid: 
     "comm_monoid (<carrier(G), add_field(G), zero(G), 0>)"
 
@@ -57,7 +57,7 @@
   assumes a_comm_group: 
     "comm_group (<carrier(G), add_field(G), zero(G), 0>)"
 
-locale ring = abelian_group R + monoid R +
+locale ring = abelian_group R + monoid R for R (structure) +
   assumes l_distr: "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R); z \<in> carrier(R)\<rbrakk>
       \<Longrightarrow> (x \<oplus> y) \<cdot> z = x \<cdot> z \<oplus> y \<cdot> z"
     and r_distr: "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R); z \<in> carrier(R)\<rbrakk>
@@ -262,7 +262,8 @@
 lemma ring_hom_one: "h \<in> ring_hom(R,S) \<Longrightarrow> h ` \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>"
 by (simp add: ring_hom_def)
 
-locale ring_hom_cring = cring R + cring S + var h +
+locale ring_hom_cring = R: cring R + S: cring S
+  for R (structure) and S (structure) and h +
   assumes homh [simp, intro]: "h \<in> ring_hom(R,S)"
   notes hom_closed [simp, intro] = ring_hom_closed [OF homh]
     and hom_mult [simp] = ring_hom_mult [OF homh]
--- a/src/ZF/int_arith.ML	Mon Dec 29 11:04:27 2008 -0800
+++ b/src/ZF/int_arith.ML	Mon Jan 05 07:54:16 2009 -0800
@@ -1,7 +1,5 @@
 (*  Title:      ZF/int_arith.ML
-    ID:         $Id$
     Author:     Larry Paulson
-    Copyright   2000  University of Cambridge
 
 Simprocs for linear arithmetic.
 *)
@@ -72,7 +70,7 @@
 (*Express t as a product of (possibly) a numeral with other sorted terms*)
 fun dest_coeff sign (Const (@{const_name "zminus"}, _) $ t) = dest_coeff (~sign) t
   | dest_coeff sign t =
-    let val ts = sort Term.term_ord (dest_prod t)
+    let val ts = sort TermOrd.term_ord (dest_prod t)
         val (n, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, ts)
     in (sign*n, mk_prod ts') end;