merged
authorpaulson
Wed, 01 Jul 2009 16:19:44 +0100
changeset 31909 d3b020134006
parent 31898 82d5190ff7c8 (diff)
parent 31908 67224d7d4448 (current diff)
child 31910 a8e9ccfc427a
merged
--- a/Admin/build	Wed Jun 24 16:28:30 2009 +0100
+++ b/Admin/build	Wed Jul 01 16:19:44 2009 +0100
@@ -8,6 +8,9 @@
 PATH="/usr/local/dist/DIR/j2sdk1.5.0/bin:$PATH"
 
 PATH="/home/scala/current/bin:$PATH"
+if [ -z "$SCALA_HOME" ]; then
+  export SCALA_HOME="$(dirname "$(dirname "$(type -p scalac)")")"
+fi
 
 
 ## directory layout
@@ -32,7 +35,7 @@
     all             all modules below
     browser         graph browser (requires jdk)
     doc             documentation (requires latex and rail)
-    jars            JVM components (requires jdk and scala)
+    jars            Scala/JVM components (requires scala)
 
 EOF
   exit 1
@@ -93,13 +96,13 @@
 function build_jars ()
 {
   echo "###"
-  echo "### Building JVM components ..."
+  echo "### Building Scala/JVM components ..."
   echo "###"
 
-  type -p scalac >/dev/null || fail "Scala compiler unavailable"
+  [ -z "$SCALA_HOME" ] && fail "Scala unavailable: unknown SCALA_HOME"
 
   pushd "$ISABELLE_HOME/src/Pure" >/dev/null
-  "$ISABELLE_TOOL" make jar || fail "Failed to build Pure.jar!"
+  "$ISABELLE_TOOL" make jars || fail "Failed to build isabelle-scala.jar"
   popd >/dev/null
 }
 
--- a/Admin/isatest/isatest-makedist	Wed Jun 24 16:28:30 2009 +0100
+++ b/Admin/isatest/isatest-makedist	Wed Jul 01 16:19:44 2009 +0100
@@ -106,7 +106,7 @@
 sleep 15
 $SSH atbroy101 "$MAKEALL $HOME/settings/at64-poly"
 sleep 15
-$SSH macbroy2 "$MAKEALL $HOME/settings/mac-poly-M4; $MAKEALL $HOME/settings/mac-poly-M8"
+$SSH macbroy2 "$MAKEALL $HOME/settings/mac-poly-M4; $MAKEALL $HOME/settings/mac-poly-M8; $MAKEALL $HOME/settings/mac-poly64-M4"
 sleep 15
 $SSH macbroy5 "$MAKEALL $HOME/settings/mac-poly"
 sleep 15
--- a/Admin/isatest/isatest-stats	Wed Jun 24 16:28:30 2009 +0100
+++ b/Admin/isatest/isatest-stats	Wed Jul 01 16:19:44 2009 +0100
@@ -24,6 +24,7 @@
   HOL-MetisExamples \
   HOL-MicroJava \
   HOL-NSA \
+  HOL-NewNumberTheory \
   HOL-Nominal-Examples \
   HOL-NumberTheory \
   HOL-SET-Protocol \
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/isatest/settings/mac-poly64-M4	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,28 @@
+# -*- shell-script -*- :mode=shellscript:
+
+  POLYML_HOME="/home/polyml/polyml-svn"
+  ML_SYSTEM="polyml-experimental"
+  ML_PLATFORM="x86_64-darwin"
+  ML_HOME="$POLYML_HOME/$ML_PLATFORM"
+  ML_OPTIONS="--mutable 2000 --immutable 2000"
+
+
+ISABELLE_HOME_USER=~/isabelle-mac-poly64-M4
+
+# Where to look for isabelle tools (multiple dirs separated by ':').
+ISABELLE_TOOLS="$ISABELLE_HOME/lib/Tools"
+
+# Location for temporary files (should be on a local file system).
+ISABELLE_TMP_PREFIX="/tmp/isabelle-$USER"
+
+
+# Heap input locations. ML system identifier is included in lookup.
+ISABELLE_PATH="$ISABELLE_HOME_USER/heaps:$ISABELLE_HOME/heaps"
+
+# Heap output location. ML system identifier is appended automatically later on.
+ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps"
+ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info"
+
+ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4"
+
+HOL_USEDIR_OPTIONS="-p 2 -Q false"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/launch4j/README	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,5 @@
+Cross-platform Java executable wrapper
+======================================
+
+* http://launch4j.sourceforge.net
+
Binary file Admin/launch4j/isabelle.ico has changed
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/launch4j/isabelle.xml	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,23 @@
+<launch4jConfig>
+  <dontWrapJar>true</dontWrapJar>
+  <headerType>gui</headerType>
+  <jar>lib/classes/isabelle-scala.jar</jar>
+  <outfile>Isabelle.exe</outfile>
+  <errTitle></errTitle>
+  <cmdLine></cmdLine>
+  <chdir></chdir>
+  <priority>normal</priority>
+  <downloadUrl>http://java.com/download</downloadUrl>
+  <supportUrl></supportUrl>
+  <customProcName>false</customProcName>
+  <stayAlive>false</stayAlive>
+  <manifest></manifest>
+  <icon>isabelle.ico</icon>
+  <jre>
+    <path></path>
+    <minVersion>1.6.0</minVersion>
+    <maxVersion></maxVersion>
+    <jdkPreference>preferJre</jdkPreference>
+    <opt>-Disabelle.home=&quot;%EXEDIR%&quot;</opt>
+  </jre>
+</launch4jConfig>
\ No newline at end of file
--- a/Admin/makedist	Wed Jun 24 16:28:30 2009 +0100
+++ b/Admin/makedist	Wed Jul 01 16:19:44 2009 +0100
@@ -116,7 +116,6 @@
 DISTBASE="$DISTPREFIX/dist-$DISTNAME"
 mkdir -p "$DISTBASE" || { purge_tmp; fail "Unable to create distribution base dir $DISTBASE!"; }
 [ -e "$DISTBASE/$DISTNAME" ] && { purge_tmp; fail "$DISTBASE/$DISTNAME already exists!"; }
-[ -e "$DISTBASE/pdf/$DISTNAME" ] && { purge_tmp; fail "$DISTBASE/pdf/$DISTNAME already exists!"; }
 
 cd "$DISTBASE"
 mv "$DISTPREFIX/$TMP/isabelle-$IDENT" "$DISTNAME"
@@ -141,7 +140,7 @@
 MOVE=$(find doc-src \( -type f -a -not -type l -a -not -name isabelle_isar.pdf -a -not -name pghead.pdf -a \( -name \*.dvi -o -name \*.eps -o -name \*.ps -o -name \*.pdf \) -a -print \) | grep -v 'gfx/.*pdf')
 mv -f $MOVE doc
 rm doc/Isa-logics.eps
-rm doc/adaption.dvi doc/adaption.pdf doc/architecture.dvi doc/architecture.pdf
+rm doc/adaptation.dvi doc/adaptation.pdf doc/architecture.dvi doc/architecture.pdf
 rm -rf doc-src
 
 mkdir -p contrib
@@ -189,18 +188,9 @@
 chmod -R g=o "$DISTNAME"
 chgrp -R isabelle "$DISTNAME" Isabelle
 
-mkdir -p "pdf/$DISTNAME/doc"
-mv "$DISTNAME/doc/"*.pdf "pdf/$DISTNAME/doc"
-
 echo "$DISTNAME.tar.gz"
 tar -czf "$DISTNAME.tar.gz" Isabelle "$DISTNAME"
 
-echo "${DISTNAME}_pdf.tar.gz"
-tar -C pdf -czf "${DISTNAME}_pdf.tar.gz" "$DISTNAME"
-
-mv "pdf/$DISTNAME/doc/"*.pdf "$DISTNAME/doc"
-rmdir "pdf/$DISTNAME/doc" "pdf/$DISTNAME" pdf
-
 
 # cleanup dist
 
--- a/NEWS	Wed Jun 24 16:28:30 2009 +0100
+++ b/NEWS	Wed Jul 01 16:19:44 2009 +0100
@@ -48,6 +48,10 @@
 Suc_eq_add_numeral_1_left -> Suc_eq_plus1_left
 Suc_plus1 -> Suc_eq_plus1
 
+* New sledgehammer option "Full Types" in Proof General settings menu.
+Causes full type information to be output to the ATPs. This slows ATPs down
+considerably but eliminates a source of unsound "proofs" that fail later.
+
 * Discontinued ancient tradition to suffix certain ML module names with "_package", e.g.:
 
     DatatypePackage ~> Datatype
@@ -62,6 +66,23 @@
 
 * Simplified interfaces of datatype module.  INCOMPATIBILITY.
 
+* Abbreviation "arbitrary" of "undefined" has disappeared; use "undefined" directly.
+INCOMPATIBILITY.
+
+* New evaluator "approximate" approximates an real valued term using the same method as the
+approximation method. 
+
+* "approximate" supports now arithmetic expressions as boundaries of intervals and implements
+interval splitting and taylor series expansion.
+
+* Changed DERIV_intros to a NamedThmsFun. Each of the theorems in DERIV_intros
+assumes composition with an additional function and matches a variable to the
+derivative, which has to be solved by the simplifier. Hence
+(auto intro!: DERIV_intros) computes the derivative of most elementary terms.
+
+* Maclauren.DERIV_tac and Maclauren.deriv_tac was removed, they are replaced by:
+(auto intro!: DERIV_intros)
+INCOMPATIBILITY.
 
 *** ML ***
 
--- a/bin/isabelle-process	Wed Jun 24 16:28:30 2009 +0100
+++ b/bin/isabelle-process	Wed Jul 01 16:19:44 2009 +0100
@@ -214,7 +214,7 @@
 NICE="nice"
 if [ -n "$WRAPPER_OUTPUT" ]; then
   [ "$WRAPPER_OUTPUT" = "-" -o -p "$WRAPPER_OUTPUT" ] || fail "Bad named pipe: $WRAPPER_OUTPUT"
-  MLTEXT="$MLTEXT; IsabelleProcess.init \"$WRAPPER_OUTPUT\";"
+  MLTEXT="$MLTEXT; Isabelle_Process.init \"$WRAPPER_OUTPUT\";"
 elif [ -n "$PGIP" ]; then
   MLTEXT="$MLTEXT; ProofGeneralPgip.init_pgip $ISAR;"
 elif [ -n "$PROOFGENERAL" ]; then
--- a/doc-src/Codegen/Thy/document/Introduction.tex	Wed Jun 24 16:28:30 2009 +0100
+++ b/doc-src/Codegen/Thy/document/Introduction.tex	Wed Jul 01 16:19:44 2009 +0100
@@ -249,9 +249,9 @@
 \hspace*{0pt}dequeue (AQueue [] []) = (Nothing,~AQueue [] []);\\
 \hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\
 \hspace*{0pt}dequeue (AQueue (v :~va) []) =\\
-\hspace*{0pt} ~(let {\char123}\\
+\hspace*{0pt} ~let {\char123}\\
 \hspace*{0pt} ~~~(y :~ys) = rev (v :~va);\\
-\hspace*{0pt} ~{\char125}~in (Just y,~AQueue [] ys) );\\
+\hspace*{0pt} ~{\char125}~in (Just y,~AQueue [] ys);\\
 \hspace*{0pt}\\
 \hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\
 \hspace*{0pt}enqueue x (AQueue xs ys) = AQueue (x :~xs) ys;\\
--- a/doc-src/Codegen/Thy/document/Program.tex	Wed Jun 24 16:28:30 2009 +0100
+++ b/doc-src/Codegen/Thy/document/Program.tex	Wed Jul 01 16:19:44 2009 +0100
@@ -966,9 +966,9 @@
 \noindent%
 \hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\
 \hspace*{0pt}strict{\char95}dequeue (AQueue xs []) =\\
-\hspace*{0pt} ~(let {\char123}\\
+\hspace*{0pt} ~let {\char123}\\
 \hspace*{0pt} ~~~(y :~ys) = rev xs;\\
-\hspace*{0pt} ~{\char125}~in (y,~AQueue [] ys) );\\
+\hspace*{0pt} ~{\char125}~in (y,~AQueue [] ys);\\
 \hspace*{0pt}strict{\char95}dequeue (AQueue xs (y :~ys)) = (y,~AQueue xs ys);%
 \end{isamarkuptext}%
 \isamarkuptrue%
--- a/doc-src/Codegen/Thy/examples/Example.hs	Wed Jun 24 16:28:30 2009 +0100
+++ b/doc-src/Codegen/Thy/examples/Example.hs	Wed Jul 01 16:19:44 2009 +0100
@@ -23,9 +23,9 @@
 dequeue (AQueue [] []) = (Nothing, AQueue [] []);
 dequeue (AQueue xs (y : ys)) = (Just y, AQueue xs ys);
 dequeue (AQueue (v : va) []) =
-  (let {
+  let {
     (y : ys) = rev (v : va);
-  } in (Just y, AQueue [] ys) );
+  } in (Just y, AQueue [] ys);
 
 enqueue :: forall a. a -> Queue a -> Queue a;
 enqueue x (AQueue xs ys) = AQueue (x : xs) ys;
--- a/doc-src/IsarImplementation/Thy/Proof.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/doc-src/IsarImplementation/Thy/Proof.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -93,7 +93,7 @@
   @{index_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\
   @{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\
   @{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
-  @{index_ML Variable.import_thms: "bool -> thm list -> Proof.context ->
+  @{index_ML Variable.import: "bool -> thm list -> Proof.context ->
   ((ctyp list * cterm list) * thm list) * Proof.context"} \\
   @{index_ML Variable.focus: "cterm -> Proof.context -> (cterm list * cterm) * Proof.context"} \\
   \end{mldecls}
@@ -132,7 +132,7 @@
   with @{ML Variable.polymorphic}: here the given terms are detached
   from the context as far as possible.
 
-  \item @{ML Variable.import_thms}~@{text "open thms ctxt"} invents fixed
+  \item @{ML Variable.import}~@{text "open thms ctxt"} invents fixed
   type and term variables for the schematic ones occurring in @{text
   "thms"}.  The @{text "open"} flag indicates whether the fixed names
   should be accessible to the user, otherwise newly introduced names
--- a/doc-src/IsarImplementation/Thy/document/Proof.tex	Wed Jun 24 16:28:30 2009 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Proof.tex	Wed Jul 01 16:19:44 2009 +0100
@@ -111,7 +111,7 @@
   \indexdef{}{ML}{Variable.declare\_constraints}\verb|Variable.declare_constraints: term -> Proof.context -> Proof.context| \\
   \indexdef{}{ML}{Variable.export}\verb|Variable.export: Proof.context -> Proof.context -> thm list -> thm list| \\
   \indexdef{}{ML}{Variable.polymorphic}\verb|Variable.polymorphic: Proof.context -> term list -> term list| \\
-  \indexdef{}{ML}{Variable.import\_thms}\verb|Variable.import_thms: bool -> thm list -> Proof.context ->|\isasep\isanewline%
+  \indexdef{}{ML}{Variable.import}\verb|Variable.import: bool -> thm list -> Proof.context ->|\isasep\isanewline%
 \verb|  ((ctyp list * cterm list) * thm list) * Proof.context| \\
   \indexdef{}{ML}{Variable.focus}\verb|Variable.focus: cterm -> Proof.context -> (cterm list * cterm) * Proof.context| \\
   \end{mldecls}
@@ -149,7 +149,7 @@
   with \verb|Variable.polymorphic|: here the given terms are detached
   from the context as far as possible.
 
-  \item \verb|Variable.import_thms|~\isa{open\ thms\ ctxt} invents fixed
+  \item \verb|Variable.import|~\isa{open\ thms\ ctxt} invents fixed
   type and term variables for the schematic ones occurring in \isa{thms}.  The \isa{open} flag indicates whether the fixed names
   should be accessible to the user, otherwise newly introduced names
   are marked as ``internal'' (\secref{sec:names}).
--- a/doc-src/more_antiquote.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/doc-src/more_antiquote.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -81,7 +81,7 @@
 fun no_vars ctxt thm =
   let
     val ctxt' = Variable.set_body false ctxt;
-    val ((_, [thm]), _) = Variable.import_thms true [thm] ctxt';
+    val ((_, [thm]), _) = Variable.import true [thm] ctxt';
   in thm end;
 
 fun pretty_code_thm src ctxt raw_const =
--- a/etc/symbols	Wed Jun 24 16:28:30 2009 +0100
+++ b/etc/symbols	Wed Jul 01 16:19:44 2009 +0100
@@ -211,8 +211,8 @@
 \<or>                   code: 0x002228  font: Isabelle  abbrev: \/
 \<Or>                   code: 0x0022c1  font: Isabelle  abbrev: ??
 \<forall>               code: 0x002200  font: Isabelle  abbrev: !
-\<exists>               code: 0x002203  font: Isabelle  abbrev: ?.
-\<nexists>              code: 0x002204  font: Isabelle  abbrev: ?~
+\<exists>               code: 0x002203  font: Isabelle  abbrev: ?
+\<nexists>              code: 0x002204  font: Isabelle  abbrev: ~?
 \<not>                  code: 0x0000ac  font: Isabelle  abbrev: ~
 \<box>                  code: 0x0025a1  font: Isabelle
 \<diamond>              code: 0x0025c7  font: Isabelle
--- a/lib/scripts/SystemOnTPTP	Wed Jun 24 16:28:30 2009 +0100
+++ b/lib/scripts/SystemOnTPTP	Wed Jul 01 16:19:44 2009 +0100
@@ -10,20 +10,20 @@
 use HTTP::Request::Common;
 use LWP;
 
-my $SystemOnTPTPFormReplyURL = "http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTPFormReply";
+my $SystemOnTPTPFormReplyURL =
+  "http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTPFormReply";
 
 # default parameters
 my %URLParameters = (
     "NoHTML" => 1,
     "QuietFlag" => "-q01",
-    "X2TPTP" => "-S",
     "SubmitButton" => "RunSelectedSystems",
     "ProblemSource" => "UPLOAD",
     );
 
 #----Get format and transform options if specified
 my %Options;
-getopts("hws:t:c:",\%Options);
+getopts("hwxs:t:c:",\%Options);
 
 #----Usage
 sub usage() {
@@ -31,6 +31,7 @@
   print("    <options> are ...\n");
   print("    -h            - print this help\n");
   print("    -w            - list available ATP systems\n");
+  print("    -x            - use X2TPTP to convert output of prover\n");
   print("    -s<system>    - specified system to use\n");
   print("    -t<timelimit> - CPU time limit for system\n");
   print("    -c<command>   - custom command for system\n");
@@ -40,11 +41,18 @@
 if (exists($Options{'h'})) {
   usage();
 }
+
 #----What systems flag
 if (exists($Options{'w'})) {
     $URLParameters{"SubmitButton"} = "ListSystems";
     delete($URLParameters{"ProblemSource"});
 }
+
+#----X2TPTP
+if (exists($Options{'x'})) {
+    $URLParameters{"X2TPTP"} = "-S";
+}
+
 #----Selected system
 my $System;
 if (exists($Options{'s'})) {
@@ -86,7 +94,7 @@
 my $Response = $Agent->request($Request);
 
 #catch errors / failure
-if(! $Response->is_success){
+if(!$Response->is_success) {
   print "HTTP-Error: " . $Response->message . "\n";
   exit(-1);
 } elsif (exists($Options{'w'})) {
@@ -95,7 +103,12 @@
 } elsif ($Response->content =~ /WARNING: (\S*) does not exist/) {
   print "Specified System $1 does not exist\n";
   exit(-1);
-} elsif ($Response->content =~ /%\s*Result\s*:\s*Unsatisfiable.*\n%\s*Output\s*:\s*(CNF)?Refutation.*\n%/) {
+} elsif (exists($Options{'x'}) &&
+  $Response->content =~
+    /%\s*Result\s*:\s*Unsatisfiable.*\n%\s*Output\s*:\s*(CNF)?Refutation.*\n%/ &&
+  $Response->content !~ /ERROR: Could not form TPTP format derivation/ )
+{
+  # converted output: extract proof
   my @lines = split( /\n/, $Response->content);
   my $extract = "";
   foreach my $line (@lines){
@@ -108,12 +121,20 @@
   $extract =~ s/\s//g;
   $extract =~ s/\)\.cnf/\)\.\ncnf/g;
 
+  print "========== ~~/lib/scripts/SystemOnTPTP extracted proof: ==========\n";
   # orientation for res_reconstruct.ML
   print "# SZS output start CNFRefutation.\n";
   print "$extract\n";
   print "# SZS output end CNFRefutation.\n";
+  # can be useful for debugging; Isabelle ignores this
+  print "============== original response from SystemOnTPTP: ==============\n";
+  print $Response->content;
   exit(0);
-} else {
+} elsif (!exists($Options{'x'})) {
+  # pass output directly to Isabelle
+  print $Response->content;
+  exit(0);
+}else {
   print "Remote-script could not extract proof:\n".$Response->content;
   exit(-1);
 }
--- a/src/HOL/ATP_Linkup.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/ATP_Linkup.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -115,11 +115,11 @@
 
 text {* remote provers via SystemOnTPTP *}
 setup {* AtpManager.add_prover "remote_vampire"
-  (AtpWrapper.remote_prover "-s Vampire---9.0") *}
+  (AtpWrapper.remote_prover_opts 60 false "" "Vampire---9") *}
 setup {* AtpManager.add_prover "remote_spass"
-  (AtpWrapper.remote_prover "-s SPASS---3.01") *}
+  (AtpWrapper.remote_prover_opts 40 true "-x" "SPASS---") *}
 setup {* AtpManager.add_prover "remote_e"
-  (AtpWrapper.remote_prover "-s EP---1.0") *}
+  (AtpWrapper.remote_prover_opts 100 false "" "EP---") *}
   
 
 
--- a/src/HOL/Decision_Procs/Approximation.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -38,7 +38,7 @@
   and lb_Suc: "\<And> n i k x. lb (Suc n) i k x = lapprox_rat prec 1 (int k) - x * (ub n (F i) (G i k) x)"
   and ub_0: "\<And> i k x. ub 0 i k x = 0"
   and ub_Suc: "\<And> n i k x. ub (Suc n) i k x = rapprox_rat prec 1 (int k) - x * (lb n (F i) (G i k) x)"
-  shows "real (lb n ((F ^^ j') s) (f j') x) \<le> horner F G n ((F ^^ j') s) (f j') (real x) \<and> 
+  shows "real (lb n ((F ^^ j') s) (f j') x) \<le> horner F G n ((F ^^ j') s) (f j') (real x) \<and>
          horner F G n ((F ^^ j') s) (f j') (real x) \<le> real (ub n ((F ^^ j') s) (f j') x)"
   (is "?lb n j' \<le> ?horner n j' \<and> ?horner n j' \<le> ?ub n j'")
 proof (induct n arbitrary: j')
@@ -56,7 +56,7 @@
   proof (rule add_mono)
     show "1 / real (f j') \<le> real (rapprox_rat prec 1 (int (f j')))" using rapprox_rat[of 1 "int (f j')" prec] by auto
     from Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc, THEN conjunct1] `0 \<le> real x`
-    show "- (real x * horner F G n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) (real x)) \<le> 
+    show "- (real x * horner F G n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) (real x)) \<le>
           - real (x * lb n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x)"
       unfolding real_of_float_mult neg_le_iff_le by (rule mult_left_mono)
   qed
@@ -78,10 +78,10 @@
   and lb_Suc: "\<And> n i k x. lb (Suc n) i k x = lapprox_rat prec 1 (int k) - x * (ub n (F i) (G i k) x)"
   and ub_0: "\<And> i k x. ub 0 i k x = 0"
   and ub_Suc: "\<And> n i k x. ub (Suc n) i k x = rapprox_rat prec 1 (int k) - x * (lb n (F i) (G i k) x)"
-  shows "real (lb n ((F ^^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. -1 ^ j * (1 / real (f (j' + j))) * real x ^ j)" (is "?lb") and 
+  shows "real (lb n ((F ^^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. -1 ^ j * (1 / real (f (j' + j))) * real x ^ j)" (is "?lb") and
     "(\<Sum>j=0..<n. -1 ^ j * (1 / real (f (j' + j))) * (real x ^ j)) \<le> real (ub n ((F ^^ j') s) (f j') x)" (is "?ub")
 proof -
-  have "?lb  \<and> ?ub" 
+  have "?lb  \<and> ?ub"
     using horner_bounds'[where lb=lb, OF `0 \<le> real x` f_Suc lb_0 lb_Suc ub_0 ub_Suc]
     unfolding horner_schema[where f=f, OF f_Suc] .
   thus "?lb" and "?ub" by auto
@@ -93,7 +93,7 @@
   and lb_Suc: "\<And> n i k x. lb (Suc n) i k x = lapprox_rat prec 1 (int k) + x * (ub n (F i) (G i k) x)"
   and ub_0: "\<And> i k x. ub 0 i k x = 0"
   and ub_Suc: "\<And> n i k x. ub (Suc n) i k x = rapprox_rat prec 1 (int k) + x * (lb n (F i) (G i k) x)"
-  shows "real (lb n ((F ^^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. (1 / real (f (j' + j))) * real x ^ j)" (is "?lb") and 
+  shows "real (lb n ((F ^^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. (1 / real (f (j' + j))) * real x ^ j)" (is "?lb") and
     "(\<Sum>j=0..<n. (1 / real (f (j' + j))) * (real x ^ j)) \<le> real (ub n ((F ^^ j') s) (f j') x)" (is "?ub")
 proof -
   { fix x y z :: float have "x - y * z = x + - y * z"
@@ -104,7 +104,7 @@
 
   have move_minus: "real (-x) = -1 * real x" by auto
 
-  have sum_eq: "(\<Sum>j=0..<n. (1 / real (f (j' + j))) * real x ^ j) = 
+  have sum_eq: "(\<Sum>j=0..<n. (1 / real (f (j' + j))) * real x ^ j) =
     (\<Sum>j = 0..<n. -1 ^ j * (1 / real (f (j' + j))) * real (- x) ^ j)"
   proof (rule setsum_cong, simp)
     fix j assume "j \<in> {0 ..< n}"
@@ -174,7 +174,7 @@
     show ?thesis
     proof (cases "u < 0")
       case True hence "0 \<le> - real u" and "- real u \<le> - x" and "0 \<le> - x" and "-x \<le> - real l" using assms unfolding less_float_def by auto
-      hence "real u ^ n \<le> x ^ n" and "x ^ n \<le> real l ^ n" using power_mono[of  "-x" "-real l" n] power_mono[of "-real u" "-x" n] 
+      hence "real u ^ n \<le> x ^ n" and "x ^ n \<le> real l ^ n" using power_mono[of  "-x" "-real l" n] power_mono[of "-real u" "-x" n]
 	unfolding power_minus_even[OF `even n`] by auto
       moreover have u1: "u1 = l ^ n" and l1: "l1 = u ^ n" using assms unfolding float_power_bnds_def if_not_P[OF P] if_P[OF True] by auto
       ultimately show ?thesis using float_power by auto
@@ -315,7 +315,7 @@
   shows "0 \<le> real (lb_sqrt prec x)"
 proof (cases "0 < x")
   case True hence "0 < real x" and "0 \<le> x" using `0 \<le> real x` unfolding less_float_def le_float_def by auto
-  hence "0 < sqrt_iteration prec prec x" unfolding less_float_def using sqrt_iteration_lower_bound by auto 
+  hence "0 < sqrt_iteration prec prec x" unfolding less_float_def using sqrt_iteration_lower_bound by auto
   hence "0 \<le> real (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF `0 \<le> x`] unfolding le_float_def by auto
   thus ?thesis unfolding lb_sqrt.simps using True by auto
 next
@@ -336,7 +336,7 @@
     also have "\<dots> < real x / sqrt (real x)"
       by (rule divide_strict_left_mono[OF sqrt_ub `0 < real x`
                mult_pos_pos[OF order_less_trans[OF sqrt_gt0 sqrt_ub] sqrt_gt0]])
-    also have "\<dots> = sqrt (real x)" 
+    also have "\<dots> = sqrt (real x)"
       unfolding inverse_eq_iff_eq[of _ "sqrt (real x)", symmetric]
 	        sqrt_divide_self_eq[OF `0 \<le> real x`, symmetric] by auto
     finally have "real (lb_sqrt prec x) \<le> sqrt (real x)"
@@ -357,7 +357,7 @@
     case True with lb ub show ?thesis by auto
   next case False show ?thesis
   proof (cases "real x = 0")
-    case True thus ?thesis 
+    case True thus ?thesis
       by (auto simp add: less_float_def lb_sqrt.simps ub_sqrt.simps)
   next
     case False with `\<not> 0 < x` have "x < 0" and "0 < -x"
@@ -399,10 +399,10 @@
 fun ub_arctan_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float"
 and lb_arctan_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where
   "ub_arctan_horner prec 0 k x = 0"
-| "ub_arctan_horner prec (Suc n) k x = 
+| "ub_arctan_horner prec (Suc n) k x =
     (rapprox_rat prec 1 (int k)) - x * (lb_arctan_horner prec n (k + 2) x)"
 | "lb_arctan_horner prec 0 k x = 0"
-| "lb_arctan_horner prec (Suc n) k x = 
+| "lb_arctan_horner prec (Suc n) k x =
     (lapprox_rat prec 1 (int k)) - x * (ub_arctan_horner prec n (k + 2) x)"
 
 lemma arctan_0_1_bounds': assumes "0 \<le> real x" "real x \<le> 1" and "even n"
@@ -413,12 +413,12 @@
 
   have "0 \<le> real (x * x)" by auto
   from `even n` obtain m where "2 * m = n" unfolding even_mult_two_ex by auto
-  
+
   have "arctan (real x) \<in> { ?S n .. ?S (Suc n) }"
   proof (cases "real x = 0")
     case False
     hence "0 < real x" using `0 \<le> real x` by auto
-    hence prem: "0 < 1 / real (0 * 2 + (1::nat)) * real x ^ (0 * 2 + 1)" by auto 
+    hence prem: "0 < 1 / real (0 * 2 + (1::nat)) * real x ^ (0 * 2 + 1)" by auto
 
     have "\<bar> real x \<bar> \<le> 1"  using `0 \<le> real x` `real x \<le> 1` by auto
     from mp[OF summable_Leibniz(2)[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] prem, THEN spec, of m, unfolded `2 * m = n`]
@@ -428,9 +428,9 @@
 
   have F: "\<And>n. 2 * Suc n + 1 = 2 * n + 1 + 2" by auto
 
-  note bounds = horner_bounds[where s=1 and f="\<lambda>i. 2 * i + 1" and j'=0 
+  note bounds = horner_bounds[where s=1 and f="\<lambda>i. 2 * i + 1" and j'=0
     and lb="\<lambda>n i k x. lb_arctan_horner prec n k x"
-    and ub="\<lambda>n i k x. ub_arctan_horner prec n k x", 
+    and ub="\<lambda>n i k x. ub_arctan_horner prec n k x",
     OF `0 \<le> real (x*x)` F lb_arctan_horner.simps ub_arctan_horner.simps]
 
   { have "real (x * lb_arctan_horner prec n 1 (x*x)) \<le> ?S n"
@@ -481,15 +481,15 @@
 subsection "Compute \<pi>"
 
 definition ub_pi :: "nat \<Rightarrow> float" where
-  "ub_pi prec = (let A = rapprox_rat prec 1 5 ; 
+  "ub_pi prec = (let A = rapprox_rat prec 1 5 ;
                      B = lapprox_rat prec 1 239
-                 in ((Float 1 2) * ((Float 1 2) * A * (ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (A * A)) - 
+                 in ((Float 1 2) * ((Float 1 2) * A * (ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (A * A)) -
                                                   B * (lb_arctan_horner prec (get_even (prec div 14 + 1)) 1 (B * B)))))"
 
 definition lb_pi :: "nat \<Rightarrow> float" where
-  "lb_pi prec = (let A = lapprox_rat prec 1 5 ; 
+  "lb_pi prec = (let A = lapprox_rat prec 1 5 ;
                      B = rapprox_rat prec 1 239
-                 in ((Float 1 2) * ((Float 1 2) * A * (lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (A * A)) - 
+                 in ((Float 1 2) * ((Float 1 2) * A * (lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (A * A)) -
                                                   B * (ub_arctan_horner prec (get_odd (prec div 14 + 1)) 1 (B * B)))))"
 
 lemma pi_boundaries: "pi \<in> {real (lb_pi n) .. real (ub_pi n)}"
@@ -499,7 +499,7 @@
   { fix prec n :: nat fix k :: int assume "1 < k" hence "0 \<le> k" and "0 < k" and "1 \<le> k" by auto
     let ?k = "rapprox_rat prec 1 k"
     have "1 div k = 0" using div_pos_pos_trivial[OF _ `1 < k`] by auto
-      
+
     have "0 \<le> real ?k" by (rule order_trans[OF _ rapprox_rat], auto simp add: `0 \<le> k`)
     have "real ?k \<le> 1" unfolding rapprox_rat.simps(2)[OF zero_le_one `0 < k`]
       by (rule rapprox_posrat_le1, auto simp add: `0 < k` `1 \<le> k`)
@@ -616,7 +616,7 @@
 	using arctan_0_1_bounds[OF `0 \<le> real ?DIV` `real ?DIV \<le> 1`] by auto
       also have "\<dots> \<le> 2 * arctan (real x / ?R)"
 	using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono)
-      also have "2 * arctan (real x / ?R) = arctan (real x)" using arctan_half[symmetric] unfolding numeral_2_eq_2 power_Suc2 power_0 real_mult_1 . 
+      also have "2 * arctan (real x / ?R) = arctan (real x)" using arctan_half[symmetric] unfolding numeral_2_eq_2 power_Suc2 power_0 real_mult_1 .
       finally show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF True] .
     next
       case False
@@ -629,7 +629,7 @@
       show ?thesis
       proof (cases "1 < ?invx")
 	case True
-	show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF False] if_P[OF True] 
+	show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF False] if_P[OF True]
 	  using `0 \<le> arctan (real x)` by auto
       next
 	case False
@@ -731,7 +731,7 @@
       have "real (?lb_horner ?invx) \<le> arctan (real ?invx)" using arctan_0_1_bounds[OF `0 \<le> real ?invx` `real ?invx \<le> 1`] by auto
       also have "\<dots> \<le> arctan (1 / real x)" unfolding real_of_float_1[symmetric] by (rule arctan_monotone', rule float_divl)
       finally have "arctan (real x) \<le> pi / 2 - real (?lb_horner ?invx)"
-	using `0 \<le> arctan (real x)` arctan_inverse[OF `1 / real x \<noteq> 0`] 
+	using `0 \<le> arctan (real x)` arctan_inverse[OF `1 / real x \<noteq> 0`]
 	unfolding real_sgn_pos[OF `0 < 1 / real x`] le_diff_eq by auto
       moreover
       have "pi / 2 \<le> real (ub_pi prec * Float 1 -1)" unfolding real_of_float_mult Float_num times_divide_eq_right mult_1_right using pi_boundaries by auto
@@ -780,10 +780,10 @@
 fun ub_sin_cos_aux :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float"
 and lb_sin_cos_aux :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where
   "ub_sin_cos_aux prec 0 i k x = 0"
-| "ub_sin_cos_aux prec (Suc n) i k x = 
+| "ub_sin_cos_aux prec (Suc n) i k x =
     (rapprox_rat prec 1 (int k)) - x * (lb_sin_cos_aux prec n (i + 2) (k * i * (i + 1)) x)"
 | "lb_sin_cos_aux prec 0 i k x = 0"
-| "lb_sin_cos_aux prec (Suc n) i k x = 
+| "lb_sin_cos_aux prec (Suc n) i k x =
     (lapprox_rat prec 1 (int k)) - x * (ub_sin_cos_aux prec n (i + 2) (k * i * (i + 1)) x)"
 
 lemma cos_aux:
@@ -793,12 +793,12 @@
   have "0 \<le> real (x * x)" unfolding real_of_float_mult by auto
   let "?f n" = "fact (2 * n)"
 
-  { fix n 
+  { fix n
     have F: "\<And>m. ((\<lambda>i. i + 2) ^^ n) m = m + 2 * n" by (induct n arbitrary: m, auto)
     have "?f (Suc n) = ?f n * ((\<lambda>i. i + 2) ^^ n) 1 * (((\<lambda>i. i + 2) ^^ n) 1 + 1)"
       unfolding F by auto } note f_eq = this
-    
-  from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, 
+
+  from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0,
     OF `0 \<le> real (x * x)` f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
   show "?lb" and "?ub" by (auto simp add: power_mult power2_eq_square[of "real x"])
 qed
@@ -815,7 +815,7 @@
     = (\<Sum> i = 0 ..< 2 * n. (if even(i) then (-1 ^ (i div 2))/(real (fact i)) else 0) * x ^ i)" (is "?sum = ?ifsum")
   proof -
     have "?sum = ?sum + (\<Sum> j = 0 ..< n. 0)" by auto
-    also have "\<dots> = 
+    also have "\<dots> =
       (\<Sum> j = 0 ..< n. -1 ^ ((2 * j) div 2) / (real (fact (2 * j))) * x ^(2 * j)) + (\<Sum> j = 0 ..< n. 0)" by auto
     also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. if even i then -1 ^ (i div 2) / (real (fact i)) * x ^ i else 0)"
       unfolding sum_split_even_odd ..
@@ -828,8 +828,8 @@
   { fix n :: nat assume "0 < n"
     hence "0 < 2 * n" by auto
     obtain t where "0 < t" and "t < real x" and
-      cos_eq: "cos (real x) = (\<Sum> i = 0 ..< 2 * n. (if even(i) then (-1 ^ (i div 2))/(real (fact i)) else 0) * (real x) ^ i) 
-      + (cos (t + 1/2 * real (2 * n) * pi) / real (fact (2*n))) * (real x)^(2*n)" 
+      cos_eq: "cos (real x) = (\<Sum> i = 0 ..< 2 * n. (if even(i) then (-1 ^ (i div 2))/(real (fact i)) else 0) * (real x) ^ i)
+      + (cos (t + 1/2 * real (2 * n) * pi) / real (fact (2*n))) * (real x)^(2*n)"
       (is "_ = ?SUM + ?rest / ?fact * ?pow")
       using Maclaurin_cos_expansion2[OF `0 < real x` `0 < 2 * n`] by auto
 
@@ -848,7 +848,7 @@
     {
       assume "even n"
       have "real (lb_sin_cos_aux prec n 1 1 (x * x)) \<le> ?SUM"
-	unfolding morph_to_if_power[symmetric] using cos_aux by auto 
+	unfolding morph_to_if_power[symmetric] using cos_aux by auto
       also have "\<dots> \<le> cos (real x)"
       proof -
 	from even[OF `even n`] `0 < ?fact` `0 < ?pow`
@@ -874,7 +874,7 @@
   } note ub = this(1) and lb = this(2)
 
   have "cos (real x) \<le> real (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))" using ub[OF odd_pos[OF get_odd] get_odd] .
-  moreover have "real (lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) \<le> cos (real x)" 
+  moreover have "real (lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) \<le> cos (real x)"
   proof (cases "0 < get_even n")
     case True show ?thesis using lb[OF True get_even] .
   next
@@ -889,7 +889,7 @@
   case True
   show ?thesis
   proof (cases "n = 0")
-    case True 
+    case True
     thus ?thesis unfolding `n = 0` get_even_def get_odd_def using `real x = 0` lapprox_rat[where x="-1" and y=1] by auto
   next
     case False with not0_implies_Suc obtain m where "n = Suc m" by blast
@@ -904,11 +904,11 @@
   have "0 \<le> real (x * x)" unfolding real_of_float_mult by auto
   let "?f n" = "fact (2 * n + 1)"
 
-  { fix n 
+  { fix n
     have F: "\<And>m. ((\<lambda>i. i + 2) ^^ n) m = m + 2 * n" by (induct n arbitrary: m, auto)
     have "?f (Suc n) = ?f n * ((\<lambda>i. i + 2) ^^ n) 2 * (((\<lambda>i. i + 2) ^^ n) 2 + 1)"
       unfolding F by auto } note f_eq = this
-    
+
   from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0,
     OF `0 \<le> real (x * x)` f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
   show "?lb" and "?ub" using `0 \<le> real x` unfolding real_of_float_mult
@@ -940,8 +940,8 @@
   { fix n :: nat assume "0 < n"
     hence "0 < 2 * n + 1" by auto
     obtain t where "0 < t" and "t < real x" and
-      sin_eq: "sin (real x) = (\<Sum> i = 0 ..< 2 * n + 1. (if even(i) then 0 else (-1 ^ ((i - Suc 0) div 2))/(real (fact i))) * (real x) ^ i) 
-      + (sin (t + 1/2 * real (2 * n + 1) * pi) / real (fact (2*n + 1))) * (real x)^(2*n + 1)" 
+      sin_eq: "sin (real x) = (\<Sum> i = 0 ..< 2 * n + 1. (if even(i) then 0 else (-1 ^ ((i - Suc 0) div 2))/(real (fact i))) * (real x) ^ i)
+      + (sin (t + 1/2 * real (2 * n + 1) * pi) / real (fact (2*n + 1))) * (real x)^(2*n + 1)"
       (is "_ = ?SUM + ?rest / ?fact * ?pow")
       using Maclaurin_sin_expansion3[OF `0 < 2 * n + 1` `0 < real x`] by auto
 
@@ -956,7 +956,7 @@
 
     {
       assume "even n"
-      have "real (x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le> 
+      have "real (x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le>
             (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else (-1 ^ ((i - Suc 0) div 2))/(real (fact i))) * (real x) ^ i)"
 	using sin_aux[OF `0 \<le> real x`] unfolding setsum_morph[symmetric] by auto
       also have "\<dots> \<le> ?SUM" by auto
@@ -980,14 +980,14 @@
       qed
       also have "\<dots> \<le> (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else (-1 ^ ((i - Suc 0) div 2))/(real (fact i))) * (real x) ^ i)"
 	 by auto
-      also have "\<dots> \<le> real (x * ub_sin_cos_aux prec n 2 1 (x * x))" 
+      also have "\<dots> \<le> real (x * ub_sin_cos_aux prec n 2 1 (x * x))"
 	using sin_aux[OF `0 \<le> real x`] unfolding setsum_morph[symmetric] by auto
       finally have "sin (real x) \<le> real (x * ub_sin_cos_aux prec n 2 1 (x * x))" .
     } note ub = this and lb
   } note ub = this(1) and lb = this(2)
 
   have "sin (real x) \<le> real (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))" using ub[OF odd_pos[OF get_odd] get_odd] .
-  moreover have "real (x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) \<le> sin (real x)" 
+  moreover have "real (x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) \<le> sin (real x)"
   proof (cases "0 < get_even n")
     case True show ?thesis using lb[OF True get_even] .
   next
@@ -1001,7 +1001,7 @@
   case True
   show ?thesis
   proof (cases "n = 0")
-    case True 
+    case True
     thus ?thesis unfolding `n = 0` get_even_def get_odd_def using `real x = 0` lapprox_rat[where x="-1" and y=1] by auto
   next
     case False with not0_implies_Suc obtain m where "n = Suc m" by blast
@@ -1106,7 +1106,7 @@
       moreover have "?cos x \<le> real (?ub x)"
       proof -
 	from ub_half[OF ub `-pi \<le> real x` `real x \<le> pi`]
-	show ?thesis unfolding ub_cos_def[where x=x] Let_def using `\<not> x < 0` `\<not> x < Float 1 -1` `x < 1` by auto 
+	show ?thesis unfolding ub_cos_def[where x=x] Let_def using `\<not> x < 0` `\<not> x < Float 1 -1` `x < 1` by auto
       qed
       ultimately show ?thesis by auto
     next
@@ -1435,7 +1435,7 @@
     qed
     finally have "real (lb_exp_horner prec (get_even n) 1 1 x) \<le> exp (real x)" .
   } moreover
-  { 
+  {
     have x_less_zero: "real x ^ get_odd n \<le> 0"
     proof (cases "real x = 0")
       case True
@@ -1462,12 +1462,12 @@
 
 function ub_exp :: "nat \<Rightarrow> float \<Rightarrow> float" and lb_exp :: "nat \<Rightarrow> float \<Rightarrow> float" where
 "lb_exp prec x = (if 0 < x then float_divl prec 1 (ub_exp prec (-x))
-             else let 
+             else let
                 horner = (\<lambda> x. let  y = lb_exp_horner prec (get_even (prec + 2)) 1 1 x  in if y \<le> 0 then Float 1 -2 else y)
              in if x < - 1 then (case floor_fl x of (Float m e) \<Rightarrow> (horner (float_divl prec x (- Float m e))) ^ (nat (-m) * 2 ^ nat e))
                            else horner x)" |
 "ub_exp prec x = (if 0 < x    then float_divr prec 1 (lb_exp prec (-x))
-             else if x < - 1  then (case floor_fl x of (Float m e) \<Rightarrow> 
+             else if x < - 1  then (case floor_fl x of (Float m e) \<Rightarrow>
                                     (ub_exp_horner prec (get_odd (prec + 2)) 1 1 (float_divr prec x (- Float m e))) ^ (nat (-m) * 2 ^ nat e))
                               else ub_exp_horner prec (get_odd (prec + 2)) 1 1 x)"
 by pat_completeness auto
@@ -1479,10 +1479,10 @@
 
   have "1 / 4 = real (Float 1 -2)" unfolding Float_num by auto
   also have "\<dots> \<le> real (lb_exp_horner 1 (get_even 4) 1 1 (- 1))"
-    unfolding get_even_def eq4 
+    unfolding get_even_def eq4
     by (auto simp add: lapprox_posrat_def rapprox_posrat_def normfloat.simps)
   also have "\<dots> \<le> exp (real (- 1 :: float))" using bnds_exp_horner[where x="- 1"] by auto
-  finally show ?thesis unfolding real_of_float_minus real_of_float_1 . 
+  finally show ?thesis unfolding real_of_float_minus real_of_float_1 .
 qed
 
 lemma lb_exp_pos: assumes "\<not> 0 < x" shows "0 < lb_exp prec x"
@@ -1523,10 +1523,10 @@
     qed
   next
     case True
-    
+
     obtain m e where Float_floor: "floor_fl x = Float m e" by (cases "floor_fl x", auto)
     let ?num = "nat (- m) * 2 ^ nat e"
-    
+
     have "real (floor_fl x) < - 1" using floor_fl `x < - 1` unfolding le_float_def less_float_def real_of_float_minus real_of_float_1 by (rule order_le_less_trans)
     hence "real (floor_fl x) < 0" unfolding Float_floor real_of_float_simp using zero_less_pow2[of xe] by auto
     hence "m < 0"
@@ -1544,12 +1544,12 @@
       unfolding Float_floor real_of_float_minus real_of_float_simp real_of_nat_mult pow2_int[of "nat e", unfolded e_nat] realpow_real_of_nat[symmetric] by auto
     have "0 < - floor_fl x" using `0 < ?num`[unfolded real_of_nat_less_iff[symmetric]] unfolding less_float_def num_eq[symmetric] real_of_float_0 real_of_nat_zero .
     hence "real (floor_fl x) < 0" unfolding less_float_def by auto
-    
+
     have "exp (real x) \<le> real (ub_exp prec x)"
     proof -
-      have div_less_zero: "real (float_divr prec x (- floor_fl x)) \<le> 0" 
+      have div_less_zero: "real (float_divr prec x (- floor_fl x)) \<le> 0"
 	using float_divr_nonpos_pos_upper_bound[OF `x \<le> 0` `0 < - floor_fl x`] unfolding le_float_def real_of_float_0 .
-      
+
       have "exp (real x) = exp (real ?num * (real x / real ?num))" using `real ?num \<noteq> 0` by auto
       also have "\<dots> = exp (real x / real ?num) ^ ?num" unfolding exp_real_of_nat_mult ..
       also have "\<dots> \<le> exp (real (float_divr prec x (- floor_fl x))) ^ ?num" unfolding num_eq
@@ -1558,21 +1558,21 @@
 	by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto)
       finally show ?thesis unfolding ub_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def .
     qed
-    moreover 
+    moreover
     have "real (lb_exp prec x) \<le> exp (real x)"
     proof -
       let ?divl = "float_divl prec x (- Float m e)"
       let ?horner = "?lb_exp_horner ?divl"
-      
+
       show ?thesis
       proof (cases "?horner \<le> 0")
 	case False hence "0 \<le> real ?horner" unfolding le_float_def by auto
-	
+
 	have div_less_zero: "real (float_divl prec x (- floor_fl x)) \<le> 0"
 	  using `real (floor_fl x) < 0` `real x \<le> 0` by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg)
-	
-	have "real ((?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num) \<le>  
-          exp (real (float_divl prec x (- floor_fl x))) ^ ?num" unfolding float_power 
+
+	have "real ((?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num) \<le>
+          exp (real (float_divl prec x (- floor_fl x))) ^ ?num" unfolding float_power
 	  using `0 \<le> real ?horner`[unfolded Float_floor[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono)
 	also have "\<dots> \<le> exp (real x / real ?num) ^ ?num" unfolding num_eq
 	  using float_divl by (auto intro!: power_mono simp del: real_of_float_minus)
@@ -1602,16 +1602,16 @@
 proof -
   show ?thesis
   proof (cases "0 < x")
-    case False hence "x \<le> 0" unfolding less_float_def le_float_def by auto 
+    case False hence "x \<le> 0" unfolding less_float_def le_float_def by auto
     from exp_boundaries'[OF this] show ?thesis .
   next
     case True hence "-x \<le> 0" unfolding less_float_def le_float_def by auto
-    
+
     have "real (lb_exp prec x) \<le> exp (real x)"
     proof -
       from exp_boundaries'[OF `-x \<le> 0`]
       have ub_exp: "exp (- real x) \<le> real (ub_exp prec (-x))" unfolding atLeastAtMost_iff real_of_float_minus by auto
-      
+
       have "real (float_divl prec 1 (ub_exp prec (-x))) \<le> 1 / real (ub_exp prec (-x))" using float_divl[where x=1] by auto
       also have "\<dots> \<le> exp (real x)"
 	using ub_exp[unfolded inverse_le_iff_le[OF order_less_le_trans[OF exp_gt_zero ub_exp] exp_gt_zero, symmetric]]
@@ -1622,12 +1622,12 @@
     have "exp (real x) \<le> real (ub_exp prec x)"
     proof -
       have "\<not> 0 < -x" using `0 < x` unfolding less_float_def by auto
-      
+
       from exp_boundaries'[OF `-x \<le> 0`]
       have lb_exp: "real (lb_exp prec (-x)) \<le> exp (- real x)" unfolding atLeastAtMost_iff real_of_float_minus by auto
-      
+
       have "exp (real x) \<le> real (1 :: float) / real (lb_exp prec (-x))"
-	using lb_exp[unfolded inverse_le_iff_le[OF exp_gt_zero lb_exp_pos[OF `\<not> 0 < -x`, unfolded less_float_def real_of_float_0], 
+	using lb_exp[unfolded inverse_le_iff_le[OF exp_gt_zero lb_exp_pos[OF `\<not> 0 < -x`, unfolded less_float_def real_of_float_0],
 	                                        symmetric]]
 	unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide real_of_float_1 by auto
       also have "\<dots> \<le> real (float_divr prec 1 (lb_exp prec (-x)))" using float_divr .
@@ -1658,7 +1658,7 @@
 
 subsection "Compute the logarithm series"
 
-fun ub_ln_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" 
+fun ub_ln_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float"
 and lb_ln_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where
 "ub_ln_horner prec 0 i x       = 0" |
 "ub_ln_horner prec (Suc n) i x = rapprox_rat prec 1 (int i) - x * lb_ln_horner prec n (Suc i) x" |
@@ -1676,13 +1676,13 @@
     using ln_series[of "x + 1"] `0 \<le> x` `x < 1` by auto
 
   have "norm x < 1" using assms by auto
-  have "?a ----> 0" unfolding Suc_eq_plus1[symmetric] inverse_eq_divide[symmetric] 
+  have "?a ----> 0" unfolding Suc_eq_plus1[symmetric] inverse_eq_divide[symmetric]
     using LIMSEQ_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_Suc[OF LIMSEQ_power_zero[OF `norm x < 1`]]] by auto
   { fix n have "0 \<le> ?a n" by (rule mult_nonneg_nonneg, auto intro!: mult_nonneg_nonneg simp add: `0 \<le> x`) }
   { fix n have "?a (Suc n) \<le> ?a n" unfolding inverse_eq_divide[symmetric]
     proof (rule mult_mono)
       show "0 \<le> x ^ Suc (Suc n)" by (auto intro!: mult_nonneg_nonneg simp add: `0 \<le> x`)
-      have "x ^ Suc (Suc n) \<le> x ^ Suc n * 1" unfolding power_Suc2 real_mult_assoc[symmetric] 
+      have "x ^ Suc (Suc n) \<le> x ^ Suc n * 1" unfolding power_Suc2 real_mult_assoc[symmetric]
 	by (rule mult_left_mono, fact less_imp_le[OF `x < 1`], auto intro!: mult_nonneg_nonneg simp add: `0 \<le> x`)
       thus "x ^ Suc (Suc n) \<le> x ^ Suc n" by auto
     qed auto }
@@ -1690,7 +1690,7 @@
   show "?lb" and "?ub" by auto
 qed
 
-lemma ln_float_bounds: 
+lemma ln_float_bounds:
   assumes "0 \<le> real x" and "real x < 1"
   shows "real (x * lb_ln_horner prec (get_even n) 1 x) \<le> ln (real x + 1)" (is "?lb \<le> ?ln")
   and "ln (real x + 1) \<le> real (x * ub_ln_horner prec (get_odd n) 1 x)" (is "?ln \<le> ?ub")
@@ -1705,21 +1705,21 @@
       OF `0 \<le> real x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \<le> real x`
     by (rule mult_right_mono)
   also have "\<dots> \<le> ?ln" using ln_bounds(1)[OF `0 \<le> real x` `real x < 1`] by auto
-  finally show "?lb \<le> ?ln" . 
+  finally show "?lb \<le> ?ln" .
 
   have "?ln \<le> setsum ?s {0 ..< 2 * od + 1}" using ln_bounds(2)[OF `0 \<le> real x` `real x < 1`] by auto
   also have "\<dots> \<le> ?ub" unfolding power_Suc2 real_mult_assoc[symmetric] real_of_float_mult setsum_left_distrib[symmetric] unfolding real_mult_commute[of "real x"] od
     using horner_bounds(2)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1",
       OF `0 \<le> real x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \<le> real x`
     by (rule mult_right_mono)
-  finally show "?ln \<le> ?ub" . 
+  finally show "?ln \<le> ?ub" .
 qed
 
 lemma ln_add: assumes "0 < x" and "0 < y" shows "ln (x + y) = ln x + ln (1 + y / x)"
 proof -
   have "x \<noteq> 0" using assms by auto
   have "x + y = x * (1 + y / x)" unfolding right_distrib times_divide_eq_right nonzero_mult_divide_cancel_left[OF `x \<noteq> 0`] by auto
-  moreover 
+  moreover
   have "0 < y / x" using assms divide_pos_pos by auto
   hence "0 < 1 + y / x" by auto
   ultimately show ?thesis using ln_mult assms by auto
@@ -1727,11 +1727,11 @@
 
 subsection "Compute the logarithm of 2"
 
-definition ub_ln2 where "ub_ln2 prec = (let third = rapprox_rat (max prec 1) 1 3 
-                                        in (Float 1 -1 * ub_ln_horner prec (get_odd prec) 1 (Float 1 -1)) + 
+definition ub_ln2 where "ub_ln2 prec = (let third = rapprox_rat (max prec 1) 1 3
+                                        in (Float 1 -1 * ub_ln_horner prec (get_odd prec) 1 (Float 1 -1)) +
                                            (third * ub_ln_horner prec (get_odd prec) 1 third))"
-definition lb_ln2 where "lb_ln2 prec = (let third = lapprox_rat prec 1 3 
-                                        in (Float 1 -1 * lb_ln_horner prec (get_even prec) 1 (Float 1 -1)) + 
+definition lb_ln2 where "lb_ln2 prec = (let third = lapprox_rat prec 1 3
+                                        in (Float 1 -1 * lb_ln_horner prec (get_even prec) 1 (Float 1 -1)) +
                                            (third * lb_ln_horner prec (get_even prec) 1 third))"
 
 lemma ub_ln2: "ln 2 \<le> real (ub_ln2 prec)" (is "?ub_ln2")
@@ -2069,8 +2069,7 @@
   | Atom nat
   | Num float
 
-fun interpret_floatarith :: "floatarith \<Rightarrow> real list \<Rightarrow> real"
-where
+fun interpret_floatarith :: "floatarith \<Rightarrow> real list \<Rightarrow> real" where
 "interpret_floatarith (Add a b) vs   = (interpret_floatarith a vs) + (interpret_floatarith b vs)" |
 "interpret_floatarith (Minus a) vs    = - (interpret_floatarith a vs)" |
 "interpret_floatarith (Mult a b) vs   = (interpret_floatarith a vs) * (interpret_floatarith b vs)" |
@@ -2088,6 +2087,35 @@
 "interpret_floatarith (Num f) vs      = real f" |
 "interpret_floatarith (Atom n) vs     = vs ! n"
 
+lemma interpret_floatarith_divide: "interpret_floatarith (Mult a (Inverse b)) vs = (interpret_floatarith a vs) / (interpret_floatarith b vs)"
+  unfolding real_divide_def interpret_floatarith.simps ..
+
+lemma interpret_floatarith_diff: "interpret_floatarith (Add a (Minus b)) vs = (interpret_floatarith a vs) - (interpret_floatarith b vs)"
+  unfolding real_diff_def interpret_floatarith.simps ..
+
+lemma interpret_floatarith_sin: "interpret_floatarith (Cos (Add (Mult Pi (Num (Float 1 -1))) (Minus a))) vs =
+  sin (interpret_floatarith a vs)"
+  unfolding sin_cos_eq interpret_floatarith.simps
+            interpret_floatarith_divide interpret_floatarith_diff real_diff_def
+  by auto
+
+lemma interpret_floatarith_tan:
+  "interpret_floatarith (Mult (Cos (Add (Mult Pi (Num (Float 1 -1))) (Minus a))) (Inverse (Cos a))) vs =
+   tan (interpret_floatarith a vs)"
+  unfolding interpret_floatarith.simps(3,4) interpret_floatarith_sin tan_def real_divide_def
+  by auto
+
+lemma interpret_floatarith_powr: "interpret_floatarith (Exp (Mult b (Ln a))) vs = (interpret_floatarith a vs) powr (interpret_floatarith b vs)"
+  unfolding powr_def interpret_floatarith.simps ..
+
+lemma interpret_floatarith_log: "interpret_floatarith ((Mult (Ln x) (Inverse (Ln b)))) vs = log (interpret_floatarith b vs) (interpret_floatarith x vs)"
+  unfolding log_def interpret_floatarith.simps real_divide_def ..
+
+lemma interpret_floatarith_num:
+  shows "interpret_floatarith (Num (Float 0 0)) vs = 0"
+  and "interpret_floatarith (Num (Float 1 0)) vs = 1"
+  and "interpret_floatarith (Num (Float (number_of a) 0)) vs = number_of a" by auto
+
 subsection "Implement approximation function"
 
 fun lift_bin' :: "(float * float) option \<Rightarrow> (float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> (float * float)) \<Rightarrow> (float * float) option" where
@@ -2103,35 +2131,50 @@
 "lift_un' (Some (l1, u1)) f = Some (f l1 u1)" |
 "lift_un' b f = None"
 
-fun bounded_by :: "real list \<Rightarrow> (float * float) list \<Rightarrow> bool " where
-bounded_by_Cons: "bounded_by (v#vs) ((l, u)#bs) = ((real l \<le> v \<and> v \<le> real u) \<and> bounded_by vs bs)" |
-bounded_by_Nil: "bounded_by [] [] = True" |
-"bounded_by _ _ = False"
-
-lemma bounded_by: assumes "bounded_by vs bs" and "i < length bs"
-  shows "real (fst (bs ! i)) \<le> vs ! i \<and> vs ! i \<le> real (snd (bs ! i))"
-  using `bounded_by vs bs` and `i < length bs`
-proof (induct arbitrary: i rule: bounded_by.induct)
-  fix v :: real and vs :: "real list" and l u :: float and bs :: "(float * float) list" and i :: nat
-  assume hyp: "\<And>i. \<lbrakk>bounded_by vs bs; i < length bs\<rbrakk> \<Longrightarrow> real (fst (bs ! i)) \<le> vs ! i \<and> vs ! i \<le> real (snd (bs ! i))"
-  assume bounded: "bounded_by (v # vs) ((l, u) # bs)" and length: "i < length ((l, u) # bs)"
-  show "real (fst (((l, u) # bs) ! i)) \<le> (v # vs) ! i \<and> (v # vs) ! i \<le> real (snd (((l, u) # bs) ! i))"
-  proof (cases i)
-    case 0
-    show ?thesis using bounded unfolding 0 nth_Cons_0 fst_conv snd_conv bounded_by.simps ..
-  next
-    case (Suc i) with length have "i < length bs" by auto
-    show ?thesis unfolding Suc nth_Cons_Suc bounded_by.simps
-      using hyp[OF bounded[unfolded bounded_by.simps, THEN conjunct2] `i < length bs`] .
-  qed
-qed auto
-
-fun approx approx' :: "nat \<Rightarrow> floatarith \<Rightarrow> (float * float) list \<Rightarrow> (float * float) option" where
+definition
+"bounded_by xs vs \<longleftrightarrow>
+  (\<forall> i < length vs. case vs ! i of None \<Rightarrow> True
+         | Some (l, u) \<Rightarrow> xs ! i \<in> { real l .. real u })"
+
+lemma bounded_byE:
+  assumes "bounded_by xs vs"
+  shows "\<And> i. i < length vs \<Longrightarrow> case vs ! i of None \<Rightarrow> True
+         | Some (l, u) \<Rightarrow> xs ! i \<in> { real l .. real u }"
+  using assms bounded_by_def by blast
+
+lemma bounded_by_update:
+  assumes "bounded_by xs vs"
+  and bnd: "xs ! i \<in> { real l .. real u }"
+  shows "bounded_by xs (vs[i := Some (l,u)])"
+proof -
+{ fix j
+  let ?vs = "vs[i := Some (l,u)]"
+  assume "j < length ?vs" hence [simp]: "j < length vs" by simp
+  have "case ?vs ! j of None \<Rightarrow> True | Some (l, u) \<Rightarrow> xs ! j \<in> { real l .. real u }"
+  proof (cases "?vs ! j")
+    case (Some b)
+    thus ?thesis
+    proof (cases "i = j")
+      case True
+      thus ?thesis using `?vs ! j = Some b` and bnd by auto
+    next
+      case False
+      thus ?thesis using `bounded_by xs vs` unfolding bounded_by_def by auto
+    qed
+  qed auto }
+  thus ?thesis unfolding bounded_by_def by auto
+qed
+
+lemma bounded_by_None:
+  shows "bounded_by xs (replicate (length xs) None)"
+  unfolding bounded_by_def by auto
+
+fun approx approx' :: "nat \<Rightarrow> floatarith \<Rightarrow> (float * float) option list \<Rightarrow> (float * float) option" where
 "approx' prec a bs          = (case (approx prec a bs) of Some (l, u) \<Rightarrow> Some (round_down prec l, round_up prec u) | None \<Rightarrow> None)" |
-"approx prec (Add a b) bs  = lift_bin' (approx' prec a bs) (approx' prec b bs) (\<lambda> l1 u1 l2 u2. (l1 + l2, u1 + u2))" | 
+"approx prec (Add a b) bs   = lift_bin' (approx' prec a bs) (approx' prec b bs) (\<lambda> l1 u1 l2 u2. (l1 + l2, u1 + u2))" |
 "approx prec (Minus a) bs   = lift_un' (approx' prec a bs) (\<lambda> l u. (-u, -l))" |
 "approx prec (Mult a b) bs  = lift_bin' (approx' prec a bs) (approx' prec b bs)
-                                    (\<lambda> a1 a2 b1 b2. (float_nprt a1 * float_pprt b2 + float_nprt a2 * float_nprt b2 + float_pprt a1 * float_pprt b1 + float_pprt a2 * float_nprt b1, 
+                                    (\<lambda> a1 a2 b1 b2. (float_nprt a1 * float_pprt b2 + float_nprt a2 * float_nprt b2 + float_pprt a1 * float_pprt b1 + float_pprt a2 * float_nprt b1,
                                                      float_pprt a2 * float_pprt b2 + float_pprt a1 * float_nprt b2 + float_nprt a2 * float_pprt b1 + float_nprt a1 * float_nprt b1))" |
 "approx prec (Inverse a) bs = lift_un (approx' prec a bs) (\<lambda> l u. if (0 < l \<or> u < 0) then (Some (float_divl prec 1 u), Some (float_divr prec 1 l)) else (None, None))" |
 "approx prec (Cos a) bs     = lift_un' (approx' prec a bs) (bnds_cos prec)" |
@@ -2145,7 +2188,7 @@
 "approx prec (Ln a) bs      = lift_un (approx' prec a bs) (\<lambda> l u. (lb_ln prec l, ub_ln prec u))" |
 "approx prec (Power a n) bs = lift_un' (approx' prec a bs) (float_power_bnds n)" |
 "approx prec (Num f) bs     = Some (f, f)" |
-"approx prec (Atom i) bs    = (if i < length bs then Some (bs ! i) else None)"
+"approx prec (Atom i) bs    = (if i < length bs then bs ! i else None)"
 
 lemma lift_bin'_ex:
   assumes lift_bin'_Some: "Some (l, u) = lift_bin' a b f"
@@ -2174,9 +2217,9 @@
 proof -
   obtain l1 u1 l2 u2
     where Sa: "Some (l1, u1) = g a" and Sb: "Some (l2, u2) = g b" using lift_bin'_ex[OF assms(1)] by auto
-  have lu: "(l, u) = f l1 u1 l2 u2" using lift_bin'_Some[unfolded Sa[symmetric] Sb[symmetric] lift_bin'.simps] by auto 
+  have lu: "(l, u) = f l1 u1 l2 u2" using lift_bin'_Some[unfolded Sa[symmetric] Sb[symmetric] lift_bin'.simps] by auto
   have "l = fst (f l1 u1 l2 u2)" and "u = snd (f l1 u1 l2 u2)" unfolding lu[symmetric] by auto
-  thus ?thesis using Pa[OF Sa] Pb[OF Sb] by auto 
+  thus ?thesis using Pa[OF Sa] Pb[OF Sb] by auto
 qed
 
 lemma approx_approx':
@@ -2188,7 +2231,7 @@
     using approx' unfolding approx'.simps by (cases "approx prec a vs", auto)
   have l': "l = round_down prec l'" and u': "u = round_up prec u'"
     using approx' unfolding approx'.simps S[symmetric] by auto
-  show ?thesis unfolding l' u' 
+  show ?thesis unfolding l' u'
     using order_trans[OF Pa[OF S, THEN conjunct2] round_up[of u']]
     using order_trans[OF round_down[of _ l'] Pa[OF S, THEN conjunct1]] by auto
 qed
@@ -2197,8 +2240,8 @@
   assumes lift_bin'_Some: "Some (l, u) = lift_bin' (approx' prec a bs) (approx' prec b bs) f"
   and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow> real l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u" (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")
   and Pb: "\<And>l u. Some (l, u) = approx prec b bs \<Longrightarrow> real l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> real u"
-  shows "\<exists> l1 u1 l2 u2. (real l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u1) \<and> 
-                        (real l2 \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> real u2) \<and> 
+  shows "\<exists> l1 u1 l2 u2. (real l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u1) \<and>
+                        (real l2 \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> real u2) \<and>
                         l = fst (f l1 u1 l2 u2) \<and> u = snd (f l1 u1 l2 u2)"
 proof -
   { fix l u assume "Some (l, u) = approx' prec a bs"
@@ -2238,7 +2281,7 @@
 lemma lift_un':
   assumes lift_un'_Some: "Some (l, u) = lift_un' (approx' prec a bs) f"
   and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow> real l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u" (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")
-  shows "\<exists> l1 u1. (real l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u1) \<and> 
+  shows "\<exists> l1 u1. (real l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u1) \<and>
                         l = fst (f l1 u1) \<and> u = snd (f l1 u1)"
 proof -
   { fix l u assume "Some (l, u) = approx' prec a bs"
@@ -2282,7 +2325,7 @@
   proof (rule ccontr)
     assume "\<not> (fst (f l1 u1) \<noteq> None \<and> snd (f l1 u1) \<noteq> None)"
     hence or: "fst (f l1 u1) = None \<or> snd (f l1 u1) = None" by auto
-    hence "lift_un (g a) f = None" 
+    hence "lift_un (g a) f = None"
     proof (cases "fst (f l1 u1) = None")
       case True
       then obtain b where b: "f l1 u1 = (None, b)" by (cases "f l1 u1", auto)
@@ -2303,7 +2346,7 @@
 lemma lift_un:
   assumes lift_un_Some: "Some (l, u) = lift_un (approx' prec a bs) f"
   and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow> real l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u" (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")
-  shows "\<exists> l1 u1. (real l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u1) \<and> 
+  shows "\<exists> l1 u1. (real l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real u1) \<and>
                   Some l = fst (f l1 u1) \<and> Some u = snd (f l1 u1)"
 proof -
   { fix l u assume "Some (l, u) = approx' prec a bs"
@@ -2329,7 +2372,7 @@
   assumes "bounded_by xs vs"
   and "Some (l, u) = approx prec arith vs" (is "_ = ?g arith")
   shows "real l \<le> interpret_floatarith arith xs \<and> interpret_floatarith arith xs \<le> real u" (is "?P l u arith")
-  using `Some (l, u) = approx prec arith vs` 
+  using `Some (l, u) = approx prec arith vs`
 proof (induct arith arbitrary: l u x)
   case (Add a b)
   from lift_bin'[OF Add.prems[unfolded approx.simps]] Add.hyps
@@ -2346,17 +2389,17 @@
 next
   case (Mult a b)
   from lift_bin'[OF Mult.prems[unfolded approx.simps]] Mult.hyps
-  obtain l1 u1 l2 u2 
+  obtain l1 u1 l2 u2
     where l: "l = float_nprt l1 * float_pprt u2 + float_nprt u1 * float_nprt u2 + float_pprt l1 * float_pprt l2 + float_pprt u1 * float_nprt l2"
     and u: "u = float_pprt u1 * float_pprt u2 + float_pprt l1 * float_nprt u2 + float_nprt u1 * float_pprt l2 + float_nprt l1 * float_nprt l2"
     and "real l1 \<le> interpret_floatarith a xs" and "interpret_floatarith a xs \<le> real u1"
     and "real l2 \<le> interpret_floatarith b xs" and "interpret_floatarith b xs \<le> real u2" unfolding fst_conv snd_conv by blast
-  thus ?case unfolding interpret_floatarith.simps l u real_of_float_add real_of_float_mult real_of_float_nprt real_of_float_pprt 
+  thus ?case unfolding interpret_floatarith.simps l u real_of_float_add real_of_float_mult real_of_float_nprt real_of_float_pprt
     using mult_le_prts mult_ge_prts by auto
 next
   case (Inverse a)
   from lift_un[OF Inverse.prems[unfolded approx.simps], unfolded if_distrib[of fst] if_distrib[of snd] fst_conv snd_conv] Inverse.hyps
-  obtain l1 u1 where l': "Some l = (if 0 < l1 \<or> u1 < 0 then Some (float_divl prec 1 u1) else None)" 
+  obtain l1 u1 where l': "Some l = (if 0 < l1 \<or> u1 < 0 then Some (float_divl prec 1 u1) else None)"
     and u': "Some u = (if 0 < l1 \<or> u1 < 0 then Some (float_divr prec 1 l1) else None)"
     and l1: "real l1 \<le> interpret_floatarith a xs" and u1: "interpret_floatarith a xs \<le> real u1" by blast
   have either: "0 < l1 \<or> u1 < 0" proof (rule ccontr) assume P: "\<not> (0 < l1 \<or> u1 < 0)" show False using l' unfolding if_not_P[OF P] by auto qed
@@ -2366,7 +2409,7 @@
   have inv: "inverse (real u1) \<le> inverse (interpret_floatarith a xs)
            \<and> inverse (interpret_floatarith a xs) \<le> inverse (real l1)"
   proof (cases "0 < l1")
-    case True hence "0 < real u1" and "0 < real l1" "0 < interpret_floatarith a xs" 
+    case True hence "0 < real u1" and "0 < real l1" "0 < interpret_floatarith a xs"
       unfolding less_float_def using l1_le_u1 l1 by auto
     show ?thesis
       unfolding inverse_le_iff_le[OF `0 < real u1` `0 < interpret_floatarith a xs`]
@@ -2374,7 +2417,7 @@
       using l1 u1 by auto
   next
     case False hence "u1 < 0" using either by blast
-    hence "real u1 < 0" and "real l1 < 0" "interpret_floatarith a xs < 0" 
+    hence "real u1 < 0" and "real l1 < 0" "interpret_floatarith a xs < 0"
       unfolding less_float_def using l1_le_u1 u1 by auto
     show ?thesis
       unfolding inverse_le_iff_le_neg[OF `real u1 < 0` `interpret_floatarith a xs < 0`]
@@ -2420,119 +2463,730 @@
 next case (Power a n) with lift_un'_bnds[OF bnds_power] show ?case by auto
 next case (Num f) thus ?case by auto
 next
-  case (Atom n) 
-  show ?case
-  proof (cases "n < length vs")
-    case True
-    with Atom have "vs ! n = (l, u)" by auto
-    thus ?thesis using bounded_by[OF assms(1) True] by auto
+  case (Atom n)
+  from this[symmetric] `bounded_by xs vs`[THEN bounded_byE, of n]
+  show ?case by (cases "n < length vs", auto)
+qed
+
+datatype form = Bound floatarith floatarith floatarith form
+              | Assign floatarith floatarith form
+              | Less floatarith floatarith
+              | LessEqual floatarith floatarith
+              | AtLeastAtMost floatarith floatarith floatarith
+
+fun interpret_form :: "form \<Rightarrow> real list \<Rightarrow> bool" where
+"interpret_form (Bound x a b f) vs = (interpret_floatarith x vs \<in> { interpret_floatarith a vs .. interpret_floatarith b vs } \<longrightarrow> interpret_form f vs)" |
+"interpret_form (Assign x a f) vs  = (interpret_floatarith x vs = interpret_floatarith a vs \<longrightarrow> interpret_form f vs)" |
+"interpret_form (Less a b) vs      = (interpret_floatarith a vs < interpret_floatarith b vs)" |
+"interpret_form (LessEqual a b) vs = (interpret_floatarith a vs \<le> interpret_floatarith b vs)" |
+"interpret_form (AtLeastAtMost x a b) vs = (interpret_floatarith x vs \<in> { interpret_floatarith a vs .. interpret_floatarith b vs })"
+
+fun approx_form' and approx_form :: "nat \<Rightarrow> form \<Rightarrow> (float * float) option list \<Rightarrow> nat list \<Rightarrow> bool" where
+"approx_form' prec f 0 n l u bs ss = approx_form prec f (bs[n := Some (l, u)]) ss" |
+"approx_form' prec f (Suc s) n l u bs ss =
+  (let m = (l + u) * Float 1 -1
+   in approx_form' prec f s n l m bs ss \<and>
+      approx_form' prec f s n m u bs ss)" |
+"approx_form prec (Bound (Atom n) a b f) bs ss =
+   (case (approx prec a bs, approx prec b bs)
+   of (Some (l, _), Some (_, u)) \<Rightarrow> approx_form' prec f (ss ! n) n l u bs ss
+    | _ \<Rightarrow> False)" |
+"approx_form prec (Assign (Atom n) a f) bs ss =
+   (case (approx prec a bs)
+   of (Some (l, u)) \<Rightarrow> approx_form' prec f (ss ! n) n l u bs ss
+    | _ \<Rightarrow> False)" |
+"approx_form prec (Less a b) bs ss =
+   (case (approx prec a bs, approx prec b bs)
+   of (Some (l, u), Some (l', u')) \<Rightarrow> u < l'
+    | _ \<Rightarrow> False)" |
+"approx_form prec (LessEqual a b) bs ss =
+   (case (approx prec a bs, approx prec b bs)
+   of (Some (l, u), Some (l', u')) \<Rightarrow> u \<le> l'
+    | _ \<Rightarrow> False)" |
+"approx_form prec (AtLeastAtMost x a b) bs ss =
+   (case (approx prec x bs, approx prec a bs, approx prec b bs)
+   of (Some (lx, ux), Some (l, u), Some (l', u')) \<Rightarrow> u \<le> lx \<and> ux \<le> l'
+    | _ \<Rightarrow> False)" |
+"approx_form _ _ _ _ = False"
+
+lemma approx_form_approx_form':
+  assumes "approx_form' prec f s n l u bs ss" and "x \<in> { real l .. real u }"
+  obtains l' u' where "x \<in> { real l' .. real u' }"
+  and "approx_form prec f (bs[n := Some (l', u')]) ss"
+using assms proof (induct s arbitrary: l u)
+  case 0
+  from this(1)[of l u] this(2,3)
+  show thesis by auto
+next
+  case (Suc s)
+
+  let ?m = "(l + u) * Float 1 -1"
+  have "real l \<le> real ?m" and "real ?m \<le> real u"
+    unfolding le_float_def using Suc.prems by auto
+
+  with `x \<in> { real l .. real u }`
+  have "x \<in> { real l .. real ?m} \<or> x \<in> { real ?m .. real u }" by auto
+  thus thesis
+  proof (rule disjE)
+    assume *: "x \<in> { real l .. real ?m }"
+    with Suc.hyps[OF _ _ *] Suc.prems
+    show thesis by (simp add: Let_def)
   next
-    case False thus ?thesis using Atom by auto
+    assume *: "x \<in> { real ?m .. real u }"
+    with Suc.hyps[OF _ _ *] Suc.prems
+    show thesis by (simp add: Let_def)
   qed
 qed
 
-datatype inequality = Less floatarith floatarith 
-                    | LessEqual floatarith floatarith 
-
-fun interpret_inequality :: "inequality \<Rightarrow> real list \<Rightarrow> bool" where 
-"interpret_inequality (Less a b) vs                   = (interpret_floatarith a vs < interpret_floatarith b vs)" |
-"interpret_inequality (LessEqual a b) vs              = (interpret_floatarith a vs \<le> interpret_floatarith b vs)"
-
-fun approx_inequality :: "nat \<Rightarrow> inequality \<Rightarrow> (float * float) list \<Rightarrow> bool" where 
-"approx_inequality prec (Less a b) bs = (case (approx prec a bs, approx prec b bs) of (Some (l, u), Some (l', u')) \<Rightarrow> u < l' | _ \<Rightarrow> False)" |
-"approx_inequality prec (LessEqual a b) bs = (case (approx prec a bs, approx prec b bs) of (Some (l, u), Some (l', u')) \<Rightarrow> u \<le> l' | _ \<Rightarrow> False)"
-
-lemma approx_inequality: fixes m :: nat assumes "bounded_by vs bs" and "approx_inequality prec eq bs"
-  shows "interpret_inequality eq vs"
-proof (cases eq)
+lemma approx_form_aux:
+  assumes "approx_form prec f vs ss"
+  and "bounded_by xs vs"
+  shows "interpret_form f xs"
+using assms proof (induct f arbitrary: vs)
+  case (Bound x a b f)
+  then obtain n
+    where x_eq: "x = Atom n" by (cases x) auto
+
+  with Bound.prems obtain l u' l' u
+    where l_eq: "Some (l, u') = approx prec a vs"
+    and u_eq: "Some (l', u) = approx prec b vs"
+    and approx_form': "approx_form' prec f (ss ! n) n l u vs ss"
+    by (cases "approx prec a vs", simp,
+        cases "approx prec b vs", auto) blast
+
+  { assume "xs ! n \<in> { interpret_floatarith a xs .. interpret_floatarith b xs }"
+    with approx[OF Bound.prems(2) l_eq] and approx[OF Bound.prems(2) u_eq]
+    have "xs ! n \<in> { real l .. real u}" by auto
+
+    from approx_form_approx_form'[OF approx_form' this]
+    obtain lx ux where bnds: "xs ! n \<in> { real lx .. real ux }"
+      and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" .
+
+    from `bounded_by xs vs` bnds
+    have "bounded_by xs (vs[n := Some (lx, ux)])" by (rule bounded_by_update)
+    with Bound.hyps[OF approx_form]
+    have "interpret_form f xs" by blast }
+  thus ?case using interpret_form.simps x_eq and interpret_floatarith.simps by simp
+next
+  case (Assign x a f)
+  then obtain n
+    where x_eq: "x = Atom n" by (cases x) auto
+
+  with Assign.prems obtain l u' l' u
+    where bnd_eq: "Some (l, u) = approx prec a vs"
+    and x_eq: "x = Atom n"
+    and approx_form': "approx_form' prec f (ss ! n) n l u vs ss"
+    by (cases "approx prec a vs") auto
+
+  { assume bnds: "xs ! n = interpret_floatarith a xs"
+    with approx[OF Assign.prems(2) bnd_eq]
+    have "xs ! n \<in> { real l .. real u}" by auto
+    from approx_form_approx_form'[OF approx_form' this]
+    obtain lx ux where bnds: "xs ! n \<in> { real lx .. real ux }"
+      and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" .
+
+    from `bounded_by xs vs` bnds
+    have "bounded_by xs (vs[n := Some (lx, ux)])" by (rule bounded_by_update)
+    with Assign.hyps[OF approx_form]
+    have "interpret_form f xs" by blast }
+  thus ?case using interpret_form.simps x_eq and interpret_floatarith.simps by simp
+next
   case (Less a b)
-  show ?thesis
-  proof (cases "\<exists> u l u' l'. approx prec a bs = Some (l, u) \<and> 
-                             approx prec b bs = Some (l', u')")
-    case True
-    then obtain l u l' u' where a_approx: "approx prec a bs = Some (l, u)"
-      and b_approx: "approx prec b bs = Some (l', u') " by auto
-    with `approx_inequality prec eq bs` have "real u < real l'"
-      unfolding Less approx_inequality.simps less_float_def by auto
-    moreover from a_approx[symmetric] and b_approx[symmetric] and `bounded_by vs bs`
-    have "interpret_floatarith a vs \<le> real u" and "real l' \<le> interpret_floatarith b vs"
-      using approx by auto
-    ultimately show ?thesis unfolding interpret_inequality.simps Less by auto
-  next
-    case False
-    hence "approx prec a bs = None \<or> approx prec b bs = None"
-      unfolding not_Some_eq[symmetric] by auto
-    hence "\<not> approx_inequality prec eq bs" unfolding Less approx_inequality.simps 
-      by (cases "approx prec a bs = None", auto)
-    thus ?thesis using assms by auto
-  qed
+  then obtain l u l' u'
+    where l_eq: "Some (l, u) = approx prec a vs"
+    and u_eq: "Some (l', u') = approx prec b vs"
+    and inequality: "u < l'"
+    by (cases "approx prec a vs", auto,
+      cases "approx prec b vs", auto)
+  from inequality[unfolded less_float_def] approx[OF Less.prems(2) l_eq] approx[OF Less.prems(2) u_eq]
+  show ?case by auto
 next
   case (LessEqual a b)
-  show ?thesis
-  proof (cases "\<exists> u l u' l'. approx prec a bs = Some (l, u) \<and> 
-                             approx prec b bs = Some (l', u')")
+  then obtain l u l' u'
+    where l_eq: "Some (l, u) = approx prec a vs"
+    and u_eq: "Some (l', u') = approx prec b vs"
+    and inequality: "u \<le> l'"
+    by (cases "approx prec a vs", auto,
+      cases "approx prec b vs", auto)
+  from inequality[unfolded le_float_def] approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq]
+  show ?case by auto
+next
+  case (AtLeastAtMost x a b)
+  then obtain lx ux l u l' u'
+    where x_eq: "Some (lx, ux) = approx prec x vs"
+    and l_eq: "Some (l, u) = approx prec a vs"
+    and u_eq: "Some (l', u') = approx prec b vs"
+    and inequality: "u \<le> lx \<and> ux \<le> l'"
+    by (cases "approx prec x vs", auto,
+      cases "approx prec a vs", auto,
+      cases "approx prec b vs", auto, blast)
+  from inequality[unfolded le_float_def] approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq]
+  show ?case by auto
+qed
+
+lemma approx_form:
+  assumes "n = length xs"
+  assumes "approx_form prec f (replicate n None) ss"
+  shows "interpret_form f xs"
+  using approx_form_aux[OF _ bounded_by_None] assms by auto
+
+subsection {* Implementing Taylor series expansion *}
+
+fun isDERIV :: "nat \<Rightarrow> floatarith \<Rightarrow> real list \<Rightarrow> bool" where
+"isDERIV x (Add a b) vs         = (isDERIV x a vs \<and> isDERIV x b vs)" |
+"isDERIV x (Mult a b) vs        = (isDERIV x a vs \<and> isDERIV x b vs)" |
+"isDERIV x (Minus a) vs         = isDERIV x a vs" |
+"isDERIV x (Inverse a) vs       = (isDERIV x a vs \<and> interpret_floatarith a vs \<noteq> 0)" |
+"isDERIV x (Cos a) vs           = isDERIV x a vs" |
+"isDERIV x (Arctan a) vs        = isDERIV x a vs" |
+"isDERIV x (Min a b) vs         = False" |
+"isDERIV x (Max a b) vs         = False" |
+"isDERIV x (Abs a) vs           = False" |
+"isDERIV x Pi vs                = True" |
+"isDERIV x (Sqrt a) vs          = (isDERIV x a vs \<and> interpret_floatarith a vs > 0)" |
+"isDERIV x (Exp a) vs           = isDERIV x a vs" |
+"isDERIV x (Ln a) vs            = (isDERIV x a vs \<and> interpret_floatarith a vs > 0)" |
+"isDERIV x (Power a 0) vs       = True" |
+"isDERIV x (Power a (Suc n)) vs = isDERIV x a vs" |
+"isDERIV x (Num f) vs           = True" |
+"isDERIV x (Atom n) vs          = True"
+
+fun DERIV_floatarith :: "nat \<Rightarrow> floatarith \<Rightarrow> floatarith" where
+"DERIV_floatarith x (Add a b)         = Add (DERIV_floatarith x a) (DERIV_floatarith x b)" |
+"DERIV_floatarith x (Mult a b)        = Add (Mult a (DERIV_floatarith x b)) (Mult (DERIV_floatarith x a) b)" |
+"DERIV_floatarith x (Minus a)         = Minus (DERIV_floatarith x a)" |
+"DERIV_floatarith x (Inverse a)       = Minus (Mult (DERIV_floatarith x a) (Inverse (Power a 2)))" |
+"DERIV_floatarith x (Cos a)           = Minus (Mult (Cos (Add (Mult Pi (Num (Float 1 -1))) (Minus a))) (DERIV_floatarith x a))" |
+"DERIV_floatarith x (Arctan a)        = Mult (Inverse (Add (Num 1) (Power a 2))) (DERIV_floatarith x a)" |
+"DERIV_floatarith x (Min a b)         = Num 0" |
+"DERIV_floatarith x (Max a b)         = Num 0" |
+"DERIV_floatarith x (Abs a)           = Num 0" |
+"DERIV_floatarith x Pi                = Num 0" |
+"DERIV_floatarith x (Sqrt a)          = (Mult (Inverse (Mult (Sqrt a) (Num 2))) (DERIV_floatarith x a))" |
+"DERIV_floatarith x (Exp a)           = Mult (Exp a) (DERIV_floatarith x a)" |
+"DERIV_floatarith x (Ln a)            = Mult (Inverse a) (DERIV_floatarith x a)" |
+"DERIV_floatarith x (Power a 0)       = Num 0" |
+"DERIV_floatarith x (Power a (Suc n)) = Mult (Num (Float (int (Suc n)) 0)) (Mult (Power a n) (DERIV_floatarith x a))" |
+"DERIV_floatarith x (Num f)           = Num 0" |
+"DERIV_floatarith x (Atom n)          = (if x = n then Num 1 else Num 0)"
+
+lemma DERIV_floatarith:
+  assumes "n < length vs"
+  assumes isDERIV: "isDERIV n f (vs[n := x])"
+  shows "DERIV (\<lambda> x'. interpret_floatarith f (vs[n := x'])) x :>
+               interpret_floatarith (DERIV_floatarith n f) (vs[n := x])"
+   (is "DERIV (?i f) x :> _")
+using isDERIV proof (induct f arbitrary: x)
+     case (Inverse a) thus ?case
+    by (auto intro!: DERIV_intros
+             simp add: algebra_simps power2_eq_square)
+next case (Cos a) thus ?case
+  by (auto intro!: DERIV_intros
+           simp del: interpret_floatarith.simps(5)
+           simp add: interpret_floatarith_sin interpret_floatarith.simps(5)[of a])
+next case (Power a n) thus ?case
+  by (cases n, auto intro!: DERIV_intros
+                    simp del: power_Suc simp add: real_eq_of_nat)
+next case (Ln a) thus ?case
+    by (auto intro!: DERIV_intros simp add: divide_inverse)
+next case (Atom i) thus ?case using `n < length vs` by auto
+qed (auto intro!: DERIV_intros)
+
+declare approx.simps[simp del]
+
+fun isDERIV_approx :: "nat \<Rightarrow> nat \<Rightarrow> floatarith \<Rightarrow> (float * float) option list \<Rightarrow> bool" where
+"isDERIV_approx prec x (Add a b) vs         = (isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs)" |
+"isDERIV_approx prec x (Mult a b) vs        = (isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs)" |
+"isDERIV_approx prec x (Minus a) vs         = isDERIV_approx prec x a vs" |
+"isDERIV_approx prec x (Inverse a) vs       =
+  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l \<or> u < 0 | None \<Rightarrow> False))" |
+"isDERIV_approx prec x (Cos a) vs           = isDERIV_approx prec x a vs" |
+"isDERIV_approx prec x (Arctan a) vs        = isDERIV_approx prec x a vs" |
+"isDERIV_approx prec x (Min a b) vs         = False" |
+"isDERIV_approx prec x (Max a b) vs         = False" |
+"isDERIV_approx prec x (Abs a) vs           = False" |
+"isDERIV_approx prec x Pi vs                = True" |
+"isDERIV_approx prec x (Sqrt a) vs          =
+  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l | None \<Rightarrow> False))" |
+"isDERIV_approx prec x (Exp a) vs           = isDERIV_approx prec x a vs" |
+"isDERIV_approx prec x (Ln a) vs            =
+  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l | None \<Rightarrow> False))" |
+"isDERIV_approx prec x (Power a 0) vs       = True" |
+"isDERIV_approx prec x (Power a (Suc n)) vs = isDERIV_approx prec x a vs" |
+"isDERIV_approx prec x (Num f) vs           = True" |
+"isDERIV_approx prec x (Atom n) vs          = True"
+
+lemma isDERIV_approx:
+  assumes "bounded_by xs vs"
+  and isDERIV_approx: "isDERIV_approx prec x f vs"
+  shows "isDERIV x f xs"
+using isDERIV_approx proof (induct f)
+  case (Inverse a)
+  then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
+    and *: "0 < l \<or> u < 0"
+    by (cases "approx prec a vs", auto)
+  with approx[OF `bounded_by xs vs` approx_Some]
+  have "interpret_floatarith a xs \<noteq> 0" unfolding less_float_def by auto
+  thus ?case using Inverse by auto
+next
+  case (Ln a)
+  then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
+    and *: "0 < l"
+    by (cases "approx prec a vs", auto)
+  with approx[OF `bounded_by xs vs` approx_Some]
+  have "0 < interpret_floatarith a xs" unfolding less_float_def by auto
+  thus ?case using Ln by auto
+next
+  case (Sqrt a)
+  then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
+    and *: "0 < l"
+    by (cases "approx prec a vs", auto)
+  with approx[OF `bounded_by xs vs` approx_Some]
+  have "0 < interpret_floatarith a xs" unfolding less_float_def by auto
+  thus ?case using Sqrt by auto
+next
+  case (Power a n) thus ?case by (cases n, auto)
+qed auto
+
+lemma bounded_by_update_var:
+  assumes "bounded_by xs vs" and "vs ! i = Some (l, u)"
+  and bnd: "x \<in> { real l .. real u }"
+  shows "bounded_by (xs[i := x]) vs"
+proof (cases "i < length xs")
+  case False thus ?thesis using `bounded_by xs vs` by auto
+next
+  let ?xs = "xs[i := x]"
+  case True hence "i < length ?xs" by auto
+{ fix j
+  assume "j < length vs"
+  have "case vs ! j of None \<Rightarrow> True | Some (l, u) \<Rightarrow> ?xs ! j \<in> { real l .. real u }"
+  proof (cases "vs ! j")
+    case (Some b)
+    thus ?thesis
+    proof (cases "i = j")
+      case True
+      thus ?thesis using `vs ! i = Some (l, u)` Some and bnd `i < length ?xs`
+	by auto
+    next
+      case False
+      thus ?thesis using `bounded_by xs vs`[THEN bounded_byE, OF `j < length vs`] Some
+	by auto
+    qed
+  qed auto }
+  thus ?thesis unfolding bounded_by_def by auto
+qed
+
+lemma isDERIV_approx':
+  assumes "bounded_by xs vs"
+  and vs_x: "vs ! x = Some (l, u)" and X_in: "X \<in> { real l .. real u }"
+  and approx: "isDERIV_approx prec x f vs"
+  shows "isDERIV x f (xs[x := X])"
+proof -
+  note bounded_by_update_var[OF `bounded_by xs vs` vs_x X_in] approx
+  thus ?thesis by (rule isDERIV_approx)
+qed
+
+lemma DERIV_approx:
+  assumes "n < length xs" and bnd: "bounded_by xs vs"
+  and isD: "isDERIV_approx prec n f vs"
+  and app: "Some (l, u) = approx prec (DERIV_floatarith n f) vs" (is "_ = approx _ ?D _")
+  shows "\<exists>x. real l \<le> x \<and> x \<le> real u \<and>
+             DERIV (\<lambda> x. interpret_floatarith f (xs[n := x])) (xs!n) :> x"
+         (is "\<exists> x. _ \<and> _ \<and> DERIV (?i f) _ :> _")
+proof (rule exI[of _ "?i ?D (xs!n)"], rule conjI[OF _ conjI])
+  let "?i f x" = "interpret_floatarith f (xs[n := x])"
+  from approx[OF bnd app]
+  show "real l \<le> ?i ?D (xs!n)" and "?i ?D (xs!n) \<le> real u"
+    using `n < length xs` by auto
+  from DERIV_floatarith[OF `n < length xs`, of f "xs!n"] isDERIV_approx[OF bnd isD]
+  show "DERIV (?i f) (xs!n) :> (?i ?D (xs!n))" by simp
+qed
+
+fun lift_bin :: "(float * float) option \<Rightarrow> (float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> (float * float) option) \<Rightarrow> (float * float) option" where
+"lift_bin (Some (l1, u1)) (Some (l2, u2)) f = f l1 u1 l2 u2" |
+"lift_bin a b f = None"
+
+lemma lift_bin:
+  assumes lift_bin_Some: "Some (l, u) = lift_bin a b f"
+  obtains l1 u1 l2 u2
+  where "a = Some (l1, u1)"
+  and "b = Some (l2, u2)"
+  and "f l1 u1 l2 u2 = Some (l, u)"
+using assms by (cases a, simp, cases b, simp, auto)
+
+fun approx_tse where
+"approx_tse prec n 0 c k f bs = approx prec f bs" |
+"approx_tse prec n (Suc s) c k f bs =
+  (if isDERIV_approx prec n f bs then
+    lift_bin (approx prec f (bs[n := Some (c,c)]))
+             (approx_tse prec n s c (Suc k) (DERIV_floatarith n f) bs)
+             (\<lambda> l1 u1 l2 u2. approx prec
+                 (Add (Atom 0)
+                      (Mult (Inverse (Num (Float (int k) 0)))
+                                 (Mult (Add (Atom (Suc (Suc 0))) (Minus (Num c)))
+                                       (Atom (Suc 0))))) [Some (l1, u1), Some (l2, u2), bs!n])
+  else approx prec f bs)"
+
+lemma bounded_by_Cons:
+  assumes bnd: "bounded_by xs vs"
+  and x: "x \<in> { real l .. real u }"
+  shows "bounded_by (x#xs) ((Some (l, u))#vs)"
+proof -
+  { fix i assume *: "i < length ((Some (l, u))#vs)"
+    have "case ((Some (l,u))#vs) ! i of Some (l, u) \<Rightarrow> (x#xs)!i \<in> { real l .. real u } | None \<Rightarrow> True"
+    proof (cases i)
+      case 0 with x show ?thesis by auto
+    next
+      case (Suc i) with * have "i < length vs" by auto
+      from bnd[THEN bounded_byE, OF this]
+      show ?thesis unfolding Suc nth_Cons_Suc .
+    qed }
+  thus ?thesis by (auto simp add: bounded_by_def)
+qed
+
+lemma approx_tse_generic:
+  assumes "bounded_by xs vs"
+  and bnd_c: "bounded_by (xs[x := real c]) vs" and "x < length vs" and "x < length xs"
+  and bnd_x: "vs ! x = Some (lx, ux)"
+  and ate: "Some (l, u) = approx_tse prec x s c k f vs"
+  shows "\<exists> n. (\<forall> m < n. \<forall> z \<in> {real lx .. real ux}.
+      DERIV (\<lambda> y. interpret_floatarith ((DERIV_floatarith x ^^ m) f) (xs[x := y])) z :>
+            (interpret_floatarith ((DERIV_floatarith x ^^ (Suc m)) f) (xs[x := z])))
+   \<and> (\<forall> t \<in> {real lx .. real ux}.  (\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) *
+                  interpret_floatarith ((DERIV_floatarith x ^^ i) f) (xs[x := real c]) *
+                  (xs!x - real c)^i) +
+      inverse (real (\<Prod> j \<in> {k..<k+n}. j)) *
+      interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := t]) *
+      (xs!x - real c)^n \<in> {real l .. real u})" (is "\<exists> n. ?taylor f k l u n")
+using ate proof (induct s arbitrary: k f l u)
+  case 0
+  { fix t assume "t \<in> {real lx .. real ux}"
+    note bounded_by_update_var[OF `bounded_by xs vs` bnd_x this]
+    from approx[OF this 0[unfolded approx_tse.simps]]
+    have "(interpret_floatarith f (xs[x := t])) \<in> {real l .. real u}"
+      by (auto simp add: algebra_simps)
+  } thus ?case by (auto intro!: exI[of _ 0])
+next
+  case (Suc s)
+  show ?case
+  proof (cases "isDERIV_approx prec x f vs")
+    case False
+    note ap = Suc.prems[unfolded approx_tse.simps if_not_P[OF False]]
+
+    { fix t assume "t \<in> {real lx .. real ux}"
+      note bounded_by_update_var[OF `bounded_by xs vs` bnd_x this]
+      from approx[OF this ap]
+      have "(interpret_floatarith f (xs[x := t])) \<in> {real l .. real u}"
+	by (auto simp add: algebra_simps)
+    } thus ?thesis by (auto intro!: exI[of _ 0])
+  next
     case True
-    then obtain l u l' u' where a_approx: "approx prec a bs = Some (l, u)"
-      and b_approx: "approx prec b bs = Some (l', u') " by auto
-    with `approx_inequality prec eq bs` have "real u \<le> real l'"
-      unfolding LessEqual approx_inequality.simps le_float_def by auto
-    moreover from a_approx[symmetric] and b_approx[symmetric] and `bounded_by vs bs`
-    have "interpret_floatarith a vs \<le> real u" and "real l' \<le> interpret_floatarith b vs"
-      using approx by auto
-    ultimately show ?thesis unfolding interpret_inequality.simps LessEqual by auto
-  next
-    case False
-    hence "approx prec a bs = None \<or> approx prec b bs = None"
-      unfolding not_Some_eq[symmetric] by auto
-    hence "\<not> approx_inequality prec eq bs" unfolding LessEqual approx_inequality.simps 
-      by (cases "approx prec a bs = None", auto)
-    thus ?thesis using assms by auto
+    with Suc.prems
+    obtain l1 u1 l2 u2
+      where a: "Some (l1, u1) = approx prec f (vs[x := Some (c,c)])"
+      and ate: "Some (l2, u2) = approx_tse prec x s c (Suc k) (DERIV_floatarith x f) vs"
+      and final: "Some (l, u) = approx prec
+        (Add (Atom 0)
+             (Mult (Inverse (Num (Float (int k) 0)))
+                   (Mult (Add (Atom (Suc (Suc 0))) (Minus (Num c)))
+                         (Atom (Suc 0))))) [Some (l1, u1), Some (l2, u2), vs!x]"
+      by (auto elim!: lift_bin) blast
+
+    from bnd_c `x < length xs`
+    have bnd: "bounded_by (xs[x:=real c]) (vs[x:= Some (c,c)])"
+      by (auto intro!: bounded_by_update)
+
+    from approx[OF this a]
+    have f_c: "interpret_floatarith ((DERIV_floatarith x ^^ 0) f) (xs[x := real c]) \<in> { real l1 .. real u1 }"
+              (is "?f 0 (real c) \<in> _")
+      by auto
+
+    { fix f :: "'a \<Rightarrow> 'a" fix n :: nat fix x :: 'a
+      have "(f ^^ Suc n) x = (f ^^ n) (f x)"
+	by (induct n, auto) }
+    note funpow_Suc = this[symmetric]
+    from Suc.hyps[OF ate, unfolded this]
+    obtain n
+      where DERIV_hyp: "\<And> m z. \<lbrakk> m < n ; z \<in> { real lx .. real ux } \<rbrakk> \<Longrightarrow> DERIV (?f (Suc m)) z :> ?f (Suc (Suc m)) z"
+      and hyp: "\<forall> t \<in> {real lx .. real ux}. (\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {Suc k..<Suc k + i}. j)) * ?f (Suc i) (real c) * (xs!x - real c)^i) +
+           inverse (real (\<Prod> j \<in> {Suc k..<Suc k + n}. j)) * ?f (Suc n) t * (xs!x - real c)^n \<in> {real l2 .. real u2}"
+          (is "\<forall> t \<in> _. ?X (Suc k) f n t \<in> _")
+      by blast
+
+    { fix m z
+      assume "m < Suc n" and bnd_z: "z \<in> { real lx .. real ux }"
+      have "DERIV (?f m) z :> ?f (Suc m) z"
+      proof (cases m)
+	case 0
+	with DERIV_floatarith[OF `x < length xs` isDERIV_approx'[OF `bounded_by xs vs` bnd_x bnd_z True]]
+	show ?thesis by simp
+      next
+	case (Suc m')
+	hence "m' < n" using `m < Suc n` by auto
+	from DERIV_hyp[OF this bnd_z]
+	show ?thesis using Suc by simp
+      qed } note DERIV = this
+
+    have "\<And> k i. k < i \<Longrightarrow> {k ..< i} = insert k {Suc k ..< i}" by auto
+    hence setprod_head_Suc: "\<And> k i. \<Prod> {k ..< k + Suc i} = k * \<Prod> {Suc k ..< Suc k + i}" by auto
+    have setsum_move0: "\<And> k F. setsum F {0..<Suc k} = F 0 + setsum (\<lambda> k. F (Suc k)) {0..<k}"
+      unfolding setsum_shift_bounds_Suc_ivl[symmetric]
+      unfolding setsum_head_upt_Suc[OF zero_less_Suc] ..
+    def C \<equiv> "xs!x - real c"
+
+    { fix t assume t: "t \<in> {real lx .. real ux}"
+      hence "bounded_by [xs!x] [vs!x]"
+	using `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`]
+	by (cases "vs!x", auto simp add: bounded_by_def)
+
+      with hyp[THEN bspec, OF t] f_c
+      have "bounded_by [?f 0 (real c), ?X (Suc k) f n t, xs!x] [Some (l1, u1), Some (l2, u2), vs!x]"
+	by (auto intro!: bounded_by_Cons)
+      from approx[OF this final, unfolded atLeastAtMost_iff[symmetric]]
+      have "?X (Suc k) f n t * (xs!x - real c) * inverse (real k) + ?f 0 (real c) \<in> {real l .. real u}"
+	by (auto simp add: algebra_simps)
+      also have "?X (Suc k) f n t * (xs!x - real c) * inverse (real k) + ?f 0 (real c) =
+               (\<Sum> i = 0..<Suc n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) * ?f i (real c) * (xs!x - real c)^i) +
+               inverse (real (\<Prod> j \<in> {k..<k+Suc n}. j)) * ?f (Suc n) t * (xs!x - real c)^Suc n" (is "_ = ?T")
+	unfolding funpow_Suc C_def[symmetric] setsum_move0 setprod_head_Suc
+	by (auto simp add: algebra_simps setsum_right_distrib[symmetric])
+      finally have "?T \<in> {real l .. real u}" . }
+    thus ?thesis using DERIV by blast
   qed
 qed
 
-lemma interpret_floatarith_divide: "interpret_floatarith (Mult a (Inverse b)) vs = (interpret_floatarith a vs) / (interpret_floatarith b vs)"
-  unfolding real_divide_def interpret_floatarith.simps ..
-
-lemma interpret_floatarith_diff: "interpret_floatarith (Add a (Minus b)) vs = (interpret_floatarith a vs) - (interpret_floatarith b vs)"
-  unfolding real_diff_def interpret_floatarith.simps ..
-
-lemma interpret_floatarith_sin: "interpret_floatarith (Cos (Add (Mult Pi (Num (Float 1 -1))) (Minus a))) vs = 
-  sin (interpret_floatarith a vs)"
-  unfolding sin_cos_eq interpret_floatarith.simps
-            interpret_floatarith_divide interpret_floatarith_diff real_diff_def
-  by auto
-
-lemma interpret_floatarith_tan:
-  "interpret_floatarith (Mult (Cos (Add (Mult Pi (Num (Float 1 -1))) (Minus a))) (Inverse (Cos a))) vs =
-   tan (interpret_floatarith a vs)"
-  unfolding interpret_floatarith.simps(3,4) interpret_floatarith_sin tan_def real_divide_def
-  by auto
-
-lemma interpret_floatarith_powr: "interpret_floatarith (Exp (Mult b (Ln a))) vs = (interpret_floatarith a vs) powr (interpret_floatarith b vs)"
-  unfolding powr_def interpret_floatarith.simps ..
-
-lemma interpret_floatarith_log: "interpret_floatarith ((Mult (Ln x) (Inverse (Ln b)))) vs = log (interpret_floatarith b vs) (interpret_floatarith x vs)"
-  unfolding log_def interpret_floatarith.simps real_divide_def ..
-
-lemma interpret_floatarith_num:
-  shows "interpret_floatarith (Num (Float 0 0)) vs = 0"
-  and "interpret_floatarith (Num (Float 1 0)) vs = 1"
-  and "interpret_floatarith (Num (Float (number_of a) 0)) vs = number_of a" by auto
+lemma setprod_fact: "\<Prod> {1..<1 + k} = fact (k :: nat)"
+proof (induct k)
+  case (Suc k)
+  have "{ 1 ..< Suc (Suc k) } = insert (Suc k) { 1 ..< Suc k }" by auto
+  hence "\<Prod> { 1 ..< Suc (Suc k) } = (Suc k) * \<Prod> { 1 ..< Suc k }" by auto
+  thus ?case using Suc by auto
+qed simp
+
+lemma approx_tse:
+  assumes "bounded_by xs vs"
+  and bnd_x: "vs ! x = Some (lx, ux)" and bnd_c: "real c \<in> {real lx .. real ux}"
+  and "x < length vs" and "x < length xs"
+  and ate: "Some (l, u) = approx_tse prec x s c 1 f vs"
+  shows "interpret_floatarith f xs \<in> { real l .. real u }"
+proof -
+  def F \<equiv> "\<lambda> n z. interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := z])"
+  hence F0: "F 0 = (\<lambda> z. interpret_floatarith f (xs[x := z]))" by auto
+
+  hence "bounded_by (xs[x := real c]) vs" and "x < length vs" "x < length xs"
+    using `bounded_by xs vs` bnd_x bnd_c `x < length vs` `x < length xs`
+    by (auto intro!: bounded_by_update_var)
+
+  from approx_tse_generic[OF `bounded_by xs vs` this bnd_x ate]
+  obtain n
+    where DERIV: "\<forall> m z. m < n \<and> real lx \<le> z \<and> z \<le> real ux \<longrightarrow> DERIV (F m) z :> F (Suc m) z"
+    and hyp: "\<And> t. t \<in> {real lx .. real ux} \<Longrightarrow>
+           (\<Sum> j = 0..<n. inverse (real (fact j)) * F j (real c) * (xs!x - real c)^j) +
+             inverse (real (fact n)) * F n t * (xs!x - real c)^n
+             \<in> {real l .. real u}" (is "\<And> t. _ \<Longrightarrow> ?taylor t \<in> _")
+    unfolding F_def atLeastAtMost_iff[symmetric] setprod_fact by blast
+
+  have bnd_xs: "xs ! x \<in> { real lx .. real ux }"
+    using `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`] bnd_x by auto
+
+  show ?thesis
+  proof (cases n)
+    case 0 thus ?thesis using hyp[OF bnd_xs] unfolding F_def by auto
+  next
+    case (Suc n')
+    show ?thesis
+    proof (cases "xs ! x = real c")
+      case True
+      from True[symmetric] hyp[OF bnd_xs] Suc show ?thesis
+	unfolding F_def Suc setsum_head_upt_Suc[OF zero_less_Suc] setsum_shift_bounds_Suc_ivl by auto
+    next
+      case False
+
+      have "real lx \<le> real c" "real c \<le> real ux" "real lx \<le> xs!x" "xs!x \<le> real ux"
+	using Suc bnd_c `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`] bnd_x by auto
+      from Taylor.taylor[OF zero_less_Suc, of F, OF F0 DERIV[unfolded Suc] this False]
+      obtain t where t_bnd: "if xs ! x < real c then xs ! x < t \<and> t < real c else real c < t \<and> t < xs ! x"
+	and fl_eq: "interpret_floatarith f (xs[x := xs ! x]) =
+	   (\<Sum>m = 0..<Suc n'. F m (real c) / real (fact m) * (xs ! x - real c) ^ m) +
+           F (Suc n') t / real (fact (Suc n')) * (xs ! x - real c) ^ Suc n'"
+	by blast
+
+      from t_bnd bnd_xs bnd_c have *: "t \<in> {real lx .. real ux}"
+	by (cases "xs ! x < real c", auto)
+
+      have "interpret_floatarith f (xs[x := xs ! x]) = ?taylor t"
+	unfolding fl_eq Suc by (auto simp add: algebra_simps divide_inverse)
+      also have "\<dots> \<in> {real l .. real u}" using * by (rule hyp)
+      finally show ?thesis by simp
+    qed
+  qed
+qed
+
+fun approx_tse_form' where
+"approx_tse_form' prec t f 0 l u cmp =
+  (case approx_tse prec 0 t ((l + u) * Float 1 -1) 1 f [Some (l, u)]
+     of Some (l, u) \<Rightarrow> cmp l u | None \<Rightarrow> False)" |
+"approx_tse_form' prec t f (Suc s) l u cmp =
+  (let m = (l + u) * Float 1 -1
+   in approx_tse_form' prec t f s l m cmp \<and>
+      approx_tse_form' prec t f s m u cmp)"
+
+lemma approx_tse_form':
+  assumes "approx_tse_form' prec t f s l u cmp" and "x \<in> {real l .. real u}"
+  shows "\<exists> l' u' ly uy. x \<in> { real l' .. real u' } \<and> real l \<le> real l' \<and> real u' \<le> real u \<and> cmp ly uy \<and>
+                  approx_tse prec 0 t ((l' + u') * Float 1 -1) 1 f [Some (l', u')] = Some (ly, uy)"
+using assms proof (induct s arbitrary: l u)
+  case 0
+  then obtain ly uy
+    where *: "approx_tse prec 0 t ((l + u) * Float 1 -1) 1 f [Some (l, u)] = Some (ly, uy)"
+    and **: "cmp ly uy" by (auto elim!: option_caseE)
+  with 0 show ?case by (auto intro!: exI)
+next
+  case (Suc s)
+  let ?m = "(l + u) * Float 1 -1"
+  from Suc.prems
+  have l: "approx_tse_form' prec t f s l ?m cmp"
+    and u: "approx_tse_form' prec t f s ?m u cmp"
+    by (auto simp add: Let_def)
+
+  have m_l: "real l \<le> real ?m" and m_u: "real ?m \<le> real u"
+    unfolding le_float_def using Suc.prems by auto
+
+  with `x \<in> { real l .. real u }`
+  have "x \<in> { real l .. real ?m} \<or> x \<in> { real ?m .. real u }" by auto
+  thus ?case
+  proof (rule disjE)
+    assume "x \<in> { real l .. real ?m}"
+    from Suc.hyps[OF l this]
+    obtain l' u' ly uy
+      where "x \<in> { real l' .. real u' } \<and> real l \<le> real l' \<and> real u' \<le> real ?m \<and> cmp ly uy \<and>
+                  approx_tse prec 0 t ((l' + u') * Float 1 -1) 1 f [Some (l', u')] = Some (ly, uy)" by blast
+    with m_u show ?thesis by (auto intro!: exI)
+  next
+    assume "x \<in> { real ?m .. real u }"
+    from Suc.hyps[OF u this]
+    obtain l' u' ly uy
+      where "x \<in> { real l' .. real u' } \<and> real ?m \<le> real l' \<and> real u' \<le> real u \<and> cmp ly uy \<and>
+                  approx_tse prec 0 t ((l' + u') * Float 1 -1) 1 f [Some (l', u')] = Some (ly, uy)" by blast
+    with m_u show ?thesis by (auto intro!: exI)
+  qed
+qed
+
+lemma approx_tse_form'_less:
+  assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s l u (\<lambda> l u. 0 < l)"
+  and x: "x \<in> {real l .. real u}"
+  shows "interpret_floatarith b [x] < interpret_floatarith a [x]"
+proof -
+  from approx_tse_form'[OF tse x]
+  obtain l' u' ly uy
+    where x': "x \<in> { real l' .. real u' }" and "real l \<le> real l'"
+    and "real u' \<le> real u" and "0 < ly"
+    and tse: "approx_tse prec 0 t ((l' + u') * Float 1 -1) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)"
+    by blast
+
+  hence "bounded_by [x] [Some (l', u')]" by (auto simp add: bounded_by_def)
+
+  from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
+  have "real ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
+    by (auto simp add: diff_minus)
+  from order_less_le_trans[OF `0 < ly`[unfolded less_float_def] this]
+  show ?thesis by auto
+qed
+
+lemma approx_tse_form'_le:
+  assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s l u (\<lambda> l u. 0 \<le> l)"
+  and x: "x \<in> {real l .. real u}"
+  shows "interpret_floatarith b [x] \<le> interpret_floatarith a [x]"
+proof -
+  from approx_tse_form'[OF tse x]
+  obtain l' u' ly uy
+    where x': "x \<in> { real l' .. real u' }" and "real l \<le> real l'"
+    and "real u' \<le> real u" and "0 \<le> ly"
+    and tse: "approx_tse prec 0 t ((l' + u') * Float 1 -1) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)"
+    by blast
+
+  hence "bounded_by [x] [Some (l', u')]" by (auto simp add: bounded_by_def)
+
+  from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
+  have "real ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
+    by (auto simp add: diff_minus)
+  from order_trans[OF `0 \<le> ly`[unfolded le_float_def] this]
+  show ?thesis by auto
+qed
+
+definition
+"approx_tse_form prec t s f =
+  (case f
+   of (Bound x a b f) \<Rightarrow> x = Atom 0 \<and>
+     (case (approx prec a [None], approx prec b [None])
+      of (Some (l, u), Some (l', u')) \<Rightarrow>
+        (case f
+         of Less lf rt \<Rightarrow> approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 < l)
+          | LessEqual lf rt \<Rightarrow> approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)
+          | AtLeastAtMost x lf rt \<Rightarrow>
+            approx_tse_form' prec t (Add x (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l) \<and>
+            approx_tse_form' prec t (Add rt (Minus x)) s l u' (\<lambda> l u. 0 \<le> l)
+          | _ \<Rightarrow> False)
+       | _ \<Rightarrow> False)
+   | _ \<Rightarrow> False)"
+
+lemma approx_tse_form:
+  assumes "approx_tse_form prec t s f"
+  shows "interpret_form f [x]"
+proof (cases f)
+  case (Bound i a b f') note f_def = this
+  with assms obtain l u l' u'
+    where a: "approx prec a [None] = Some (l, u)"
+    and b: "approx prec b [None] = Some (l', u')"
+    unfolding approx_tse_form_def by (auto elim!: option_caseE)
+
+  from Bound assms have "i = Atom 0" unfolding approx_tse_form_def by auto
+  hence i: "interpret_floatarith i [x] = x" by auto
+
+  { let "?f z" = "interpret_floatarith z [x]"
+    assume "?f i \<in> { ?f a .. ?f b }"
+    with approx[OF _ a[symmetric], of "[x]"] approx[OF _ b[symmetric], of "[x]"]
+    have bnd: "x \<in> { real l .. real u'}" unfolding bounded_by_def i by auto
+
+    have "interpret_form f' [x]"
+    proof (cases f')
+      case (Less lf rt)
+      with Bound a b assms
+      have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 < l)"
+	unfolding approx_tse_form_def by auto
+      from approx_tse_form'_less[OF this bnd]
+      show ?thesis using Less by auto
+    next
+      case (LessEqual lf rt)
+      with Bound a b assms
+      have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"
+	unfolding approx_tse_form_def by auto
+      from approx_tse_form'_le[OF this bnd]
+      show ?thesis using LessEqual by auto
+    next
+      case (AtLeastAtMost x lf rt)
+      with Bound a b assms
+      have "approx_tse_form' prec t (Add rt (Minus x)) s l u' (\<lambda> l u. 0 \<le> l)"
+	and "approx_tse_form' prec t (Add x (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"
+	unfolding approx_tse_form_def by auto
+      from approx_tse_form'_le[OF this(1) bnd] approx_tse_form'_le[OF this(2) bnd]
+      show ?thesis using AtLeastAtMost by auto
+    next
+      case (Bound x a b f') with assms
+      show ?thesis by (auto elim!: option_caseE simp add: f_def approx_tse_form_def)
+    next
+      case (Assign x a f') with assms
+      show ?thesis by (auto elim!: option_caseE simp add: f_def approx_tse_form_def)
+    qed } thus ?thesis unfolding f_def by auto
+next case Assign with assms show ?thesis by (auto simp add: approx_tse_form_def)
+next case LessEqual with assms show ?thesis by (auto simp add: approx_tse_form_def)
+next case Less with assms show ?thesis by (auto simp add: approx_tse_form_def)
+next case AtLeastAtMost with assms show ?thesis by (auto simp add: approx_tse_form_def)
+qed
 
 subsection {* Implement proof method \texttt{approximation} *}
 
-lemma bounded_divl: assumes "real a / real b \<le> x" shows "real (float_divl p a b) \<le> x" by (rule order_trans[OF _ assms], rule float_divl)
-lemma bounded_divr: assumes "x \<le> real a / real b" shows "x \<le> real (float_divr p a b)" by (rule order_trans[OF assms _], rule float_divr)
-lemma bounded_num: shows "real (Float 5 1) = 10" and "real (Float 0 0) = 0" and "real (Float 1 0) = 1" and "real (Float (number_of n) 0) = (number_of n)"
-                     and "0 * pow2 e = real (Float 0 e)" and "1 * pow2 e = real (Float 1 e)" and "number_of m * pow2 e = real (Float (number_of m) e)"
-                     and "real (Float (number_of A) (int B)) = (number_of A) * 2^B"
-                     and "real (Float 1 (int B)) = 2^B"
-                     and "real (Float (number_of A) (- int B)) = (number_of A) / 2^B"
-                     and "real (Float 1 (- int B)) = 1 / 2^B"
-  by (auto simp add: real_of_float_simp pow2_def real_divide_def)
-
-lemmas bounded_by_equations = bounded_by_Cons bounded_by_Nil float_power bounded_divl bounded_divr bounded_num HOL.simp_thms
-lemmas interpret_inequality_equations = interpret_inequality.simps interpret_floatarith.simps interpret_floatarith_num
+lemmas interpret_form_equations = interpret_form.simps interpret_floatarith.simps interpret_floatarith_num
   interpret_floatarith_divide interpret_floatarith_diff interpret_floatarith_tan interpret_floatarith_powr interpret_floatarith_log
   interpret_floatarith_sin
 
@@ -2543,9 +3197,11 @@
 @{code_datatype float = Float}
 @{code_datatype floatarith = Add | Minus | Mult | Inverse | Cos | Arctan
                            | Abs | Max | Min | Pi | Sqrt | Exp | Ln | Power | Atom | Num }
-@{code_datatype inequality = Less | LessEqual }
-
-val approx_inequality = @{code approx_inequality}
+@{code_datatype form = Bound | Assign | Less | LessEqual | AtLeastAtMost}
+
+val approx_form = @{code approx_form}
+val approx_tse_form = @{code approx_tse_form}
+val approx' = @{code approx'}
 
 end
 *}
@@ -2565,101 +3221,212 @@
         "Float'_Arith.Exp" and "Float'_Arith.Ln" and "Float'_Arith.Power/ (_,/ _)" and
         "Float'_Arith.Atom" and "Float'_Arith.Num")
 
-code_type inequality (Eval "Float'_Arith.inequality")
-code_const Less and LessEqual (Eval "Float'_Arith.Less/ (_,/ _)" and "Float'_Arith.LessEqual/ (_,/ _)")
-
-code_const approx_inequality (Eval "Float'_Arith.approx'_inequality")
+code_type form (Eval "Float'_Arith.form")
+code_const Bound and Assign and Less and LessEqual and AtLeastAtMost
+      (Eval "Float'_Arith.Bound/ (_,/ _,/ _,/ _)" and "Float'_Arith.Assign/ (_,/ _,/ _)" and
+            "Float'_Arith.Less/ (_,/ _)" and "Float'_Arith.LessEqual/ (_,/ _)"  and
+            "Float'_Arith.AtLeastAtMost/ (_,/ _,/ _)")
+
+code_const approx_form (Eval "Float'_Arith.approx'_form")
+code_const approx_tse_form (Eval "Float'_Arith.approx'_tse'_form")
+code_const approx' (Eval "Float'_Arith.approx'")
 
 ML {*
-  val ineq_equations = PureThy.get_thms @{theory} "interpret_inequality_equations";
-  val bounded_by_equations = PureThy.get_thms @{theory} "bounded_by_equations";
-  val bounded_by_simpset = (HOL_basic_ss addsimps bounded_by_equations)
-
-  fun reify_ineq ctxt i = (fn st =>
+  fun reorder_bounds_tac i prems =
     let
-      val to = HOLogic.dest_Trueprop (Logic.strip_imp_concl (List.nth (prems_of st, i - 1)))
-    in (Reflection.genreify_tac ctxt ineq_equations (SOME to) i) st
-    end)
-
-  fun rule_ineq ctxt prec i thm = let
-    fun conv_num typ = HOLogic.dest_number #> snd #> HOLogic.mk_number typ
-    val to_natc = conv_num @{typ "nat"} #> Thm.cterm_of (ProofContext.theory_of ctxt)
-    val to_nat = conv_num @{typ "nat"}
-    val to_int = conv_num @{typ "int"}
-    fun int_to_float A = @{term "Float"} $ to_int A $ @{term "0 :: int"}
-
-    val prec' = to_nat prec
-
-    fun bot_float (Const (@{const_name "times"}, _) $ mantisse $ (Const (@{const_name "pow2"}, _) $ exp))
-                   = @{term "Float"} $ to_int mantisse $ to_int exp
-      | bot_float (Const (@{const_name "divide"}, _) $ mantisse $ (@{term "power 2 :: nat \<Rightarrow> real"} $ exp))
-                   = @{term "Float"} $ to_int mantisse $ (@{term "uminus :: int \<Rightarrow> int"} $ (@{term "int :: nat \<Rightarrow> int"} $ to_nat exp))
-      | bot_float (Const (@{const_name "times"}, _) $ mantisse $ (@{term "power 2 :: nat \<Rightarrow> real"} $ exp))
-                   = @{term "Float"} $ to_int mantisse $ (@{term "int :: nat \<Rightarrow> int"} $ to_nat exp)
-      | bot_float (Const (@{const_name "divide"}, _) $ A $ (@{term "power 10 :: nat \<Rightarrow> real"} $ exp))
-                   = @{term "float_divl"} $ prec' $ (int_to_float A) $ (@{term "power (Float 5 1)"} $ to_nat exp)
-      | bot_float (Const (@{const_name "divide"}, _) $ A $ B)
-                   = @{term "float_divl"} $ prec' $ int_to_float A $ int_to_float B
-      | bot_float (@{term "power 2 :: nat \<Rightarrow> real"} $ exp)
-                   = @{term "Float 1"} $ (@{term "int :: nat \<Rightarrow> int"} $ to_nat exp)
-      | bot_float A = int_to_float A
-
-    fun top_float (Const (@{const_name "times"}, _) $ mantisse $ (Const (@{const_name "pow2"}, _) $ exp))
-                   = @{term "Float"} $ to_int mantisse $ to_int exp
-      | top_float (Const (@{const_name "divide"}, _) $ mantisse $ (@{term "power 2 :: nat \<Rightarrow> real"} $ exp))
-                   = @{term "Float"} $ to_int mantisse $ (@{term "uminus :: int \<Rightarrow> int"} $ (@{term "int :: nat \<Rightarrow> int"} $ to_nat exp))
-      | top_float (Const (@{const_name "times"}, _) $ mantisse $ (@{term "power 2 :: nat \<Rightarrow> real"} $ exp))
-                   = @{term "Float"} $ to_int mantisse $ (@{term "int :: nat \<Rightarrow> int"} $ to_nat exp)
-      | top_float (Const (@{const_name "divide"}, _) $ A $ (@{term "power 10 :: nat \<Rightarrow> real"} $ exp))
-                   = @{term "float_divr"} $ prec' $ (int_to_float A) $ (@{term "power (Float 5 1)"} $ to_nat exp)
-      | top_float (Const (@{const_name "divide"}, _) $ A $ B)
-                   = @{term "float_divr"} $ prec' $ int_to_float A $ int_to_float B
-      | top_float (@{term "power 2 :: nat \<Rightarrow> real"} $ exp)
-                   = @{term "Float 1"} $ (@{term "int :: nat \<Rightarrow> int"} $ to_nat exp)
-      | top_float A = int_to_float A
-
-    val goal' : term = List.nth (prems_of thm, i - 1)
-
-    fun lift_bnd (t as (Const (@{const_name "op &"}, _) $
-                        (Const (@{const_name "less_eq"}, _) $
-                         bottom $ (Free (name, _))) $
-                        (Const (@{const_name "less_eq"}, _) $ _ $ top)))
-         = ((name, HOLogic.mk_prod (bot_float bottom, top_float top))
-            handle TERM (txt, ts) => raise TERM ("Invalid bound number format: " ^
-                                  (Syntax.string_of_term ctxt t), [t]))
-      | lift_bnd t = raise TERM ("Premisse needs format '<num> <= <var> & <var> <= <num>', but found " ^
-                                 (Syntax.string_of_term ctxt t), [t])
-    val bound_eqs = map (HOLogic.dest_Trueprop #> lift_bnd)  (Logic.strip_imp_prems goal')
-
-    fun lift_var (Free (varname, _)) = (case AList.lookup (op =) bound_eqs varname of
-                                          SOME bound => bound
-                                        | NONE => raise TERM ("No bound equations found for " ^ varname, []))
-      | lift_var t = raise TERM ("Can not convert expression " ^
-                                 (Syntax.string_of_term ctxt t), [t])
-
-    val _ $ vs = HOLogic.dest_Trueprop (Logic.strip_imp_concl goal')
-
-    val bs = (HOLogic.dest_list #> map lift_var #> HOLogic.mk_list @{typ "float * float"}) vs
-    val map = [(@{cpat "?prec::nat"}, to_natc prec),
-               (@{cpat "?bs::(float * float) list"}, Thm.cterm_of (ProofContext.theory_of ctxt) bs)]
-  in rtac (Thm.instantiate ([], map) @{thm "approx_inequality"}) i thm end
-
-  val eval_tac = CSUBGOAL (fn (ct, i) => rtac (eval_oracle ct) i)
-
+      fun variable_of_bound (Const ("Trueprop", _) $
+                             (Const (@{const_name "op :"}, _) $
+                              Free (name, _) $ _)) = name
+        | variable_of_bound (Const ("Trueprop", _) $
+                             (Const ("op =", _) $
+                              Free (name, _) $ _)) = name
+        | variable_of_bound t = raise TERM ("variable_of_bound", [t])
+
+      val variable_bounds
+        = map (` (variable_of_bound o prop_of)) prems
+
+      fun add_deps (name, bnds)
+        = Graph.add_deps_acyclic
+            (name, remove (op =) name (Term.add_free_names (prop_of bnds) []))
+      val order = Graph.empty
+                  |> fold Graph.new_node variable_bounds
+                  |> fold add_deps variable_bounds
+                  |> Graph.topological_order |> rev
+                  |> map_filter (AList.lookup (op =) variable_bounds)
+
+      fun prepend_prem th tac
+        = tac THEN rtac (th RSN (2, @{thm mp})) i
+    in
+      fold prepend_prem order all_tac
+    end
+
+  (* Should be in HOL.thy ? *)
   fun gen_eval_tac conv ctxt = CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt)
                                THEN' rtac TrueI
 
+  val form_equations = PureThy.get_thms @{theory} "interpret_form_equations";
+
+  fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let
+      fun lookup_splitting (Free (name, typ))
+        = case AList.lookup (op =) splitting name
+          of SOME s => HOLogic.mk_number @{typ nat} s
+           | NONE => @{term "0 :: nat"}
+      val vs = nth (prems_of st) (i - 1)
+               |> Logic.strip_imp_concl
+               |> HOLogic.dest_Trueprop
+               |> Term.strip_comb |> snd |> List.last
+               |> HOLogic.dest_list
+      val p = prec
+              |> HOLogic.mk_number @{typ nat}
+              |> Thm.cterm_of (ProofContext.theory_of ctxt)
+    in case taylor
+    of NONE => let
+         val n = vs |> length
+                 |> HOLogic.mk_number @{typ nat}
+                 |> Thm.cterm_of (ProofContext.theory_of ctxt)
+         val s = vs
+                 |> map lookup_splitting
+                 |> HOLogic.mk_list @{typ nat}
+                 |> Thm.cterm_of (ProofContext.theory_of ctxt)
+       in
+         (rtac (Thm.instantiate ([], [(@{cpat "?n::nat"}, n),
+                                     (@{cpat "?prec::nat"}, p),
+                                     (@{cpat "?ss::nat list"}, s)])
+              @{thm "approx_form"}) i
+          THEN simp_tac @{simpset} i) st
+       end
+
+     | SOME t => if length vs <> 1 then raise (TERM ("More than one variable used for taylor series expansion", [prop_of st]))
+       else let
+         val t = t
+              |> HOLogic.mk_number @{typ nat}
+              |> Thm.cterm_of (ProofContext.theory_of ctxt)
+         val s = vs |> map lookup_splitting |> hd
+              |> Thm.cterm_of (ProofContext.theory_of ctxt)
+       in
+         rtac (Thm.instantiate ([], [(@{cpat "?s::nat"}, s),
+                                     (@{cpat "?t::nat"}, t),
+                                     (@{cpat "?prec::nat"}, p)])
+              @{thm "approx_tse_form"}) i st
+       end
+    end
+
+  (* copied from Tools/induct.ML should probably in args.ML *)
+  val free = Args.context -- Args.term >> (fn (_, Free (n, t)) => n | (ctxt, t) =>
+    error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
+
 *}
 
+lemma intervalE: "a \<le> x \<and> x \<le> b \<Longrightarrow> \<lbrakk> x \<in> { a .. b } \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+  by auto
+
+lemma meta_eqE: "x \<equiv> a \<Longrightarrow> \<lbrakk> x = a \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+  by auto
+
 method_setup approximation = {*
-  Args.term >>
-  (fn prec => fn ctxt =>
+  Scan.lift (OuterParse.nat)
+  --
+  Scan.optional (Scan.lift (Args.$$$ "splitting" |-- Args.colon)
+    |-- OuterParse.and_list' (free --| Scan.lift (Args.$$$ "=") -- Scan.lift OuterParse.nat)) []
+  --
+  Scan.option (Scan.lift (Args.$$$ "taylor" |-- Args.colon)
+    |-- (free |-- Scan.lift (Args.$$$ "=") |-- Scan.lift OuterParse.nat))
+  >>
+  (fn ((prec, splitting), taylor) => fn ctxt =>
     SIMPLE_METHOD' (fn i =>
-     (DETERM (reify_ineq ctxt i)
-      THEN rule_ineq ctxt prec i
-      THEN Simplifier.asm_full_simp_tac bounded_by_simpset i
-      THEN (TRY (filter_prems_tac (fn t => false) i))
-      THEN (gen_eval_tac eval_oracle ctxt) i)))
-*} "real number approximation"
+      REPEAT (FIRST' [etac @{thm intervalE},
+                      etac @{thm meta_eqE},
+                      rtac @{thm impI}] i)
+      THEN METAHYPS (reorder_bounds_tac i) i
+      THEN TRY (filter_prems_tac (K false) i)
+      THEN DETERM (Reflection.genreify_tac ctxt form_equations NONE i)
+      THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
+      THEN gen_eval_tac eval_oracle ctxt i))
+ *} "real number approximation"
+
+ML {*
+  fun dest_interpret (@{const "interpret_floatarith"} $ b $ xs) = (b, xs)
+  | dest_interpret t = raise TERM ("dest_interpret", [t])
+
+  fun mk_approx' prec t = (@{const "approx'"}
+                         $ HOLogic.mk_number @{typ nat} prec
+                         $ t $ @{term "[] :: (float * float) list"})
+
+  fun dest_result (Const (@{const_name "Some"}, _) $
+                   ((Const (@{const_name "Pair"}, _)) $
+                    (@{const "Float"} $ lm $ le) $
+                    (@{const "Float"} $ um $ ue)))
+                   = SOME ((snd (HOLogic.dest_number lm), snd (HOLogic.dest_number le)),
+                           (snd (HOLogic.dest_number um), snd (HOLogic.dest_number ue)))
+    | dest_result (Const (@{const_name "None"}, _)) = NONE
+    | dest_result t = raise TERM ("dest_result", [t])
+
+  fun float2_float10 prec round_down (m, e) = (
+    let
+      val (m, e) = (if e < 0 then (m,e) else (m * Integer.pow e 2, 0))
+
+      fun frac c p 0 digits cnt = (digits, cnt, 0)
+        | frac c 0 r digits cnt = (digits, cnt, r)
+        | frac c p r digits cnt = (let
+          val (d, r) = Integer.div_mod (r * 10) (Integer.pow (~e) 2)
+        in frac (c orelse d <> 0) (if d <> 0 orelse c then p - 1 else p) r
+                (digits * 10 + d) (cnt + 1)
+        end)
+
+      val sgn = Int.sign m
+      val m = abs m
+
+      val round_down = (sgn = 1 andalso round_down) orelse
+                       (sgn = ~1 andalso not round_down)
+
+      val (x, r) = Integer.div_mod m (Integer.pow (~e) 2)
+
+      val p = ((if x = 0 then prec else prec - (IntInf.log2 x + 1)) * 3) div 10 + 1
+
+      val (digits, e10, r) = if p > 0 then frac (x <> 0) p r 0 0 else (0,0,0)
+
+      val digits = if round_down orelse r = 0 then digits else digits + 1
+
+    in (sgn * (digits + x * (Integer.pow e10 10)), ~e10)
+    end)
+
+  fun mk_result prec (SOME (l, u)) = (let
+      fun mk_float10 rnd x = (let val (m, e) = float2_float10 prec rnd x
+                         in if e = 0 then HOLogic.mk_number @{typ real} m
+                       else if e = 1 then @{term "divide :: real \<Rightarrow> real \<Rightarrow> real"} $
+                                          HOLogic.mk_number @{typ real} m $
+                                          @{term "10"}
+                                     else @{term "divide :: real \<Rightarrow> real \<Rightarrow> real"} $
+                                          HOLogic.mk_number @{typ real} m $
+                                          (@{term "power 10 :: nat \<Rightarrow> real"} $
+                                           HOLogic.mk_number @{typ nat} (~e)) end)
+      in @{term "atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set"} $
+         mk_float10 true l $ mk_float10 false u end)
+    | mk_result prec NONE = @{term "UNIV :: real set"}
+
+
+  fun realify t = let
+      val t = Logic.varify t
+      val m = map (fn (name, sort) => (name, @{typ real})) (Term.add_tvars t [])
+      val t = Term.subst_TVars m t
+    in t end
+
+  fun approx prec ctxt t = let val t = realify t in
+          t
+       |> Reflection.genreif ctxt form_equations
+       |> prop_of
+       |> HOLogic.dest_Trueprop
+       |> HOLogic.dest_eq |> snd
+       |> dest_interpret |> fst
+       |> mk_approx' prec
+       |> Codegen.eval_term @{theory}
+       |> dest_result
+       |> mk_result prec
+    end
+*}
+
+setup {*
+  Value.add_evaluator ("approximate", approx 30)
+*}
 
 end
--- a/src/HOL/Decision_Procs/ex/Approximation_Ex.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Decision_Procs/ex/Approximation_Ex.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -8,13 +8,28 @@
 
 Here are some examples how to use the approximation method.
 
-The parameter passed to the method specifies the precision used by the computations, it is specified
-as number of bits to calculate. When a variable is used it needs to be bounded by an interval. This
-interval is specified as a conjunction of the lower and upper bound. It must have the form
-@{text "\<lbrakk> l\<^isub>1 \<le> x\<^isub>1 \<and> x\<^isub>1 \<le> u\<^isub>1 ; \<dots> ; l\<^isub>n \<le> x\<^isub>n \<and> x\<^isub>n \<le> u\<^isub>n \<rbrakk> \<Longrightarrow> F"} where @{term F} is the formula, and
-@{text "x\<^isub>1, \<dots>, x\<^isub>n"} are the variables. The lower bounds @{text "l\<^isub>1, \<dots>, l\<^isub>n"} and the upper bounds
-@{text "u\<^isub>1, \<dots>, u\<^isub>n"} must either be integer numerals, floating point numbers or of the form
-@{term "m * pow2 e"} to specify a exact floating point value.
+The approximation method has the following syntax:
+
+approximate "prec" (splitting: "x" = "depth" and "y" = "depth" ...)? (taylor: "z" = "derivates")?
+
+Here "prec" specifies the precision used in all computations, it is specified as
+number of bits to calculate. In the proposition to prove an arbitrary amount of
+variables can be used, but each one need to be bounded by an upper and lower
+bound.
+
+To specify the bounds either @{term "l\<^isub>1 \<le> x \<and> x \<le> u\<^isub>1"},
+@{term "x \<in> { l\<^isub>1 .. u\<^isub>1 }"} or @{term "x = bnd"} can be used. Where the
+bound specification are again arithmetic formulas containing variables. They can
+be connected using either meta level or HOL equivalence.
+
+To use interval splitting add for each variable whos interval should be splitted
+to the "splitting:" parameter. The parameter specifies how often each interval
+should be divided, e.g. when x = 16 is specified, there will be @{term "65536 = 2^16"}
+intervals to be calculated.
+
+To use taylor series expansion specify the variable to derive. You need to
+specify the amount of derivations to compute. When using taylor series expansion
+only one variable can be used.
 
 *}
 
@@ -40,6 +55,9 @@
 lemma "0.5 \<le> x \<and> x \<le> 4.5 \<Longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
   by (approximation 10)
 
+lemma "x \<in> {0.5 .. 4.5} \<longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
+  by (approximation 10)
+
 lemma "0.49 \<le> x \<and> x \<le> 4.49 \<Longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
   by (approximation 20)
 
@@ -49,5 +67,12 @@
 lemma "3.2 \<le> x \<and> x \<le> 6.2 \<Longrightarrow> sin x \<le> 0"
   by (approximation 9)
 
+lemma
+  defines "g \<equiv> 9.80665" and "v \<equiv> 128.61" and "d \<equiv> pi / 180"
+  shows "g / v * tan (35 * d) \<in> { 3 * d .. 3.1 * d }"
+  using assms by (approximation 80)
+
+lemma "\<phi> \<in> { 0 .. 1 :: real } \<longrightarrow> \<phi> ^ 2 \<le> \<phi>"
+  by (approximation 30 splitting: \<phi>=1 taylor: \<phi> = 3)
+
 end
-
--- a/src/HOL/Deriv.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Deriv.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -115,12 +115,16 @@
 
 text{*Differentiation of finite sum*}
 
+lemma DERIV_setsum:
+  assumes "finite S"
+  and "\<And> n. n \<in> S \<Longrightarrow> DERIV (%x. f x n) x :> (f' x n)"
+  shows "DERIV (%x. setsum (f x) S) x :> setsum (f' x) S"
+  using assms by induct (auto intro!: DERIV_add)
+
 lemma DERIV_sumr [rule_format (no_asm)]:
      "(\<forall>r. m \<le> r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x))
       --> DERIV (%x. \<Sum>n=m..<n::nat. f n x :: real) x :> (\<Sum>r=m..<n. f' r x)"
-apply (induct "n")
-apply (auto intro: DERIV_add)
-done
+  by (auto intro: DERIV_setsum)
 
 text{*Alternative definition for differentiability*}
 
@@ -216,7 +220,6 @@
   shows "DERIV (\<lambda>x. f x ^ n) x :> of_nat n * (D * f x ^ (n - Suc 0))"
 by (cases "n", simp, simp add: DERIV_power_Suc f del: power_Suc)
 
-
 text {* Caratheodory formulation of derivative at a point *}
 
 lemma CARAT_DERIV:
@@ -308,6 +311,30 @@
 lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E"
 by auto
 
+text {* DERIV_intros *}
+
+ML{*
+structure DerivIntros =
+  NamedThmsFun(val name = "DERIV_intros"
+               val description = "DERIV introduction rules");
+*}
+setup DerivIntros.setup
+
+lemma DERIV_cong: "\<lbrakk> DERIV f x :> X ; X = Y \<rbrakk> \<Longrightarrow> DERIV f x :> Y"
+  by simp
+
+declare
+  DERIV_const[THEN DERIV_cong, DERIV_intros]
+  DERIV_ident[THEN DERIV_cong, DERIV_intros]
+  DERIV_add[THEN DERIV_cong, DERIV_intros]
+  DERIV_minus[THEN DERIV_cong, DERIV_intros]
+  DERIV_mult[THEN DERIV_cong, DERIV_intros]
+  DERIV_diff[THEN DERIV_cong, DERIV_intros]
+  DERIV_inverse'[THEN DERIV_cong, DERIV_intros]
+  DERIV_divide[THEN DERIV_cong, DERIV_intros]
+  DERIV_power[where 'a=real, THEN DERIV_cong,
+              unfolded real_of_nat_def[symmetric], DERIV_intros]
+  DERIV_setsum[THEN DERIV_cong, DERIV_intros]
 
 subsection {* Differentiability predicate *}
 
--- a/src/HOL/GCD.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/GCD.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -1,6 +1,6 @@
 (*  Title:      GCD.thy
     Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
-                Thomas M. Rasmussen, Jeremy Avigad
+                Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow
 
 
 This file deals with the functions gcd and lcm, and properties of
@@ -20,6 +20,7 @@
 the congruence relations on the integers. The notion was extended to
 the natural numbers by Chiaeb.
 
+Tobias Nipkow cleaned up a lot.
 *)
 
 
@@ -180,8 +181,17 @@
 lemma int_gcd_neg2 [simp]: "gcd (x::int) (-y) = gcd x y"
   by (simp add: gcd_int_def)
 
+lemma abs_gcd_int[simp]: "abs(gcd (x::int) y) = gcd x y"
+by(simp add: gcd_int_def)
+
 lemma int_gcd_abs: "gcd (x::int) y = gcd (abs x) (abs y)"
-  by (simp add: gcd_int_def)
+by (simp add: gcd_int_def)
+
+lemma gcd_abs1_int[simp]: "gcd (abs x) (y::int) = gcd x y"
+by (metis abs_idempotent int_gcd_abs)
+
+lemma gcd_abs2_int[simp]: "gcd x (abs y::int) = gcd x y"
+by (metis abs_idempotent int_gcd_abs)
 
 lemma int_gcd_cases:
   fixes x :: int and y
@@ -204,6 +214,15 @@
 lemma int_lcm_abs: "lcm (x::int) y = lcm (abs x) (abs y)"
   by (simp add: lcm_int_def)
 
+lemma abs_lcm_int [simp]: "abs (lcm i j::int) = lcm i j"
+by(simp add:lcm_int_def)
+
+lemma lcm_abs1_int[simp]: "lcm (abs x) (y::int) = lcm x y"
+by (metis abs_idempotent lcm_int_def)
+
+lemma lcm_abs2_int[simp]: "lcm x (abs y::int) = lcm x y"
+by (metis abs_idempotent lcm_int_def)
+
 lemma int_lcm_cases:
   fixes x :: int and y
   assumes "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (lcm x y)"
@@ -247,11 +266,11 @@
 lemma int_gcd_1 [simp]: "gcd (m::int) 1 = 1"
   by (simp add: gcd_int_def)
 
-lemma nat_gcd_self [simp]: "gcd (x::nat) x = x"
-  by simp
+lemma nat_gcd_idem: "gcd (x::nat) x = x"
+by simp
 
-lemma int_gcd_def [simp]: "gcd (x::int) x = abs x"
-  by (subst int_gcd_abs, auto simp add: gcd_int_def)
+lemma int_gcd_idem: "gcd (x::int) x = abs x"
+by (auto simp add: gcd_int_def)
 
 declare gcd_nat.simps [simp del]
 
@@ -267,21 +286,11 @@
   apply (blast dest: dvd_mod_imp_dvd)
 done
 
-thm nat_gcd_dvd1 [transferred]
-
 lemma int_gcd_dvd1 [iff]: "gcd (x::int) y dvd x"
-  apply (subst int_gcd_abs)
-  apply (rule dvd_trans)
-  apply (rule nat_gcd_dvd1 [transferred])
-  apply auto
-done
+by (metis gcd_int_def int_dvd_iff nat_gcd_dvd1)
 
 lemma int_gcd_dvd2 [iff]: "gcd (x::int) y dvd y"
-  apply (subst int_gcd_abs)
-  apply (rule dvd_trans)
-  apply (rule nat_gcd_dvd2 [transferred])
-  apply auto
-done
+by (metis gcd_int_def int_dvd_iff nat_gcd_dvd2)
 
 lemma dvd_gcd_D1_nat: "k dvd gcd m n \<Longrightarrow> (k::nat) dvd m"
 by(metis nat_gcd_dvd1 dvd_trans)
@@ -308,16 +317,14 @@
   by (rule zdvd_imp_le, auto)
 
 lemma nat_gcd_greatest: "(k::nat) dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
-  by (induct m n rule: nat_gcd_induct) (simp_all add: nat_gcd_non_0 dvd_mod)
+by (induct m n rule: nat_gcd_induct) (simp_all add: nat_gcd_non_0 dvd_mod)
 
 lemma int_gcd_greatest:
-  assumes "(k::int) dvd m" and "k dvd n"
-  shows "k dvd gcd m n"
-
+  "(k::int) dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
   apply (subst int_gcd_abs)
   apply (subst abs_dvd_iff [symmetric])
   apply (rule nat_gcd_greatest [transferred])
-  using prems apply auto
+  apply auto
 done
 
 lemma nat_gcd_greatest_iff [iff]: "(k dvd gcd (m::nat) n) =
@@ -364,15 +371,6 @@
 
 lemmas int_gcd_ac = int_gcd_assoc int_gcd_commute int_gcd_left_commute
 
-lemma nat_gcd_1_left [simp]: "gcd (1::nat) m = 1"
-  by (subst nat_gcd_commute, simp)
-
-lemma nat_gcd_Suc_0_left [simp]: "gcd (Suc 0) m = Suc 0"
-  by (subst nat_gcd_commute, simp add: One_nat_def)
-
-lemma int_gcd_1_left [simp]: "gcd (1::int) m = 1"
-  by (subst int_gcd_commute, simp)
-
 lemma nat_gcd_unique: "(d::nat) dvd a \<and> d dvd b \<and>
     (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
   apply auto
@@ -395,6 +393,19 @@
   apply (auto intro: int_gcd_greatest)
 done
 
+lemma gcd_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> gcd x y = x"
+by (metis dvd.eq_iff nat_gcd_unique)
+
+lemma gcd_proj2_if_dvd_nat [simp]: "(y::nat) dvd x \<Longrightarrow> gcd x y = y"
+by (metis dvd.eq_iff nat_gcd_unique)
+
+lemma gcd_proj1_if_dvd_int[simp]: "x dvd y \<Longrightarrow> gcd (x::int) y = abs x"
+by (metis abs_dvd_iff abs_eq_0 int_gcd_0_left int_gcd_abs int_gcd_unique)
+
+lemma gcd_proj2_if_dvd_int[simp]: "y dvd x \<Longrightarrow> gcd (x::int) y = abs y"
+by (metis gcd_proj1_if_dvd_int int_gcd_commute)
+
+
 text {*
   \medskip Multiplication laws
 *}
@@ -409,17 +420,11 @@
 
 lemma int_gcd_mult_distrib: "abs (k::int) * gcd m n = gcd (k * m) (k * n)"
   apply (subst (1 2) int_gcd_abs)
-  apply (simp add: abs_mult)
+  apply (subst (1 2) abs_mult)
   apply (rule nat_gcd_mult_distrib [transferred])
   apply auto
 done
 
-lemma nat_gcd_mult [simp]: "gcd (k::nat) (k * n) = k"
-  by (rule nat_gcd_mult_distrib [of k 1 n, simplified, symmetric])
-
-lemma int_gcd_mult [simp]: "gcd (k::int) (k * n) = abs k"
-  by (rule int_gcd_mult_distrib [of k 1 n, simplified, symmetric])
-
 lemma nat_coprime_dvd_mult: "coprime (k::nat) n \<Longrightarrow> k dvd m * n \<Longrightarrow> k dvd m"
   apply (insert nat_gcd_mult_distrib [of m k n])
   apply simp
@@ -428,16 +433,14 @@
   done
 
 lemma int_coprime_dvd_mult:
-  assumes "coprime (k::int) n" and "k dvd m * n"
-  shows "k dvd m"
-
-  using prems
-  apply (subst abs_dvd_iff [symmetric])
-  apply (subst dvd_abs_iff [symmetric])
-  apply (subst (asm) int_gcd_abs)
-  apply (rule nat_coprime_dvd_mult [transferred])
-  apply auto
-  apply (subst abs_mult [symmetric], auto)
+  "coprime (k::int) n \<Longrightarrow> k dvd m * n \<Longrightarrow> k dvd m"
+apply (subst abs_dvd_iff [symmetric])
+apply (subst dvd_abs_iff [symmetric])
+apply (subst (asm) int_gcd_abs)
+apply (rule nat_coprime_dvd_mult [transferred])
+    prefer 4 apply assumption
+   apply auto
+apply (subst abs_mult [symmetric], auto)
 done
 
 lemma nat_coprime_dvd_mult_iff: "coprime (k::nat) n \<Longrightarrow>
@@ -458,14 +461,10 @@
 done
 
 lemma int_gcd_mult_cancel:
-  assumes "coprime (k::int) n"
-  shows "gcd (k * m) n = gcd m n"
-
-  using prems
-  apply (subst (1 2) int_gcd_abs)
-  apply (subst abs_mult)
-  apply (rule nat_gcd_mult_cancel [transferred])
-  apply (auto simp add: int_gcd_abs [symmetric])
+  "coprime (k::int) n \<Longrightarrow> gcd (k * m) n = gcd m n"
+apply (subst (1 2) int_gcd_abs)
+apply (subst abs_mult)
+apply (rule nat_gcd_mult_cancel [transferred], auto)
 done
 
 text {* \medskip Addition laws *}
@@ -517,42 +516,22 @@
 done
 
 lemma int_gcd_add1 [simp]: "gcd ((m::int) + n) n = gcd m n"
-  apply (case_tac "n = 0", force)
-  apply (subst (1 2) int_gcd_red)
-  apply auto
-done
+by (metis int_gcd_red mod_add_self1 zadd_commute)
 
 lemma int_gcd_add2 [simp]: "gcd m ((m::int) + n) = gcd m n"
-  apply (subst int_gcd_commute)
-  apply (subst add_commute)
-  apply (subst int_gcd_add1)
-  apply (subst int_gcd_commute)
-  apply (rule refl)
-done
+by (metis int_gcd_add1 int_gcd_commute zadd_commute)
 
 lemma nat_gcd_add_mult: "gcd (m::nat) (k * m + n) = gcd m n"
-  by (induct k, simp_all add: ring_simps)
+by (metis mod_mult_self3 nat_gcd_commute nat_gcd_red)
 
 lemma int_gcd_add_mult: "gcd (m::int) (k * m + n) = gcd m n"
-  apply (subgoal_tac "ALL s. ALL m. ALL n.
-      gcd m (int (s::nat) * m + n) = gcd m n")
-  apply (case_tac "k >= 0")
-  apply (drule_tac x = "nat k" in spec, force)
-  apply (subst (1 2) int_gcd_neg2 [symmetric])
-  apply (drule_tac x = "nat (- k)" in spec)
-  apply (drule_tac x = "m" in spec)
-  apply (drule_tac x = "-n" in spec)
-  apply auto
-  apply (rule nat_induct)
-  apply auto
-  apply (auto simp add: left_distrib)
-  apply (subst add_assoc)
-  apply simp
-done
+by (metis int_gcd_commute int_gcd_red mod_mult_self1 zadd_commute)
+
 
 (* to do: differences, and all variations of addition rules
     as simplification rules for nat and int *)
 
+(* FIXME remove iff *)
 lemma nat_gcd_dvd_prod [iff]: "gcd (m::nat) n dvd k * n"
   using mult_dvd_mono [of 1] by auto
 
@@ -596,7 +575,7 @@
 
 subsection {* Coprimality *}
 
-lemma nat_div_gcd_coprime [intro]:
+lemma nat_div_gcd_coprime:
   assumes nz: "(a::nat) \<noteq> 0 \<or> b \<noteq> 0"
   shows "coprime (a div gcd a b) (b div gcd a b)"
 proof -
@@ -620,16 +599,16 @@
   with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
 qed
 
-lemma int_div_gcd_coprime [intro]:
+lemma int_div_gcd_coprime:
   assumes nz: "(a::int) \<noteq> 0 \<or> b \<noteq> 0"
   shows "coprime (a div gcd a b) (b div gcd a b)"
-
-  apply (subst (1 2 3) int_gcd_abs)
-  apply (subst (1 2) abs_div)
-  apply auto
-  prefer 3
-  apply (rule nat_div_gcd_coprime [transferred])
-  using nz apply (auto simp add: int_gcd_abs [symmetric])
+apply (subst (1 2 3) int_gcd_abs)
+apply (subst (1 2) abs_div)
+  apply simp
+ apply simp
+apply(subst (1 2) abs_gcd_int)
+apply (rule nat_div_gcd_coprime [transferred])
+using nz apply (auto simp add: int_gcd_abs [symmetric])
 done
 
 lemma nat_coprime: "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
@@ -714,20 +693,20 @@
 qed
 
 lemma int_coprime_lmult:
-  assumes dab: "coprime (d::int) (a * b)" shows "coprime d a"
+  assumes "coprime (d::int) (a * b)" shows "coprime d a"
 proof -
   have "gcd d a dvd gcd d (a * b)"
     by (rule int_gcd_greatest, auto)
-  with dab show ?thesis
+  with assms show ?thesis
     by auto
 qed
 
 lemma nat_coprime_rmult:
-  assumes dab: "coprime (d::nat) (a * b)" shows "coprime d b"
+  assumes "coprime (d::nat) (a * b)" shows "coprime d b"
 proof -
   have "gcd d b dvd gcd d (a * b)"
     by (rule nat_gcd_greatest, auto intro: dvd_mult)
-  with dab show ?thesis
+  with assms show ?thesis
     by auto
 qed
 
@@ -794,7 +773,7 @@
   thus ?thesis by simp
   next assume "~(a = 0 & b = 0)"
   hence "coprime ((a div gcd a b)^n) ((b div gcd a b)^n)"
-    by auto
+    by (auto simp:nat_div_gcd_coprime)
   hence "gcd ((a div gcd a b)^n * (gcd a b)^n)
       ((b div gcd a b)^n * (gcd a b)^n) = (gcd a b)^n"
     apply (subst (1 2) mult_commute)
@@ -937,6 +916,7 @@
   ultimately show ?thesis by blast
 qed
 
+(* FIXME move to Divides(?) *)
 lemma nat_pow_divides_eq [simp]: "n ~= 0 \<Longrightarrow> ((a::nat)^n dvd b^n) = (a dvd b)"
   by (auto intro: nat_pow_divides_pow dvd_power_same)
 
@@ -1316,30 +1296,12 @@
 lemma int_lcm_0 [simp]: "lcm (m::int) 0 = 0"
   unfolding lcm_int_def by simp
 
-lemma nat_lcm_1 [simp]: "lcm (m::nat) 1 = m"
-  unfolding lcm_nat_def by simp
-
-lemma nat_lcm_Suc_0 [simp]: "lcm (m::nat) (Suc 0) = m"
-  unfolding lcm_nat_def by (simp add: One_nat_def)
-
-lemma int_lcm_1 [simp]: "lcm (m::int) 1 = abs m"
-  unfolding lcm_int_def by simp
-
 lemma nat_lcm_0_left [simp]: "lcm (0::nat) n = 0"
   unfolding lcm_nat_def by simp
 
 lemma int_lcm_0_left [simp]: "lcm (0::int) n = 0"
   unfolding lcm_int_def by simp
 
-lemma nat_lcm_1_left [simp]: "lcm (1::nat) m = m"
-  unfolding lcm_nat_def by simp
-
-lemma nat_lcm_Suc_0_left [simp]: "lcm (Suc 0) m = m"
-  unfolding lcm_nat_def by (simp add: One_nat_def)
-
-lemma int_lcm_1_left [simp]: "lcm (1::int) m = abs m"
-  unfolding lcm_int_def by simp
-
 lemma nat_lcm_commute: "lcm (m::nat) n = lcm n m"
   unfolding lcm_nat_def by (simp add: nat_gcd_commute ring_simps)
 
@@ -1348,34 +1310,14 @@
 
 
 lemma nat_lcm_pos:
-  assumes mpos: "(m::nat) > 0"
-  and npos: "n>0"
-  shows "lcm m n > 0"
-proof(rule ccontr, simp add: lcm_nat_def nat_gcd_zero)
-  assume h:"m*n div gcd m n = 0"
-  from mpos npos have "gcd m n \<noteq> 0" using nat_gcd_zero by simp
-  hence gcdp: "gcd m n > 0" by simp
-  with h
-  have "m*n < gcd m n"
-    by (cases "m * n < gcd m n")
-       (auto simp add: div_if[OF gcdp, where m="m*n"])
-  moreover
-  have "gcd m n dvd m" by simp
-  with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
-  with npos have t1:"gcd m n*n \<le> m*n" by simp
-  have "gcd m n \<le> gcd m n*n" using npos by simp
-  with t1 have "gcd m n \<le> m*n" by arith
-  ultimately show "False" by simp
-qed
+  "(m::nat) > 0 \<Longrightarrow> n>0 \<Longrightarrow> lcm m n > 0"
+by (metis gr0I mult_is_0 nat_prod_gcd_lcm)
 
 lemma int_lcm_pos:
-  assumes mneq0: "(m::int) ~= 0"
-  and npos: "n ~= 0"
-  shows "lcm m n > 0"
-
+  "(m::int) ~= 0 \<Longrightarrow> n ~= 0 \<Longrightarrow> lcm m n > 0"
   apply (subst int_lcm_abs)
   apply (rule nat_lcm_pos [transferred])
-  using prems apply auto
+  apply auto
 done
 
 lemma nat_dvd_pos:
@@ -1417,13 +1359,11 @@
 qed
 
 lemma int_lcm_least:
-  assumes "(m::int) dvd k" and "n dvd k"
-  shows "lcm m n dvd k"
-
-  apply (subst int_lcm_abs)
-  apply (rule dvd_trans)
-  apply (rule nat_lcm_least [transferred, of _ "abs k" _])
-  using prems apply auto
+  "(m::int) dvd k \<Longrightarrow> n dvd k \<Longrightarrow> lcm m n dvd k"
+apply (subst int_lcm_abs)
+apply (rule dvd_trans)
+apply (rule nat_lcm_least [transferred, of _ "abs k" _])
+apply auto
 done
 
 lemma nat_lcm_dvd1: "(m::nat) dvd lcm m n"
@@ -1481,23 +1421,23 @@
     (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
   by (auto intro: dvd_anti_sym [transferred] int_lcm_least)
 
-lemma nat_lcm_dvd_eq [simp]: "(x::nat) dvd y \<Longrightarrow> lcm x y = y"
+lemma lcm_proj2_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> lcm x y = y"
   apply (rule sym)
   apply (subst nat_lcm_unique [symmetric])
   apply auto
 done
 
-lemma int_lcm_dvd_eq [simp]: "0 <= y \<Longrightarrow> (x::int) dvd y \<Longrightarrow> lcm x y = y"
+lemma lcm_proj2_if_dvd_int [simp]: "(x::int) dvd y \<Longrightarrow> lcm x y = abs y"
   apply (rule sym)
   apply (subst int_lcm_unique [symmetric])
   apply auto
 done
 
-lemma nat_lcm_dvd_eq' [simp]: "(x::nat) dvd y \<Longrightarrow> lcm y x = y"
-  by (subst nat_lcm_commute, erule nat_lcm_dvd_eq)
+lemma lcm_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> lcm y x = y"
+by (subst nat_lcm_commute, erule lcm_proj2_if_dvd_nat)
 
-lemma int_lcm_dvd_eq' [simp]: "y >= 0 \<Longrightarrow> (x::int) dvd y \<Longrightarrow> lcm y x = y"
-  by (subst int_lcm_commute, erule (1) int_lcm_dvd_eq)
+lemma lcm_proj1_if_dvd_int [simp]: "(x::int) dvd y \<Longrightarrow> lcm y x = abs y"
+by (subst int_lcm_commute, erule lcm_proj2_if_dvd_int)
 
 
 lemma lcm_assoc_nat: "lcm (lcm n m) (p::nat) = lcm n (lcm m p)"
--- a/src/HOL/HOL.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/HOL.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -198,9 +198,6 @@
 axiomatization
   undefined :: 'a
 
-abbreviation (input)
-  "arbitrary \<equiv> undefined"
-
 
 subsubsection {* Generic classes and algebraic operations *}
 
--- a/src/HOL/HahnBanach/Bounds.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,82 +0,0 @@
-(*  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
--- a/src/HOL/HahnBanach/FunctionNorm.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,278 +0,0 @@
-(*  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
--- a/src/HOL/HahnBanach/FunctionOrder.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /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/HahnBanach/HahnBanach.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,509 +0,0 @@
-(*  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
--- a/src/HOL/HahnBanach/HahnBanachExtLemmas.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,280 +0,0 @@
-(*  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
--- a/src/HOL/HahnBanach/HahnBanachLemmas.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /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/HahnBanach/HahnBanachSupLemmas.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /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/HahnBanach/Linearform.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,60 +0,0 @@
-(*  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
--- a/src/HOL/HahnBanach/NormedSpace.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,117 +0,0 @@
-(*  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
--- a/src/HOL/HahnBanach/README.html	Wed Jun 24 16:28:30 2009 +0100
+++ /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/HahnBanach/ROOT.ML	Wed Jun 24 16:28:30 2009 +0100
+++ /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/HahnBanach/Subspace.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,513 +0,0 @@
-(*  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
--- a/src/HOL/HahnBanach/VectorSpace.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,419 +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 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
--- a/src/HOL/HahnBanach/ZornLemma.thy	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,57 +0,0 @@
-(*  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
--- a/src/HOL/HahnBanach/document/root.bib	Wed Jun 24 16:28:30 2009 +0100
+++ /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/HahnBanach/document/root.tex	Wed Jun 24 16:28:30 2009 +0100
+++ /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}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hahn_Banach/Bounds.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,82 @@
+(*  Title:      HOL/Hahn_Banach/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/Hahn_Banach/Function_Norm.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,278 @@
+(*  Title:      HOL/Hahn_Banach/Function_Norm.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The norm of a function *}
+
+theory Function_Norm
+imports Normed_Space Function_Order
+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/Hahn_Banach/Function_Order.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,141 @@
+(*  Title:      HOL/Hahn_Banach/Function_Order.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* An order on functions *}
+
+theory Function_Order
+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/Hahn_Banach/Hahn_Banach.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,509 @@
+(*  Title:      HOL/Hahn_Banach/Hahn_Banach.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The Hahn-Banach Theorem *}
+
+theory Hahn_Banach
+imports Hahn_Banach_Lemmas
+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 Hahn_Banach:
+  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-Hahn-Banach} 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_Hahn_Banach:
+  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 Hahn_Banach)
+    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_Hahn_Banach:
+  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_Hahn_Banach [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/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,280 @@
+(*  Title:      HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Extending non-maximal functions *}
+
+theory Hahn_Banach_Ext_Lemmas
+imports Function_Norm
+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/Hahn_Banach/Hahn_Banach_Lemmas.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,4 @@
+(*<*)
+theory Hahn_Banach_Lemmas imports Hahn_Banach_Sup_Lemmas Hahn_Banach_Ext_Lemmas begin
+end
+(*>*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,445 @@
+(*  Title:      HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The supremum w.r.t.~the function order *}
+
+theory Hahn_Banach_Sup_Lemmas
+imports Function_Norm Zorn_Lemma
+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_Hahn_Banach}
+  (see page \pageref{abs-Hahn_Banach}). \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/Hahn_Banach/Linearform.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,60 @@
+(*  Title:      HOL/Hahn_Banach/Linearform.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Linearforms *}
+
+theory Linearform
+imports Vector_Space
+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/Hahn_Banach/Normed_Space.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,117 @@
+(*  Title:      HOL/Hahn_Banach/Normed_Space.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Normed vector spaces *}
+
+theory Normed_Space
+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 "Normed_Space.norm E norm" ..
+    with subset show "Normed_Space.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/Hahn_Banach/README.html	Wed Jul 01 16:19:44 2009 +0100
@@ -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/Hahn_Banach/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/Hahn_Banach/ROOT.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,7 @@
+(*  Title:      HOL/Hahn_Banach/ROOT.ML
+    Author:     Gertrud Bauer, TU Munich
+
+The Hahn-Banach theorem for real vector spaces (Isabelle/Isar).
+*)
+
+use_thy "Hahn_Banach";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hahn_Banach/Subspace.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,513 @@
+(*  Title:      HOL/Hahn_Banach/Subspace.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Subspaces *}
+
+theory Subspace
+imports Vector_Space
+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/Hahn_Banach/Vector_Space.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,418 @@
+(*  Title:      HOL/Hahn_Banach/Vector_Space.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Vector spaces *}
+
+theory Vector_Space
+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/Hahn_Banach/Zorn_Lemma.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,57 @@
+(*  Title:      HOL/Hahn_Banach/Zorn_Lemma.thy
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Zorn's Lemma *}
+
+theory Zorn_Lemma
+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/Hahn_Banach/document/root.bib	Wed Jul 01 16:19:44 2009 +0100
@@ -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/Hahn_Banach/document/root.tex	Wed Jul 01 16:19:44 2009 +0100
@@ -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{Vector_Space}
+\input{Subspace}
+\input{Normed_Space}
+\input{Linearform}
+\input{Function_Order}
+\input{Function_Norm}
+\input{Zorn_Lemma}
+
+\clearpage
+\part {Lemmas for the Proof}
+
+\input{Hahn_Banach_Sup_Lemmas}
+\input{Hahn_Banach_Ext_Lemmas}
+\input{Hahn_Banach_Lemmas}
+
+\clearpage
+\part {The Main Proof}
+
+\input{Hahn_Banach}
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
--- a/src/HOL/Imperative_HOL/Array.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Imperative_HOL/Array.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -1,5 +1,4 @@
-(*  Title:      HOL/Library/Array.thy
-    ID:         $Id$
+(*  Title:      HOL/Imperative_HOL/Array.thy
     Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
 *)
 
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -306,67 +306,75 @@
 code_const "Heap_Monad.Fail" (OCaml "Failure")
 code_const "Heap_Monad.raise_exc" (OCaml "!(fun/ ()/ ->/ raise/ _)")
 
-setup {* let
-  open Code_Thingol;
+setup {*
+
+let
 
-  fun lookup naming = the o Code_Thingol.lookup_const naming;
+open Code_Thingol;
+
+fun imp_program naming =
 
-  fun imp_monad_bind'' bind' return' unit' ts =
-    let
-      val dummy_name = "";
-      val dummy_type = ITyVar dummy_name;
-      val dummy_case_term = IVar dummy_name;
-      (*assumption: dummy values are not relevant for serialization*)
-      val unitt = IConst (unit', (([], []), []));
-      fun dest_abs ((v, ty) `|=> t, _) = ((v, ty), t)
-        | dest_abs (t, ty) =
-            let
-              val vs = Code_Thingol.fold_varnames cons t [];
-              val v = Name.variant vs "x";
-              val ty' = (hd o fst o Code_Thingol.unfold_fun) ty;
-            in ((v, ty'), t `$ IVar v) end;
-      fun force (t as IConst (c, _) `$ t') = if c = return'
-            then t' else t `$ unitt
-        | force t = t `$ unitt;
-      fun tr_bind' [(t1, _), (t2, ty2)] =
-        let
-          val ((v, ty), t) = dest_abs (t2, ty2);
-        in ICase (((force t1, ty), [(IVar v, tr_bind'' t)]), dummy_case_term) end
-      and tr_bind'' t = case Code_Thingol.unfold_app t
-           of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if c = bind'
-                then tr_bind' [(x1, ty1), (x2, ty2)]
-                else force t
-            | _ => force t;
-    in (dummy_name, dummy_type) `|=> ICase (((IVar dummy_name, dummy_type),
-      [(unitt, tr_bind' ts)]), dummy_case_term) end
-  and imp_monad_bind' bind' return' unit' (const as (c, (_, tys))) ts = if c = bind' then case (ts, tys)
-     of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)]
-      | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)] `$ t3
-      | (ts, _) => imp_monad_bind bind' return' unit' (eta_expand 2 (const, ts))
-    else IConst const `$$ map (imp_monad_bind bind' return' unit') ts
-  and imp_monad_bind bind' return' unit' (IConst const) = imp_monad_bind' bind' return' unit' const []
-    | imp_monad_bind bind' return' unit' (t as IVar _) = t
-    | imp_monad_bind bind' return' unit' (t as _ `$ _) = (case unfold_app t
-       of (IConst const, ts) => imp_monad_bind' bind' return' unit' const ts
-        | (t, ts) => imp_monad_bind bind' return' unit' t `$$ map (imp_monad_bind bind' return' unit') ts)
-    | imp_monad_bind bind' return' unit' (v_ty `|=> t) = v_ty `|=> imp_monad_bind bind' return' unit' t
-    | imp_monad_bind bind' return' unit' (ICase (((t, ty), pats), t0)) = ICase
-        (((imp_monad_bind bind' return' unit' t, ty), (map o pairself) (imp_monad_bind bind' return' unit') pats), imp_monad_bind bind' return' unit' t0);
+  let
+    fun is_const c = case lookup_const naming c
+     of SOME c' => (fn c'' => c' = c'')
+      | NONE => K false;
+    val is_bindM = is_const @{const_name bindM};
+    val is_return = is_const @{const_name return};
+    val dummy_name = "";
+    val dummy_type = ITyVar dummy_name;
+    val dummy_case_term = IVar NONE;
+    (*assumption: dummy values are not relevant for serialization*)
+    val unitt = case lookup_const naming @{const_name Unity}
+     of SOME unit' => IConst (unit', (([], []), []))
+      | NONE => error ("Must include " ^ @{const_name Unity} ^ " in generated constants.");
+    fun dest_abs ((v, ty) `|=> t, _) = ((v, ty), t)
+      | dest_abs (t, ty) =
+          let
+            val vs = fold_varnames cons t [];
+            val v = Name.variant vs "x";
+            val ty' = (hd o fst o unfold_fun) ty;
+          in ((SOME v, ty'), t `$ IVar (SOME v)) end;
+    fun force (t as IConst (c, _) `$ t') = if is_return c
+          then t' else t `$ unitt
+      | force t = t `$ unitt;
+    fun tr_bind' [(t1, _), (t2, ty2)] =
+      let
+        val ((v, ty), t) = dest_abs (t2, ty2);
+      in ICase (((force t1, ty), [(IVar v, tr_bind'' t)]), dummy_case_term) end
+    and tr_bind'' t = case unfold_app t
+         of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if is_bindM c
+              then tr_bind' [(x1, ty1), (x2, ty2)]
+              else force t
+          | _ => force t;
+    fun imp_monad_bind'' ts = (SOME dummy_name, dummy_type) `|=> ICase (((IVar (SOME dummy_name), dummy_type),
+      [(unitt, tr_bind' ts)]), dummy_case_term)
+    and imp_monad_bind' (const as (c, (_, tys))) ts = if is_bindM c then case (ts, tys)
+       of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)]
+        | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)] `$ t3
+        | (ts, _) => imp_monad_bind (eta_expand 2 (const, ts))
+      else IConst const `$$ map imp_monad_bind ts
+    and imp_monad_bind (IConst const) = imp_monad_bind' const []
+      | imp_monad_bind (t as IVar _) = t
+      | imp_monad_bind (t as _ `$ _) = (case unfold_app t
+         of (IConst const, ts) => imp_monad_bind' const ts
+          | (t, ts) => imp_monad_bind t `$$ map imp_monad_bind ts)
+      | imp_monad_bind (v_ty `|=> t) = v_ty `|=> imp_monad_bind t
+      | imp_monad_bind (ICase (((t, ty), pats), t0)) = ICase
+          (((imp_monad_bind t, ty),
+            (map o pairself) imp_monad_bind pats),
+              imp_monad_bind t0);
 
-  fun imp_program naming = (Graph.map_nodes o map_terms_stmt)
-    (imp_monad_bind (lookup naming @{const_name bindM})
-      (lookup naming @{const_name return})
-      (lookup naming @{const_name Unity}));
+  in (Graph.map_nodes o map_terms_stmt) imp_monad_bind end;
 
 in
 
-  Code_Target.extend_target ("SML_imp", ("SML", imp_program))
-  #> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program))
+Code_Target.extend_target ("SML_imp", ("SML", imp_program))
+#> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program))
 
 end
+
 *}
 
-
 code_reserved OCaml Failure raise
 
 
--- a/src/HOL/Imperative_HOL/Imperative_HOL_ex.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Imperative_HOL/Imperative_HOL_ex.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -1,8 +1,9 @@
 (*  Title:      HOL/Imperative_HOL/Imperative_HOL_ex.thy
-    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
+    Author:     John Matthews, Galois Connections;
+                Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
 *)
 
-header {* Mmonadic imperative HOL with examples *}
+header {* Monadic imperative HOL with examples *}
 
 theory Imperative_HOL_ex
 imports Imperative_HOL "ex/Imperative_Quicksort"
--- a/src/HOL/IsaMakefile	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/IsaMakefile	Wed Jul 01 16:19:44 2009 +0100
@@ -16,7 +16,7 @@
   HOL-Bali \
   HOL-Decision_Procs \
   HOL-Extraction \
-  HOL-HahnBanach \
+  HOL-Hahn_Banach \
   HOL-Hoare \
   HOL-HoareParallel \
   HOL-Import \
@@ -319,7 +319,7 @@
   Library/Abstract_Rat.thy \
   Library/BigO.thy Library/ContNotDenum.thy Library/Efficient_Nat.thy	\
   Library/Euclidean_Space.thy Library/Sum_Of_Squares.thy Library/positivstellensatz.ML	\
-  Library/Convex_Euclidean_Space.thy \
+  Library/Fset.thy  Library/Convex_Euclidean_Space.thy \
   Library/sum_of_squares.ML Library/Glbs.thy Library/normarith.ML \
   Library/Executable_Set.thy Library/Infinite_Set.thy			\
   Library/FuncSet.thy Library/Permutations.thy Library/Determinants.thy\
@@ -329,7 +329,7 @@
   Library/Fundamental_Theorem_Algebra.thy \
   Library/Inner_Product.thy Library/Lattice_Syntax.thy \
   Library/Legacy_GCD.thy \
-  Library/Library.thy Library/List_Prefix.thy Library/State_Monad.thy	\
+  Library/Library.thy Library/List_Prefix.thy Library/List_Set.thy Library/State_Monad.thy	\
   Library/Nat_Int_Bij.thy Library/Multiset.thy Library/Permutation.thy	\
   Library/Primes.thy Library/Pocklington.thy Library/Quotient.thy	\
   Library/Quicksort.thy Library/Nat_Infinity.thy Library/Word.thy	\
@@ -360,21 +360,19 @@
 	@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
 
 
-## HOL-HahnBanach
+## HOL-Hahn_Banach
 
-HOL-HahnBanach: HOL $(LOG)/HOL-HahnBanach.gz
+HOL-Hahn_Banach: HOL $(LOG)/HOL-Hahn_Banach.gz
 
-$(LOG)/HOL-HahnBanach.gz: $(OUT)/HOL			\
-  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
+$(LOG)/HOL-Hahn_Banach.gz: $(OUT)/HOL Hahn_Banach/Bounds.thy		\
+  Hahn_Banach/Function_Norm.thy Hahn_Banach/Function_Order.thy		\
+  Hahn_Banach/Hahn_Banach.thy Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy	\
+  Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy Hahn_Banach/Linearform.thy	\
+  Hahn_Banach/Normed_Space.thy Hahn_Banach/README.html			\
+  Hahn_Banach/ROOT.ML Hahn_Banach/Subspace.thy				\
+  Hahn_Banach/Vector_Space.thy Hahn_Banach/Zorn_Lemma.thy		\
+  Hahn_Banach/document/root.bib Hahn_Banach/document/root.tex
+	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Hahn_Banach
 
 
 ## HOL-Subst
@@ -1138,7 +1136,7 @@
 		$(LOG)/HOL-MicroJava.gz $(LOG)/HOL-NanoJava.gz		\
 		$(LOG)/HOL-Nominal-Examples.gz $(LOG)/HOL-IOA.gz	\
 		$(LOG)/HOL-Lattice $(LOG)/HOL-Matrix			\
-		$(LOG)/HOL-HahnBanach.gz $(LOG)/HOL-SET-Protocol.gz	\
+		$(LOG)/HOL-Hahn_Banach.gz $(LOG)/HOL-SET-Protocol.gz	\
 		$(LOG)/TLA-Inc.gz $(LOG)/TLA-Buffer.gz			\
 		$(LOG)/TLA-Memory.gz $(LOG)/HOL-Library.gz		\
 		$(LOG)/HOL-Unix.gz $(OUT)/HOL-Word $(LOG)/HOL-Word.gz	\
--- a/src/HOL/Library/Executable_Set.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Library/Executable_Set.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -5,249 +5,43 @@
 header {* Implementation of finite sets by lists *}
 
 theory Executable_Set
-imports Main
+imports Main Fset
 begin
 
-subsection {* Definitional rewrites *}
+subsection {* Derived set operations *}
+
+declare member [code] 
 
 definition subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
   "subset = op \<le>"
 
 declare subset_def [symmetric, code unfold]
 
-lemma [code]: "subset A B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
-  unfolding subset_def subset_eq ..
-
-definition is_empty :: "'a set \<Rightarrow> bool" where
-  "is_empty A \<longleftrightarrow> A = {}"
+lemma [code]:
+  "subset A B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
+  by (simp add: subset_def subset_eq)
 
 definition eq_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
   [code del]: "eq_set = op ="
 
-lemma [code]: "eq_set A B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
-  unfolding eq_set_def by auto
-
 (* FIXME allow for Stefan's code generator:
 declare set_eq_subset[code unfold]
 *)
 
 lemma [code]:
-  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. x = a)"
-  unfolding bex_triv_one_point1 ..
-
-definition filter_set :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
-  "filter_set P xs = {x\<in>xs. P x}"
-
-declare filter_set_def[symmetric, code unfold] 
-
-
-subsection {* Operations on lists *}
-
-subsubsection {* Basic definitions *}
-
-definition
-  flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
-  "flip f a b = f b a"
-
-definition
-  member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
-  "member xs x \<longleftrightarrow> x \<in> set xs"
-
-definition
-  insertl :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
-  "insertl x xs = (if member xs x then xs else x#xs)"
-
-lemma [code target: List]: "member [] y \<longleftrightarrow> False"
-  and [code target: List]: "member (x#xs) y \<longleftrightarrow> y = x \<or> member xs y"
-  unfolding member_def by (induct xs) simp_all
-
-fun
-  drop_first :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
-  "drop_first f [] = []"
-| "drop_first f (x#xs) = (if f x then xs else x # drop_first f xs)"
-declare drop_first.simps [code del]
-declare drop_first.simps [code target: List]
+  "eq_set A B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
+  by (simp add: eq_set_def set_eq)
 
-declare remove1.simps [code del]
-lemma [code target: List]:
-  "remove1 x xs = (if member xs x then drop_first (\<lambda>y. y = x) xs else xs)"
-proof (cases "member xs x")
-  case False thus ?thesis unfolding member_def by (induct xs) auto
-next
-  case True
-  have "remove1 x xs = drop_first (\<lambda>y. y = x) xs" by (induct xs) simp_all
-  with True show ?thesis by simp
-qed
-
-lemma member_nil [simp]:
-  "member [] = (\<lambda>x. False)"
-proof (rule ext)
-  fix x
-  show "member [] x = False" unfolding member_def by simp
-qed
+declare inter [code]
 
-lemma member_insertl [simp]:
-  "x \<in> set (insertl x xs)"
-  unfolding insertl_def member_def mem_iff by simp
-
-lemma insertl_member [simp]:
-  fixes xs x
-  assumes member: "member xs x"
-  shows "insertl x xs = xs"
-  using member unfolding insertl_def by simp
-
-lemma insertl_not_member [simp]:
-  fixes xs x
-  assumes member: "\<not> (member xs x)"
-  shows "insertl x xs = x # xs"
-  using member unfolding insertl_def by simp
-
-lemma foldr_remove1_empty [simp]:
-  "foldr remove1 xs [] = []"
-  by (induct xs) simp_all
+declare Inter_image_eq [symmetric, code]
+declare Union_image_eq [symmetric, code]
 
 
-subsubsection {* Derived definitions *}
-
-function unionl :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "unionl [] ys = ys"
-| "unionl xs ys = foldr insertl xs ys"
-by pat_completeness auto
-termination by lexicographic_order
-
-lemmas unionl_eq = unionl.simps(2)
-
-function intersect :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "intersect [] ys = []"
-| "intersect xs [] = []"
-| "intersect xs ys = filter (member xs) ys"
-by pat_completeness auto
-termination by lexicographic_order
-
-lemmas intersect_eq = intersect.simps(3)
-
-function subtract :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "subtract [] ys = ys"
-| "subtract xs [] = []"
-| "subtract xs ys = foldr remove1 xs ys"
-by pat_completeness auto
-termination by lexicographic_order
-
-lemmas subtract_eq = subtract.simps(3)
-
-function map_distinct :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"
-where
-  "map_distinct f [] = []"
-| "map_distinct f xs = foldr (insertl o f) xs []"
-by pat_completeness auto
-termination by lexicographic_order
-
-lemmas map_distinct_eq = map_distinct.simps(2)
-
-function unions :: "'a list list \<Rightarrow> 'a list"
-where
-  "unions [] = []"
-| "unions xs = foldr unionl xs []"
-by pat_completeness auto
-termination by lexicographic_order
-
-lemmas unions_eq = unions.simps(2)
-
-consts intersects :: "'a list list \<Rightarrow> 'a list"
-primrec
-  "intersects (x#xs) = foldr intersect xs x"
-
-definition
-  map_union :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
-  "map_union xs f = unions (map f xs)"
-
-definition
-  map_inter :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
-  "map_inter xs f = intersects (map f xs)"
-
-
-subsection {* Isomorphism proofs *}
-
-lemma iso_member:
-  "member xs x \<longleftrightarrow> x \<in> set xs"
-  unfolding member_def mem_iff ..
+subsection {* Rewrites for primitive operations *}
 
-lemma iso_insert:
-  "set (insertl x xs) = insert x (set xs)"
-  unfolding insertl_def iso_member by (simp add: insert_absorb)
-
-lemma iso_remove1:
-  assumes distnct: "distinct xs"
-  shows "set (remove1 x xs) = set xs - {x}"
-  using distnct set_remove1_eq by auto
-
-lemma iso_union:
-  "set (unionl xs ys) = set xs \<union> set ys"
-  unfolding unionl_eq
-  by (induct xs arbitrary: ys) (simp_all add: iso_insert)
-
-lemma iso_intersect:
-  "set (intersect xs ys) = set xs \<inter> set ys"
-  unfolding intersect_eq Int_def by (simp add: Int_def iso_member) auto
-
-definition
-  subtract' :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
-  "subtract' = flip subtract"
-
-lemma iso_subtract:
-  fixes ys
-  assumes distnct: "distinct ys"
-  shows "set (subtract' ys xs) = set ys - set xs"
-    and "distinct (subtract' ys xs)"
-  unfolding subtract'_def flip_def subtract_eq
-  using distnct by (induct xs arbitrary: ys) auto
-
-lemma iso_map_distinct:
-  "set (map_distinct f xs) = image f (set xs)"
-  unfolding map_distinct_eq by (induct xs) (simp_all add: iso_insert)
+declare List_Set.project_def [symmetric, code unfold]
 
-lemma iso_unions:
-  "set (unions xss) = \<Union> set (map set xss)"
-  unfolding unions_eq
-proof (induct xss)
-  case Nil show ?case by simp
-next
-  case (Cons xs xss) thus ?case by (induct xs) (simp_all add: iso_insert)
-qed
-
-lemma iso_intersects:
-  "set (intersects (xs#xss)) = \<Inter> set (map set (xs#xss))"
-  by (induct xss) (simp_all add: Int_def iso_member, auto)
-
-lemma iso_UNION:
-  "set (map_union xs f) = UNION (set xs) (set o f)"
-  unfolding map_union_def iso_unions by simp
-
-lemma iso_INTER:
-  "set (map_inter (x#xs) f) = INTER (set (x#xs)) (set o f)"
-  unfolding map_inter_def iso_intersects by (induct xs) (simp_all add: iso_member, auto)
-
-definition
-  Blall :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
-  "Blall = flip list_all"
-definition
-  Blex :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
-  "Blex = flip list_ex"
-
-lemma iso_Ball:
-  "Blall xs f = Ball (set xs) f"
-  unfolding Blall_def flip_def by (induct xs) simp_all
-
-lemma iso_Bex:
-  "Blex xs f = Bex (set xs) f"
-  unfolding Blex_def flip_def by (induct xs) simp_all
-
-lemma iso_filter:
-  "set (filter P xs) = filter_set P (set xs)"
-  unfolding filter_set_def by (induct xs) auto
 
 subsection {* code generator setup *}
 
@@ -257,23 +51,33 @@
 nonfix subset;
 *}
 
-subsubsection {* const serializations *}
+definition flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
+  "flip f a b = f b a"
+
+types_code
+  fset ("(_/ \<module>fset)")
+attach {*
+datatype 'a fset = Set of 'a list;
+*}
+
+consts_code
+  Set ("\<module>Set")
 
 consts_code
-  "Set.empty" ("{*[]*}")
-  insert ("{*insertl*}")
-  is_empty ("{*null*}")
-  "op \<union>" ("{*unionl*}")
-  "op \<inter>" ("{*intersect*}")
-  "op - \<Colon> 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" ("{* flip subtract *}")
-  image ("{*map_distinct*}")
-  Union ("{*unions*}")
-  Inter ("{*intersects*}")
-  UNION ("{*map_union*}")
-  INTER ("{*map_inter*}")
-  Ball ("{*Blall*}")
-  Bex ("{*Blex*}")
-  filter_set ("{*filter*}")
-  fold ("{* foldl o flip *}")
+  "Set.empty"         ("{*Fset.empty*}")
+  "List_Set.is_empty" ("{*Fset.is_empty*}")
+  "Set.insert"        ("{*Fset.insert*}")
+  "List_Set.remove"   ("{*Fset.remove*}")
+  "Set.image"         ("{*Fset.map*}")
+  "List_Set.project"  ("{*Fset.filter*}")
+  "Ball"              ("{*flip Fset.forall*}")
+  "Bex"               ("{*flip Fset.exists*}")
+  "op \<union>"              ("{*Fset.union*}")
+  "op \<inter>"              ("{*Fset.inter*}")
+  "op - \<Colon> 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" ("{*flip Fset.subtract*}")
+  "Set.Union"         ("{*Fset.Union*}")
+  "Set.Inter"         ("{*Fset.Inter*}")
+  card                ("{*Fset.card*}")
+  fold                ("{*foldl o flip*}")
 
-end
+end
\ No newline at end of file
--- a/src/HOL/Library/Fin_Fun.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Library/Fin_Fun.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -383,7 +383,7 @@
 subsection {* Default value for FinFuns *}
 
 definition finfun_default_aux :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b"
-where [code del]: "finfun_default_aux f = (if finite (UNIV :: 'a set) then arbitrary else THE b. finite {a. f a \<noteq> b})"
+where [code del]: "finfun_default_aux f = (if finite (UNIV :: 'a set) then undefined else THE b. finite {a. f a \<noteq> b})"
 
 lemma finfun_default_aux_infinite:
   fixes f :: "'a \<Rightarrow> 'b"
@@ -453,7 +453,7 @@
 lemma finite_finfun_default: "finite {a. Rep_finfun f a \<noteq> finfun_default f}"
 unfolding finfun_default_def by(simp add: finite_finfun_default_aux)
 
-lemma finfun_default_const: "finfun_default ((\<lambda>\<^isup>f b) :: 'a \<Rightarrow>\<^isub>f 'b) = (if finite (UNIV :: 'a set) then arbitrary else b)"
+lemma finfun_default_const: "finfun_default ((\<lambda>\<^isup>f b) :: 'a \<Rightarrow>\<^isub>f 'b) = (if finite (UNIV :: 'a set) then undefined else b)"
 apply(auto simp add: finfun_default_def finfun_const_def finfun_default_aux_infinite)
 apply(simp add: finfun_default_aux_def)
 done
@@ -790,10 +790,10 @@
   ultimately show ?thesis by(simp add: finfun_rec_def)
 next
   case True
-  hence default: "finfun_default ((\<lambda>\<^isup>f c) :: 'a \<Rightarrow>\<^isub>f 'b) = arbitrary" by(simp add: finfun_default_const)
-  let ?the = "\<lambda>g :: 'a \<rightharpoonup> 'b. (\<lambda>\<^isup>f c) = Abs_finfun (map_default arbitrary g) \<and> finite (dom g) \<and> arbitrary \<notin> ran g"
+  hence default: "finfun_default ((\<lambda>\<^isup>f c) :: 'a \<Rightarrow>\<^isub>f 'b) = undefined" by(simp add: finfun_default_const)
+  let ?the = "\<lambda>g :: 'a \<rightharpoonup> 'b. (\<lambda>\<^isup>f c) = Abs_finfun (map_default undefined g) \<and> finite (dom g) \<and> undefined \<notin> ran g"
   show ?thesis
-  proof(cases "c = arbitrary")
+  proof(cases "c = undefined")
     case True
     have the: "The ?the = empty"
     proof
@@ -801,10 +801,10 @@
     next
       fix g'
       assume "?the g'"
-      hence fg: "(\<lambda>\<^isup>f c) = Abs_finfun (map_default arbitrary g')"
-        and fin: "finite (dom g')" and g: "arbitrary \<notin> ran g'" by simp_all
-      from fin have "map_default arbitrary g' \<in> finfun" by(rule map_default_in_finfun)
-      with fg have "map_default arbitrary g' = (\<lambda>a. c)"
+      hence fg: "(\<lambda>\<^isup>f c) = Abs_finfun (map_default undefined g')"
+        and fin: "finite (dom g')" and g: "undefined \<notin> ran g'" by simp_all
+      from fin have "map_default undefined g' \<in> finfun" by(rule map_default_in_finfun)
+      with fg have "map_default undefined g' = (\<lambda>a. c)"
         by(auto simp add: finfun_const_def intro: Abs_finfun_inject[THEN iffD1])
       with True show "g' = empty"
         by -(rule map_default_inject(2)[OF _ fin g], auto)
@@ -820,10 +820,10 @@
     next
       fix g' :: "'a \<rightharpoonup> 'b"
       assume "?the g'"
-      hence fg: "(\<lambda>\<^isup>f c) = Abs_finfun (map_default arbitrary g')"
-        and fin: "finite (dom g')" and g: "arbitrary \<notin> ran g'" by simp_all
-      from fin have "map_default arbitrary g' \<in> finfun" by(rule map_default_in_finfun)
-      with fg have "map_default arbitrary g' = (\<lambda>a. c)"
+      hence fg: "(\<lambda>\<^isup>f c) = Abs_finfun (map_default undefined g')"
+        and fin: "finite (dom g')" and g: "undefined \<notin> ran g'" by simp_all
+      from fin have "map_default undefined g' \<in> finfun" by(rule map_default_in_finfun)
+      with fg have "map_default undefined g' = (\<lambda>a. c)"
         by(auto simp add: finfun_const_def intro: Abs_finfun_inject[THEN iffD1])
       with True False show "g' = (\<lambda>a::'a. Some c)"
         by -(rule map_default_inject(2)[OF _ fin g], auto simp add: dom_def ran_def map_default_def_raw)
--- a/src/HOL/Library/Float.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Library/Float.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -59,6 +59,12 @@
    "real (Float -1 0) = -1" and "real (Float (number_of n) 0) = number_of n"
   by auto
 
+lemma float_number_of[simp]: "real (number_of x :: float) = number_of x"
+  by (simp only:number_of_float_def Float_num[unfolded number_of_is_id])
+
+lemma float_number_of_int[simp]: "real (Float n 0) = real n"
+  by (simp add: Float_num[unfolded number_of_is_id] real_of_float_simp pow2_def)
+
 lemma pow2_0[simp]: "pow2 0 = 1" by simp
 lemma pow2_1[simp]: "pow2 1 = 2" by simp
 lemma pow2_neg: "pow2 x = inverse (pow2 (-x))" by simp
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Fset.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,240 @@
+
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Executable finite sets *}
+
+theory Fset
+imports List_Set
+begin
+
+lemma foldl_apply_inv:
+  assumes "\<And>x. g (h x) = x"
+  shows "foldl f (g s) xs = g (foldl (\<lambda>s x. h (f (g s) x)) s xs)"
+  by (rule sym, induct xs arbitrary: s) (simp_all add: assms)
+
+declare mem_def [simp]
+
+
+subsection {* Lifting *}
+
+datatype 'a fset = Fset "'a set"
+
+primrec member :: "'a fset \<Rightarrow> 'a set" where
+  "member (Fset A) = A"
+
+lemma Fset_member [simp]:
+  "Fset (member A) = A"
+  by (cases A) simp
+
+definition Set :: "'a list \<Rightarrow> 'a fset" where
+  "Set xs = Fset (set xs)"
+
+lemma member_Set [simp]:
+  "member (Set xs) = set xs"
+  by (simp add: Set_def)
+
+code_datatype Set
+
+
+subsection {* Basic operations *}
+
+definition is_empty :: "'a fset \<Rightarrow> bool" where
+  [simp]: "is_empty A \<longleftrightarrow> List_Set.is_empty (member A)"
+
+lemma is_empty_Set [code]:
+  "is_empty (Set xs) \<longleftrightarrow> null xs"
+  by (simp add: is_empty_set)
+
+definition empty :: "'a fset" where
+  [simp]: "empty = Fset {}"
+
+lemma empty_Set [code]:
+  "empty = Set []"
+  by (simp add: Set_def)
+
+definition insert :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+  [simp]: "insert x A = Fset (Set.insert x (member A))"
+
+lemma insert_Set [code]:
+  "insert x (Set xs) = Set (List_Set.insert x xs)"
+  by (simp add: Set_def insert_set)
+
+definition remove :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+  [simp]: "remove x A = Fset (List_Set.remove x (member A))"
+
+lemma remove_Set [code]:
+  "remove x (Set xs) = Set (remove_all x xs)"
+  by (simp add: Set_def remove_set)
+
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" where
+  [simp]: "map f A = Fset (image f (member A))"
+
+lemma map_Set [code]:
+  "map f (Set xs) = Set (remdups (List.map f xs))"
+  by (simp add: Set_def)
+
+definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+  [simp]: "filter P A = Fset (List_Set.project P (member A))"
+
+lemma filter_Set [code]:
+  "filter P (Set xs) = Set (List.filter P xs)"
+  by (simp add: Set_def project_set)
+
+definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> bool" where
+  [simp]: "forall P A \<longleftrightarrow> Ball (member A) P"
+
+lemma forall_Set [code]:
+  "forall P (Set xs) \<longleftrightarrow> list_all P xs"
+  by (simp add: Set_def ball_set)
+
+definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> bool" where
+  [simp]: "exists P A \<longleftrightarrow> Bex (member A) P"
+
+lemma exists_Set [code]:
+  "exists P (Set xs) \<longleftrightarrow> list_ex P xs"
+  by (simp add: Set_def bex_set)
+
+definition card :: "'a fset \<Rightarrow> nat" where
+  [simp]: "card A = Finite_Set.card (member A)"
+
+lemma card_Set [code]:
+  "card (Set xs) = length (remdups xs)"
+proof -
+  have "Finite_Set.card (set (remdups xs)) = length (remdups xs)"
+    by (rule distinct_card) simp
+  then show ?thesis by (simp add: Set_def card_def)
+qed
+
+
+subsection {* Derived operations *}
+
+lemma member_exists [code]:
+  "member A y \<longleftrightarrow> exists (\<lambda>x. y = x) A"
+  by simp
+
+definition subfset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where
+  [simp]: "subfset_eq A B \<longleftrightarrow> member A \<subseteq> member B"
+
+lemma subfset_eq_forall [code]:
+  "subfset_eq A B \<longleftrightarrow> forall (\<lambda>x. member B x) A"
+  by (simp add: subset_eq)
+
+definition subfset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where
+  [simp]: "subfset A B \<longleftrightarrow> member A \<subset> member B"
+
+lemma subfset_subfset_eq [code]:
+  "subfset A B \<longleftrightarrow> subfset_eq A B \<and> \<not> subfset_eq B A"
+  by (simp add: subset)
+
+lemma eq_fset_subfset_eq [code]:
+  "eq_class.eq A B \<longleftrightarrow> subfset_eq A B \<and> subfset_eq B A"
+  by (cases A, cases B) (simp add: eq set_eq)
+
+definition inter :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+  [simp]: "inter A B = Fset (project (member A) (member B))"
+
+lemma inter_project [code]:
+  "inter A B = filter (member A) B"
+  by (simp add: inter)
+
+
+subsection {* Functorial operations *}
+
+definition union :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+  [simp]: "union A B = Fset (member A \<union> member B)"
+
+lemma union_insert [code]:
+  "union (Set xs) A = foldl (\<lambda>A x. insert x A) A xs"
+proof -
+  have "foldl (\<lambda>A x. Set.insert x A) (member A) xs =
+    member (foldl (\<lambda>A x. Fset (Set.insert x (member A))) A xs)"
+    by (rule foldl_apply_inv) simp
+  then show ?thesis by (simp add: union_set)
+qed
+
+definition subtract :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+  [simp]: "subtract A B = Fset (member B - member A)"
+
+lemma subtract_remove [code]:
+  "subtract (Set xs) A = foldl (\<lambda>A x. remove x A) A xs"
+proof -
+  have "foldl (\<lambda>A x. List_Set.remove x A) (member A) xs =
+    member (foldl (\<lambda>A x. Fset (List_Set.remove x (member A))) A xs)"
+    by (rule foldl_apply_inv) simp
+  then show ?thesis by (simp add: minus_set)
+qed
+
+definition Inter :: "'a fset fset \<Rightarrow> 'a fset" where
+  [simp]: "Inter A = Fset (Set.Inter (member ` member A))"
+
+lemma Inter_inter [code]:
+  "Inter (Set (A # As)) = foldl inter A As"
+proof -
+  note Inter_image_eq [simp del] set_map [simp del] set.simps [simp del]
+  have "foldl (op \<inter>) (member A) (List.map member As) = 
+    member (foldl (\<lambda>B A. Fset (member B \<inter> A)) A (List.map member As))"
+    by (rule foldl_apply_inv) simp
+  then show ?thesis
+    by (simp add: Inter_set image_set inter_def_raw inter foldl_map)
+qed
+
+definition Union :: "'a fset fset \<Rightarrow> 'a fset" where
+  [simp]: "Union A = Fset (Set.Union (member ` member A))"
+
+lemma Union_union [code]:
+  "Union (Set As) = foldl union empty As"
+proof -
+  note Union_image_eq [simp del] set_map [simp del]
+  have "foldl (op \<union>) (member empty) (List.map member As) = 
+    member (foldl (\<lambda>B A. Fset (member B \<union> A)) empty (List.map member As))"
+    by (rule foldl_apply_inv) simp
+  then show ?thesis
+    by (simp add: Union_set image_set union_def_raw foldl_map)
+qed
+
+
+subsection {* Misc operations *}
+
+lemma size_fset [code]:
+  "fset_size f A = 0"
+  "size A = 0"
+  by (cases A, simp) (cases A, simp)
+
+lemma fset_case_code [code]:
+  "fset_case f A = f (member A)"
+  by (cases A) simp
+
+lemma fset_rec_code [code]:
+  "fset_rec f A = f (member A)"
+  by (cases A) simp
+
+
+subsection {* Simplified simprules *}
+
+lemma is_empty_simp [simp]:
+  "is_empty A \<longleftrightarrow> member A = {}"
+  by (simp add: List_Set.is_empty_def)
+declare is_empty_def [simp del]
+
+lemma remove_simp [simp]:
+  "remove x A = Fset (member A - {x})"
+  by (simp add: List_Set.remove_def)
+declare remove_def [simp del]
+
+lemma filter_simp [simp]:
+  "filter P A = Fset {x \<in> member A. P x}"
+  by (simp add: List_Set.project_def)
+declare filter_def [simp del]
+
+lemma inter_simp [simp]:
+  "inter A B = Fset (member A \<inter> member B)"
+  by (simp add: inter)
+declare inter_def [simp del]
+
+declare mem_def [simp del]
+
+
+hide (open) const is_empty empty insert remove map filter forall exists card
+  subfset_eq subfset inter union subtract Inter Union
+
+end
--- a/src/HOL/Library/Library.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Library/Library.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -27,6 +27,7 @@
   Formal_Power_Series
   Fraction_Field
   FrechetDeriv
+  Fset
   FuncSet
   Fundamental_Theorem_Algebra
   Infinite_Set
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/List_Set.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,166 @@
+
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Relating (finite) sets and lists *}
+
+theory List_Set
+imports Main
+begin
+
+subsection {* Various additional list functions *}
+
+definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+  "insert x xs = (if x \<in> set xs then xs else x # xs)"
+
+definition remove_all :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+  "remove_all x xs = filter (Not o op = x) xs"
+
+
+subsection {* Various additional set functions *}
+
+definition is_empty :: "'a set \<Rightarrow> bool" where
+  "is_empty A \<longleftrightarrow> A = {}"
+
+definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+  "remove x A = A - {x}"
+
+lemma fun_left_comm_idem_remove:
+  "fun_left_comm_idem remove"
+proof -
+  have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
+  show ?thesis by (simp only: fun_left_comm_idem_remove rem)
+qed
+
+lemma minus_fold_remove:
+  assumes "finite A"
+  shows "B - A = fold remove B A"
+proof -
+  have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
+  show ?thesis by (simp only: rem assms minus_fold_remove)
+qed
+
+definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+  "project P A = {a\<in>A. P a}"
+
+
+subsection {* Basic set operations *}
+
+lemma is_empty_set:
+  "is_empty (set xs) \<longleftrightarrow> null xs"
+  by (simp add: is_empty_def null_empty)
+
+lemma ball_set:
+  "(\<forall>x\<in>set xs. P x) \<longleftrightarrow> list_all P xs"
+  by (rule list_ball_code)
+
+lemma bex_set:
+  "(\<exists>x\<in>set xs. P x) \<longleftrightarrow> list_ex P xs"
+  by (rule list_bex_code)
+
+lemma empty_set:
+  "{} = set []"
+  by simp
+
+lemma insert_set:
+  "Set.insert x (set xs) = set (insert x xs)"
+  by (auto simp add: insert_def)
+
+lemma remove_set:
+  "remove x (set xs) = set (remove_all x xs)"
+  by (auto simp add: remove_def remove_all_def)
+
+lemma image_set:
+  "image f (set xs) = set (map f xs)"
+  by simp
+
+lemma project_set:
+  "project P (set xs) = set (filter P xs)"
+  by (auto simp add: project_def)
+
+
+subsection {* Functorial set operations *}
+
+lemma union_set:
+  "set xs \<union> A = foldl (\<lambda>A x. Set.insert x A) A xs"
+proof -
+  interpret fun_left_comm_idem Set.insert
+    by (fact fun_left_comm_idem_insert)
+  show ?thesis by (simp add: union_fold_insert fold_set)
+qed
+
+lemma minus_set:
+  "A - set xs = foldl (\<lambda>A x. remove x A) A xs"
+proof -
+  interpret fun_left_comm_idem remove
+    by (fact fun_left_comm_idem_remove)
+  show ?thesis
+    by (simp add: minus_fold_remove [of _ A] fold_set)
+qed
+
+lemma Inter_set:
+  "Inter (set (A # As)) = foldl (op \<inter>) A As"
+proof -
+  have "finite (set (A # As))" by simp
+  moreover have "fold (op \<inter>) UNIV (set (A # As)) = foldl (\<lambda>y x. x \<inter> y) UNIV (A # As)"
+    by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
+  ultimately have "Inter (set (A # As)) = foldl (op \<inter>) UNIV (A # As)"
+    by (simp only: Inter_fold_inter Int_commute)
+  then show ?thesis by simp
+qed
+
+lemma Union_set:
+  "Union (set As) = foldl (op \<union>) {} As"
+proof -
+  have "fold (op \<union>) {} (set As) = foldl (\<lambda>y x. x \<union> y) {} As"
+    by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
+  then show ?thesis
+    by (simp only: Union_fold_union finite_set Un_commute)
+qed
+
+lemma INTER_set:
+  "INTER (set (A # As)) f = foldl (\<lambda>B A. f A \<inter> B) (f A) As"
+proof -
+  have "finite (set (A # As))" by simp
+  moreover have "fold (\<lambda>A. op \<inter> (f A)) UNIV (set (A # As)) = foldl (\<lambda>B A. f A \<inter> B) UNIV (A # As)"
+    by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
+  ultimately have "INTER (set (A # As)) f = foldl (\<lambda>B A. f A \<inter> B) UNIV (A # As)"
+    by (simp only: INTER_fold_inter) 
+  then show ?thesis by simp
+qed
+
+lemma UNION_set:
+  "UNION (set As) f = foldl (\<lambda>B A. f A \<union> B) {} As"
+proof -
+  have "fold (\<lambda>A. op \<union> (f A)) {} (set As) = foldl (\<lambda>B A. f A \<union> B) {} As"
+    by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
+  then show ?thesis
+    by (simp only: UNION_fold_union finite_set)
+qed
+
+
+subsection {* Derived set operations *}
+
+lemma member:
+  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
+  by simp
+
+lemma subset_eq:
+  "A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
+  by (fact subset_eq)
+
+lemma subset:
+  "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
+  by (fact less_le_not_le)
+
+lemma set_eq:
+  "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
+  by (fact eq_iff)
+
+lemma inter:
+  "A \<inter> B = project (\<lambda>x. x \<in> A) B"
+  by (auto simp add: project_def)
+
+
+hide (open) const insert
+
+end
\ No newline at end of file
--- a/src/HOL/Library/Poly_Deriv.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Library/Poly_Deriv.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -85,13 +85,7 @@
 by (rule lemma_DERIV_subst, rule DERIV_add, auto)
 
 lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x"
-apply (induct p)
-apply simp
-apply (simp add: pderiv_pCons)
-apply (rule lemma_DERIV_subst)
-apply (rule DERIV_add DERIV_mult DERIV_const DERIV_ident | assumption)+
-apply simp
-done
+  by (induct p, auto intro!: DERIV_intros simp add: pderiv_pCons)
 
 text{* Consequences of the derivative theorem above*}
 
--- a/src/HOL/Library/Topology_Euclidean_Space.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Library/Topology_Euclidean_Space.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -533,18 +533,18 @@
   fix x :: "'a ^ 'b"
   {
     fix e :: real assume "0 < e"
-    def a \<equiv> "x $ arbitrary"
+    def a \<equiv> "x $ undefined"
     have "a islimpt UNIV" by (rule islimpt_UNIV)
     with `0 < e` obtain b where "b \<noteq> a" and "dist b a < e"
       unfolding islimpt_approachable by auto
-    def y \<equiv> "Cart_lambda ((Cart_nth x)(arbitrary := b))"
+    def y \<equiv> "Cart_lambda ((Cart_nth x)(undefined := b))"
     from `b \<noteq> a` have "y \<noteq> x"
       unfolding a_def y_def by (simp add: Cart_eq)
     from `dist b a < e` have "dist y x < e"
       unfolding dist_vector_def a_def y_def
       apply simp
       apply (rule le_less_trans [OF setL2_le_setsum [OF zero_le_dist]])
-      apply (subst setsum_diff1' [where a=arbitrary], simp, simp, simp)
+      apply (subst setsum_diff1' [where a=undefined], simp, simp, simp)
       done
     from `y \<noteq> x` and `dist y x < e`
     have "\<exists>y\<in>UNIV. y \<noteq> x \<and> dist y x < e" by auto
@@ -2695,29 +2695,29 @@
 
 primrec helper_2::"(real \<Rightarrow> 'a::metric_space) \<Rightarrow> nat \<Rightarrow> 'a" where
   "helper_2 beyond 0 = beyond 0" |
-  "helper_2 beyond (Suc n) = beyond (dist arbitrary (helper_2 beyond n) + 1 )"
+  "helper_2 beyond (Suc n) = beyond (dist undefined (helper_2 beyond n) + 1 )"
 
 lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set"
   assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
   shows "bounded s"
 proof(rule ccontr)
   assume "\<not> bounded s"
-  then obtain beyond where "\<forall>a. beyond a \<in>s \<and> \<not> dist arbitrary (beyond a) \<le> a"
-    unfolding bounded_any_center [where a=arbitrary]
-    apply simp using choice[of "\<lambda>a x. x\<in>s \<and> \<not> dist arbitrary x \<le> a"] by auto
-  hence beyond:"\<And>a. beyond a \<in>s" "\<And>a. dist arbitrary (beyond a) > a"
+  then obtain beyond where "\<forall>a. beyond a \<in>s \<and> \<not> dist undefined (beyond a) \<le> a"
+    unfolding bounded_any_center [where a=undefined]
+    apply simp using choice[of "\<lambda>a x. x\<in>s \<and> \<not> dist undefined x \<le> a"] by auto
+  hence beyond:"\<And>a. beyond a \<in>s" "\<And>a. dist undefined (beyond a) > a"
     unfolding linorder_not_le by auto
   def x \<equiv> "helper_2 beyond"
 
   { fix m n ::nat assume "m<n"
-    hence "dist arbitrary (x m) + 1 < dist arbitrary (x n)"
+    hence "dist undefined (x m) + 1 < dist undefined (x n)"
     proof(induct n)
       case 0 thus ?case by auto
     next
       case (Suc n)
-      have *:"dist arbitrary (x n) + 1 < dist arbitrary (x (Suc n))"
+      have *:"dist undefined (x n) + 1 < dist undefined (x (Suc n))"
         unfolding x_def and helper_2.simps
-	using beyond(2)[of "dist arbitrary (helper_2 beyond n) + 1"] by auto
+	using beyond(2)[of "dist undefined (helper_2 beyond n) + 1"] by auto
       thus ?case proof(cases "m < n")
 	case True thus ?thesis using Suc and * by auto
       next
@@ -2729,12 +2729,12 @@
     have "1 < dist (x m) (x n)"
     proof(cases "m<n")
       case True
-      hence "1 < dist arbitrary (x n) - dist arbitrary (x m)" using *[of m n] by auto
-      thus ?thesis using dist_triangle [of arbitrary "x n" "x m"] by arith
+      hence "1 < dist undefined (x n) - dist undefined (x m)" using *[of m n] by auto
+      thus ?thesis using dist_triangle [of undefined "x n" "x m"] by arith
     next
       case False hence "n<m" using `m\<noteq>n` by auto
-      hence "1 < dist arbitrary (x m) - dist arbitrary (x n)" using *[of n m] by auto
-      thus ?thesis using dist_triangle2 [of arbitrary "x m" "x n"] by arith
+      hence "1 < dist undefined (x m) - dist undefined (x n)" using *[of n m] by auto
+      thus ?thesis using dist_triangle2 [of undefined "x m" "x n"] by arith
     qed  } note ** = this
   { fix a b assume "x a = x b" "a \<noteq> b"
     hence False using **[of a b] by auto  }
--- a/src/HOL/Library/reflection.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Library/reflection.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -10,6 +10,7 @@
   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
+  val genreif : Proof.context -> thm list -> term -> thm
 end;
 
 structure Reflection : REFLECTION =
@@ -34,7 +35,7 @@
      |> fst |> strip_comb |> fst
    val thy = ProofContext.theory_of ctxt
    val cert = Thm.cterm_of thy
-   val (((_,_),[th']), ctxt') = Variable.import_thms true [th] ctxt
+   val (((_,_),[th']), ctxt') = Variable.import 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 => Term.could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t
--- a/src/HOL/Ln.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Ln.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -343,43 +343,7 @@
 done
 
 lemma DERIV_ln: "0 < x ==> DERIV ln x :> 1 / x"
-  apply (unfold deriv_def, unfold LIM_eq, clarsimp)
-  apply (rule exI)
-  apply (rule conjI)
-  prefer 2
-  apply clarsimp
-  apply (subgoal_tac "(ln (x + xa) - ln x) / xa - (1 / x) = 
-      (ln (1 + xa / x) - xa / x) / xa")
-  apply (erule ssubst)
-  apply (subst abs_divide)
-  apply (rule mult_imp_div_pos_less)
-  apply force
-  apply (rule order_le_less_trans)
-  apply (rule abs_ln_one_plus_x_minus_x_bound)
-  apply (subst abs_divide)
-  apply (subst abs_of_pos, assumption)
-  apply (erule mult_imp_div_pos_le)
-  apply (subgoal_tac "abs xa < min (x / 2) (r * x^2 / 2)")
-  apply force
-  apply assumption
-  apply (simp add: power2_eq_square mult_compare_simps)
-  apply (rule mult_imp_div_pos_less)
-  apply (rule mult_pos_pos, assumption, assumption)
-  apply (subgoal_tac "xa * xa = abs xa * abs xa")
-  apply (erule ssubst)
-  apply (subgoal_tac "abs xa * (abs xa * 2) < abs xa * (r * (x * x))")
-  apply (simp only: mult_ac)
-  apply (rule mult_strict_left_mono)
-  apply (erule conjE, assumption)
-  apply force
-  apply simp
-  apply (subst ln_div [THEN sym])
-  apply arith
-  apply (auto simp add: algebra_simps add_frac_eq frac_eq_eq 
-    add_divide_distrib power2_eq_square)
-  apply (rule mult_pos_pos, assumption)+
-  apply assumption
-done
+  by (rule DERIV_ln[THEN DERIV_cong], simp, simp add: divide_inverse)
 
 lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)"  
 proof -
--- a/src/HOL/MacLaurin.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/MacLaurin.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -27,36 +27,6 @@
 lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))"
 by arith
 
-text{*A crude tactic to differentiate by proof.*}
-
-lemmas deriv_rulesI =
-  DERIV_ident DERIV_const DERIV_cos DERIV_cmult
-  DERIV_sin DERIV_exp DERIV_inverse DERIV_pow
-  DERIV_add DERIV_diff DERIV_mult DERIV_minus
-  DERIV_inverse_fun DERIV_quotient DERIV_fun_pow
-  DERIV_fun_exp DERIV_fun_sin DERIV_fun_cos
-  DERIV_ident DERIV_const DERIV_cos
-
-ML
-{*
-local
-exception DERIV_name;
-fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f
-|   get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f
-|   get_fun_name _ = raise DERIV_name;
-
-in
-
-fun deriv_tac ctxt = SUBGOAL (fn (prem, i) =>
-  resolve_tac @{thms deriv_rulesI} i ORELSE
-    ((rtac (read_instantiate ctxt [(("f", 0), get_fun_name prem)]
-                     @{thm DERIV_chain2}) i) handle DERIV_name => no_tac));
-
-fun DERIV_tac ctxt = ALLGOALS (fn i => REPEAT (deriv_tac ctxt i));
-
-end
-*}
-
 lemma Maclaurin_lemma2:
   assumes diff: "\<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
   assumes n: "n = Suc k"
@@ -91,13 +61,12 @@
  apply (simp (no_asm) add: divide_inverse mult_assoc del: fact_Suc power_Suc)
  apply (rule DERIV_cmult)
  apply (rule lemma_DERIV_subst)
-  apply (best intro: DERIV_chain2 intro!: DERIV_intros)
+  apply (best intro!: DERIV_intros)
  apply (subst fact_Suc)
  apply (subst real_of_nat_mult)
  apply (simp add: mult_ac)
 done
 
-
 lemma Maclaurin:
   assumes h: "0 < h"
   assumes n: "0 < n"
@@ -565,9 +534,7 @@
     apply (clarify)
     apply (subst (1 2 3) mod_Suc_eq_Suc_mod)
     apply (cut_tac m=m in mod_exhaust_less_4)
-    apply (safe, simp_all)
-    apply (rule DERIV_minus, simp)
-    apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp)
+    apply (safe, auto intro!: DERIV_intros)
     done
   from Maclaurin_all_le [OF diff_0 DERIV_diff]
   obtain t where t1: "\<bar>t\<bar> \<le> \<bar>x\<bar>" and
--- a/src/HOL/Matrix/SparseMatrix.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Matrix/SparseMatrix.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -27,8 +27,8 @@
 
 lemmas [code] = sparse_row_vector_empty [symmetric]
 
-lemma foldl_distrstart[rule_format]: "! a x y. (f (g x y) a = g x (f y a)) \<Longrightarrow> ! x y. (foldl f (g x y) l = g x (foldl f y l))"
-  by (induct l, auto)
+lemma foldl_distrstart: "! a x y. (f (g x y) a = g x (f y a)) \<Longrightarrow> (foldl f (g x y) l = g x (foldl f y l))"
+  by (induct l arbitrary: x y, auto)
 
 lemma sparse_row_vector_cons[simp]:
   "sparse_row_vector (a # arr) = (singleton_matrix 0 (fst a) (snd a)) + (sparse_row_vector arr)"
@@ -85,14 +85,14 @@
 apply (auto simp add: sorted_spvec.simps)
 done
 
-lemma sorted_sparse_row_vector_zero[rule_format]: "m <= n \<longrightarrow> sorted_spvec ((n,a)#arr) \<longrightarrow> Rep_matrix (sparse_row_vector arr) j m = 0"
+lemma sorted_sparse_row_vector_zero[rule_format]: "m <= n \<Longrightarrow> sorted_spvec ((n,a)#arr) \<longrightarrow> Rep_matrix (sparse_row_vector arr) j m = 0"
 apply (induct arr)
 apply (auto)
 apply (frule sorted_spvec_cons2,simp)+
 apply (frule sorted_spvec_cons3, simp)
 done
 
-lemma sorted_sparse_row_matrix_zero[rule_format]: "m <= n \<longrightarrow> sorted_spvec ((n,a)#arr) \<longrightarrow> Rep_matrix (sparse_row_matrix arr) m j = 0"
+lemma sorted_sparse_row_matrix_zero[rule_format]: "m <= n \<Longrightarrow> sorted_spvec ((n,a)#arr) \<longrightarrow> Rep_matrix (sparse_row_matrix arr) m j = 0"
   apply (induct arr)
   apply (auto)
   apply (frule sorted_spvec_cons2, simp)
@@ -160,19 +160,19 @@
 lemma smult_spvec_cons: "smult_spvec y (a#arr) = (fst a, y * (snd a)) # (smult_spvec y arr)"
   by (simp add: smult_spvec_def)
 
-consts addmult_spvec :: "('a::ring) * 'a spvec * 'a spvec \<Rightarrow> 'a spvec"
-recdef addmult_spvec "measure (% (y, a, b). length a + (length b))"
-  "addmult_spvec (y, arr, []) = arr"
-  "addmult_spvec (y, [], brr) = smult_spvec y brr"
-  "addmult_spvec (y, a#arr, b#brr) = (
-    if (fst a) < (fst b) then (a#(addmult_spvec (y, arr, b#brr))) 
-    else (if (fst b < fst a) then ((fst b, y * (snd b))#(addmult_spvec (y, a#arr, brr)))
-    else ((fst a, (snd a)+ y*(snd b))#(addmult_spvec (y, arr,brr)))))"
+fun addmult_spvec :: "('a::ring) \<Rightarrow> 'a spvec \<Rightarrow> 'a spvec \<Rightarrow> 'a spvec" where
+  "addmult_spvec y arr [] = arr" |
+  "addmult_spvec y [] brr = smult_spvec y brr" |
+  "addmult_spvec y ((i,a)#arr) ((j,b)#brr) = (
+    if i < j then ((i,a)#(addmult_spvec y arr ((j,b)#brr))) 
+    else (if (j < i) then ((j, y * b)#(addmult_spvec y ((i,a)#arr) brr))
+    else ((i, a + y*b)#(addmult_spvec y arr brr))))"
+(* Steven used termination "measure (% (y, a, b). length a + (length b))" *)
 
-lemma addmult_spvec_empty1[simp]: "addmult_spvec (y, [], a) = smult_spvec y a"
+lemma addmult_spvec_empty1[simp]: "addmult_spvec y [] a = smult_spvec y a"
   by (induct a) auto
 
-lemma addmult_spvec_empty2[simp]: "addmult_spvec (y, a, []) = a"
+lemma addmult_spvec_empty2[simp]: "addmult_spvec y a [] = a"
   by (induct a) auto
 
 lemma sparse_row_vector_map: "(! x y. f (x+y) = (f x) + (f y)) \<Longrightarrow> (f::'a\<Rightarrow>('a::lordered_ring)) 0 = 0 \<Longrightarrow> 
@@ -186,49 +186,43 @@
   apply (simp_all add: smult_spvec_cons scalar_mult_add)
   done
 
-lemma sparse_row_vector_addmult_spvec: "sparse_row_vector (addmult_spvec (y::'a::lordered_ring, a, b)) = 
+lemma sparse_row_vector_addmult_spvec: "sparse_row_vector (addmult_spvec (y::'a::lordered_ring) a b) = 
   (sparse_row_vector a) + (scalar_mult y (sparse_row_vector b))"
-  apply (rule addmult_spvec.induct[of _ y])
+  apply (induct y a b rule: addmult_spvec.induct)
   apply (simp add: scalar_mult_add smult_spvec_cons sparse_row_vector_smult singleton_matrix_add)+
   done
 
-lemma sorted_smult_spvec[rule_format]: "sorted_spvec a \<Longrightarrow> sorted_spvec (smult_spvec y a)"
+lemma sorted_smult_spvec: "sorted_spvec a \<Longrightarrow> sorted_spvec (smult_spvec y a)"
   apply (auto simp add: smult_spvec_def)
   apply (induct a)
   apply (auto simp add: sorted_spvec.simps split:list.split_asm)
   done
 
-lemma sorted_spvec_addmult_spvec_helper: "\<lbrakk>sorted_spvec (addmult_spvec (y, (a, b) # arr, brr)); aa < a; sorted_spvec ((a, b) # arr); 
-  sorted_spvec ((aa, ba) # brr)\<rbrakk> \<Longrightarrow> sorted_spvec ((aa, y * ba) # addmult_spvec (y, (a, b) # arr, brr))"  
+lemma sorted_spvec_addmult_spvec_helper: "\<lbrakk>sorted_spvec (addmult_spvec y ((a, b) # arr) brr); aa < a; sorted_spvec ((a, b) # arr); 
+  sorted_spvec ((aa, ba) # brr)\<rbrakk> \<Longrightarrow> sorted_spvec ((aa, y * ba) # addmult_spvec y ((a, b) # arr) brr)"  
   apply (induct brr)
   apply (auto simp add: sorted_spvec.simps)
-  apply (simp split: list.split)
-  apply (auto)
-  apply (simp split: list.split)
-  apply (auto)
   done
 
 lemma sorted_spvec_addmult_spvec_helper2: 
- "\<lbrakk>sorted_spvec (addmult_spvec (y, arr, (aa, ba) # brr)); a < aa; sorted_spvec ((a, b) # arr); sorted_spvec ((aa, ba) # brr)\<rbrakk>
-       \<Longrightarrow> sorted_spvec ((a, b) # addmult_spvec (y, arr, (aa, ba) # brr))"
+ "\<lbrakk>sorted_spvec (addmult_spvec y arr ((aa, ba) # brr)); a < aa; sorted_spvec ((a, b) # arr); sorted_spvec ((aa, ba) # brr)\<rbrakk>
+       \<Longrightarrow> sorted_spvec ((a, b) # addmult_spvec y arr ((aa, ba) # brr))"
   apply (induct arr)
   apply (auto simp add: smult_spvec_def sorted_spvec.simps)
-  apply (simp split: list.split)
-  apply (auto)
   done
 
 lemma sorted_spvec_addmult_spvec_helper3[rule_format]:
-  "sorted_spvec (addmult_spvec (y, arr, brr)) \<longrightarrow> sorted_spvec ((aa, b) # arr) \<longrightarrow> sorted_spvec ((aa, ba) # brr)
-     \<longrightarrow> sorted_spvec ((aa, b + y * ba) # (addmult_spvec (y, arr, brr)))"
-  apply (rule addmult_spvec.induct[of _ y arr brr])
-  apply (simp_all add: sorted_spvec.simps smult_spvec_def)
+  "sorted_spvec (addmult_spvec y arr brr) \<longrightarrow> sorted_spvec ((aa, b) # arr) \<longrightarrow> sorted_spvec ((aa, ba) # brr)
+     \<longrightarrow> sorted_spvec ((aa, b + y * ba) # (addmult_spvec y arr brr))"
+  apply (induct y arr brr rule: addmult_spvec.induct)
+  apply (simp_all add: sorted_spvec.simps smult_spvec_def split:list.split)
   done
 
-lemma sorted_addmult_spvec[rule_format]: "sorted_spvec a \<longrightarrow> sorted_spvec b \<longrightarrow> sorted_spvec (addmult_spvec (y, a, b))"
-  apply (rule addmult_spvec.induct[of _ y a b])
+lemma sorted_addmult_spvec: "sorted_spvec a \<Longrightarrow> sorted_spvec b \<Longrightarrow> sorted_spvec (addmult_spvec y a b)"
+  apply (induct y a b rule: addmult_spvec.induct)
   apply (simp_all add: sorted_smult_spvec)
   apply (rule conjI, intro strip)
-  apply (case_tac "~(a < aa)")
+  apply (case_tac "~(i < j)")
   apply (simp_all)
   apply (frule_tac as=brr in sorted_spvec_cons1)
   apply (simp add: sorted_spvec_addmult_spvec_helper)
@@ -242,18 +236,17 @@
   apply (simp_all add: sorted_spvec_addmult_spvec_helper3)
   done
 
-consts 
-  mult_spvec_spmat :: "('a::lordered_ring) spvec * 'a spvec * 'a spmat  \<Rightarrow> 'a spvec"
-recdef mult_spvec_spmat "measure (% (c, arr, brr). (length arr) + (length brr))"
-  "mult_spvec_spmat (c, [], brr) = c"
-  "mult_spvec_spmat (c, arr, []) = c"
-  "mult_spvec_spmat (c, a#arr, b#brr) = (
-     if ((fst a) < (fst b)) then (mult_spvec_spmat (c, arr, b#brr))
-     else (if ((fst b) < (fst a)) then (mult_spvec_spmat (c, a#arr, brr)) 
-     else (mult_spvec_spmat (addmult_spvec (snd a, c, snd b), arr, brr))))"
+fun mult_spvec_spmat :: "('a::lordered_ring) spvec \<Rightarrow> 'a spvec \<Rightarrow> 'a spmat  \<Rightarrow> 'a spvec" where
+(* recdef mult_spvec_spmat "measure (% (c, arr, brr). (length arr) + (length brr))" *)
+  "mult_spvec_spmat c [] brr = c" |
+  "mult_spvec_spmat c arr [] = c" |
+  "mult_spvec_spmat c ((i,a)#arr) ((j,b)#brr) = (
+     if (i < j) then mult_spvec_spmat c arr ((j,b)#brr)
+     else if (j < i) then mult_spvec_spmat c ((i,a)#arr) brr 
+     else mult_spvec_spmat (addmult_spvec a c b) arr brr)"
 
 lemma sparse_row_mult_spvec_spmat[rule_format]: "sorted_spvec (a::('a::lordered_ring) spvec) \<longrightarrow> sorted_spvec B \<longrightarrow> 
-  sparse_row_vector (mult_spvec_spmat (c, a, B)) = (sparse_row_vector c) + (sparse_row_vector a) * (sparse_row_matrix B)"
+  sparse_row_vector (mult_spvec_spmat c a B) = (sparse_row_vector c) + (sparse_row_vector a) * (sparse_row_matrix B)"
 proof -
   have comp_1: "!! a b. a < b \<Longrightarrow> Suc 0 <= nat ((int b)-(int a))" by arith
   have not_iff: "!! a b. a = b \<Longrightarrow> (~ a) = (~ b)" by simp
@@ -269,7 +262,7 @@
   }
   note nrows_helper = this
   show ?thesis
-    apply (rule mult_spvec_spmat.induct)
+    apply (induct c a B rule: mult_spvec_spmat.induct)
     apply simp+
     apply (rule conjI)
     apply (intro strip)
@@ -285,8 +278,8 @@
     apply (simp add: comp_1)+
     apply (subst Rep_matrix_zero_imp_mult_zero)
     apply (intro strip)
-    apply (case_tac "k <= aa")
-    apply (rule_tac m1 = k and n1 = a and a1 = b in ssubst[OF sorted_sparse_row_vector_zero])
+    apply (case_tac "k <= j")
+    apply (rule_tac m1 = k and n1 = i and a1 = a in ssubst[OF sorted_sparse_row_vector_zero])
     apply (simp_all)
     apply (rule impI)
     apply (rule disjI2)
@@ -302,11 +295,11 @@
     apply (rule disjI2)
     apply (intro strip)
     apply (simp add: sparse_row_matrix_cons neg_def)
-    apply (case_tac "a <= aa")  
+    apply (case_tac "i <= j")  
     apply (erule sorted_sparse_row_matrix_zero)  
     apply (simp_all)
     apply (intro strip)
-    apply (case_tac "a=aa")
+    apply (case_tac "i=j")
     apply (simp_all)
     apply (frule_tac as=arr in sorted_spvec_cons1)
     apply (frule_tac as=brr in sorted_spvec_cons1)
@@ -317,7 +310,7 @@
     apply (simp_all)
     apply (rule_tac A1 = "sparse_row_vector arr" in ssubst[OF Rep_matrix_zero_imp_mult_zero])
     apply (auto)
-    apply (rule_tac m=k and n = aa and a = b and arr=arr in sorted_sparse_row_vector_zero)
+    apply (rule_tac m=k and n = j and a = a and arr=arr in sorted_sparse_row_vector_zero)
     apply (simp_all)
     apply (simp add: neg_def)
     apply (drule nrows_notzero)
@@ -328,7 +321,7 @@
     apply (rule ext)+
     apply (simp)
     apply (subst Rep_matrix_mult)
-    apply (rule_tac j1=aa in ssubst[OF foldseq_almostzero])
+    apply (rule_tac j1=j in ssubst[OF foldseq_almostzero])
     apply (simp_all)
     apply (intro strip, rule conjI)
     apply (intro strip)
@@ -345,8 +338,8 @@
 qed
 
 lemma sorted_mult_spvec_spmat[rule_format]: 
-  "sorted_spvec (c::('a::lordered_ring) spvec) \<longrightarrow> sorted_spmat B \<longrightarrow> sorted_spvec (mult_spvec_spmat (c, a, B))"
-  apply (rule mult_spvec_spmat.induct[of _ c a B])
+  "sorted_spvec (c::('a::lordered_ring) spvec) \<longrightarrow> sorted_spmat B \<longrightarrow> sorted_spvec (mult_spvec_spmat c a B)"
+  apply (induct c a B rule: mult_spvec_spmat.induct)
   apply (simp_all add: sorted_addmult_spvec)
   done
 
@@ -355,10 +348,11 @@
 
 primrec 
   "mult_spmat [] A = []"
-  "mult_spmat (a#as) A = (fst a, mult_spvec_spmat ([], snd a, A))#(mult_spmat as A)"
+  "mult_spmat (a#as) A = (fst a, mult_spvec_spmat [] (snd a) A)#(mult_spmat as A)"
 
-lemma sparse_row_mult_spmat[rule_format]: 
-  "sorted_spmat A \<longrightarrow> sorted_spvec B \<longrightarrow> sparse_row_matrix (mult_spmat A B) = (sparse_row_matrix A) * (sparse_row_matrix B)"
+lemma sparse_row_mult_spmat: 
+  "sorted_spmat A \<Longrightarrow> sorted_spvec B \<Longrightarrow>
+   sparse_row_matrix (mult_spmat A B) = (sparse_row_matrix A) * (sparse_row_matrix B)"
   apply (induct A)
   apply (auto simp add: sparse_row_matrix_cons sparse_row_mult_spvec_spmat algebra_simps move_matrix_mult)
   done
@@ -372,89 +366,88 @@
   apply (auto simp add: sorted_spvec.simps)
   done
 
-lemma sorted_spmat_mult_spmat[rule_format]:
-  "sorted_spmat (B::('a::lordered_ring) spmat) \<longrightarrow> sorted_spmat (mult_spmat A B)"
+lemma sorted_spmat_mult_spmat:
+  "sorted_spmat (B::('a::lordered_ring) spmat) \<Longrightarrow> sorted_spmat (mult_spmat A B)"
   apply (induct A)
   apply (auto simp add: sorted_mult_spvec_spmat) 
   done
 
-consts
-  add_spvec :: "('a::lordered_ab_group_add) spvec * 'a spvec \<Rightarrow> 'a spvec"
-  add_spmat :: "('a::lordered_ab_group_add) spmat * 'a spmat \<Rightarrow> 'a spmat"
 
-recdef add_spvec "measure (% (a, b). length a + (length b))"
-  "add_spvec (arr, []) = arr"
-  "add_spvec ([], brr) = brr"
-  "add_spvec (a#arr, b#brr) = (
-  if (fst a) < (fst b) then (a#(add_spvec (arr, b#brr))) 
-     else (if (fst b < fst a) then (b#(add_spvec (a#arr, brr)))
-     else ((fst a, (snd a)+(snd b))#(add_spvec (arr,brr)))))"
+fun add_spvec :: "('a::lordered_ab_group_add) spvec \<Rightarrow> 'a spvec \<Rightarrow> 'a spvec" where
+(* "measure (% (a, b). length a + (length b))" *)
+  "add_spvec arr [] = arr" |
+  "add_spvec [] brr = brr" |
+  "add_spvec ((i,a)#arr) ((j,b)#brr) = (
+  if i < j then (i,a)#(add_spvec arr ((j,b)#brr)) 
+     else if (j < i) then (j,b) # add_spvec ((i,a)#arr) brr
+     else (i, a+b) # add_spvec arr brr)"
 
-lemma add_spvec_empty1[simp]: "add_spvec ([], a) = a"
-  by (induct a, auto)
+lemma add_spvec_empty1[simp]: "add_spvec [] a = a"
+by (cases a, auto)
 
-lemma add_spvec_empty2[simp]: "add_spvec (a, []) = a"
-  by (induct a, auto)
-
-lemma sparse_row_vector_add: "sparse_row_vector (add_spvec (a,b)) = (sparse_row_vector a) + (sparse_row_vector b)"
-  apply (rule add_spvec.induct[of _ a b])
+lemma sparse_row_vector_add: "sparse_row_vector (add_spvec a b) = (sparse_row_vector a) + (sparse_row_vector b)"
+  apply (induct a b rule: add_spvec.induct)
   apply (simp_all add: singleton_matrix_add)
   done
 
-recdef add_spmat "measure (% (A,B). (length A)+(length B))"
-  "add_spmat ([], bs) = bs"
-  "add_spmat (as, []) = as"
-  "add_spmat (a#as, b#bs) = (
-  if fst a < fst b then 
-    (a#(add_spmat (as, b#bs)))
-  else (if fst b < fst a then
-    (b#(add_spmat (a#as, bs)))
+fun add_spmat :: "('a::lordered_ab_group_add) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat" where
+(* "measure (% (A,B). (length A)+(length B))" *)
+  "add_spmat [] bs = bs" |
+  "add_spmat as [] = as" |
+  "add_spmat ((i,a)#as) ((j,b)#bs) = (
+  if i < j then 
+    (i,a) # add_spmat as ((j,b)#bs)
+  else if j < i then
+    (j,b) # add_spmat ((i,a)#as) bs
   else
-    ((fst a, add_spvec (snd a, snd b))#(add_spmat (as, bs)))))"
+    (i, add_spvec a b) # add_spmat as bs)"
 
-lemma sparse_row_add_spmat: "sparse_row_matrix (add_spmat (A, B)) = (sparse_row_matrix A) + (sparse_row_matrix B)"
-  apply (rule add_spmat.induct)
+lemma add_spmat_Nil2[simp]: "add_spmat as [] = as"
+by(cases as) auto
+
+lemma sparse_row_add_spmat: "sparse_row_matrix (add_spmat A B) = (sparse_row_matrix A) + (sparse_row_matrix B)"
+  apply (induct A B rule: add_spmat.induct)
   apply (auto simp add: sparse_row_matrix_cons sparse_row_vector_add move_matrix_add)
   done
 
 lemmas [code] = sparse_row_add_spmat [symmetric]
 lemmas [code] = sparse_row_vector_add [symmetric]
 
-lemma sorted_add_spvec_helper1[rule_format]: "add_spvec ((a,b)#arr, brr) = (ab, bb) # list \<longrightarrow> (ab = a | (brr \<noteq> [] & ab = fst (hd brr)))"
+lemma sorted_add_spvec_helper1[rule_format]: "add_spvec ((a,b)#arr) brr = (ab, bb) # list \<longrightarrow> (ab = a | (brr \<noteq> [] & ab = fst (hd brr)))"
   proof - 
-    have "(! x ab a. x = (a,b)#arr \<longrightarrow> add_spvec (x, brr) = (ab, bb) # list \<longrightarrow> (ab = a | (ab = fst (hd brr))))"
-      by (rule add_spvec.induct[of _ _ brr], auto)
+    have "(! x ab a. x = (a,b)#arr \<longrightarrow> add_spvec x brr = (ab, bb) # list \<longrightarrow> (ab = a | (ab = fst (hd brr))))"
+      by (induct brr rule: add_spvec.induct) (auto split:if_splits)
     then show ?thesis
       by (case_tac brr, auto)
   qed
 
-lemma sorted_add_spmat_helper1[rule_format]: "add_spmat ((a,b)#arr, brr) = (ab, bb) # list \<longrightarrow> (ab = a | (brr \<noteq> [] & ab = fst (hd brr)))"
+lemma sorted_add_spmat_helper1[rule_format]: "add_spmat ((a,b)#arr) brr = (ab, bb) # list \<longrightarrow> (ab = a | (brr \<noteq> [] & ab = fst (hd brr)))"
   proof - 
-    have "(! x ab a. x = (a,b)#arr \<longrightarrow> add_spmat (x, brr) = (ab, bb) # list \<longrightarrow> (ab = a | (ab = fst (hd brr))))"
-      by (rule add_spmat.induct[of _ _ brr], auto)
+    have "(! x ab a. x = (a,b)#arr \<longrightarrow> add_spmat x brr = (ab, bb) # list \<longrightarrow> (ab = a | (ab = fst (hd brr))))"
+      by (rule add_spmat.induct) (auto split:if_splits)
     then show ?thesis
       by (case_tac brr, auto)
   qed
 
-lemma sorted_add_spvec_helper[rule_format]: "add_spvec (arr, brr) = (ab, bb) # list \<longrightarrow> ((arr \<noteq> [] & ab = fst (hd arr)) | (brr \<noteq> [] & ab = fst (hd brr)))"
-  apply (rule add_spvec.induct[of _ arr brr])
-  apply (auto)
+lemma sorted_add_spvec_helper: "add_spvec arr brr = (ab, bb) # list \<Longrightarrow> ((arr \<noteq> [] & ab = fst (hd arr)) | (brr \<noteq> [] & ab = fst (hd brr)))"
+  apply (induct arr brr rule: add_spvec.induct)
+  apply (auto split:if_splits)
   done
 
-lemma sorted_add_spmat_helper[rule_format]: "add_spmat (arr, brr) = (ab, bb) # list \<longrightarrow> ((arr \<noteq> [] & ab = fst (hd arr)) | (brr \<noteq> [] & ab = fst (hd brr)))"
-  apply (rule add_spmat.induct[of _ arr brr])
-  apply (auto)
+lemma sorted_add_spmat_helper: "add_spmat arr brr = (ab, bb) # list \<Longrightarrow> ((arr \<noteq> [] & ab = fst (hd arr)) | (brr \<noteq> [] & ab = fst (hd brr)))"
+  apply (induct arr brr rule: add_spmat.induct)
+  apply (auto split:if_splits)
   done
 
-lemma add_spvec_commute: "add_spvec (a, b) = add_spvec (b, a)"
-  by (rule add_spvec.induct[of _ a b], auto)
+lemma add_spvec_commute: "add_spvec a b = add_spvec b a"
+by (induct a b rule: add_spvec.induct) auto
 
-lemma add_spmat_commute: "add_spmat (a, b) = add_spmat (b, a)"
-  apply (rule add_spmat.induct[of _ a b])
+lemma add_spmat_commute: "add_spmat a b = add_spmat b a"
+  apply (induct a b rule: add_spmat.induct)
   apply (simp_all add: add_spvec_commute)
   done
   
-lemma sorted_add_spvec_helper2: "add_spvec ((a,b)#arr, brr) = (ab, bb) # list \<Longrightarrow> aa < a \<Longrightarrow> sorted_spvec ((aa, ba) # brr) \<Longrightarrow> aa < ab"
+lemma sorted_add_spvec_helper2: "add_spvec ((a,b)#arr) brr = (ab, bb) # list \<Longrightarrow> aa < a \<Longrightarrow> sorted_spvec ((aa, ba) # brr) \<Longrightarrow> aa < ab"
   apply (drule sorted_add_spvec_helper1)
   apply (auto)
   apply (case_tac brr)
@@ -463,7 +456,7 @@
   apply (simp)
   done
 
-lemma sorted_add_spmat_helper2: "add_spmat ((a,b)#arr, brr) = (ab, bb) # list \<Longrightarrow> aa < a \<Longrightarrow> sorted_spvec ((aa, ba) # brr) \<Longrightarrow> aa < ab"
+lemma sorted_add_spmat_helper2: "add_spmat ((a,b)#arr) brr = (ab, bb) # list \<Longrightarrow> aa < a \<Longrightarrow> sorted_spvec ((aa, ba) # brr) \<Longrightarrow> aa < ab"
   apply (drule sorted_add_spmat_helper1)
   apply (auto)
   apply (case_tac brr)
@@ -472,51 +465,38 @@
   apply (simp)
   done
 
-lemma sorted_spvec_add_spvec[rule_format]: "sorted_spvec a \<longrightarrow> sorted_spvec b \<longrightarrow> sorted_spvec (add_spvec (a, b))"
-  apply (rule add_spvec.induct[of _ a b])
+lemma sorted_spvec_add_spvec[rule_format]: "sorted_spvec a \<longrightarrow> sorted_spvec b \<longrightarrow> sorted_spvec (add_spvec a b)"
+  apply (induct a b rule: add_spvec.induct)
   apply (simp_all)
   apply (rule conjI)
-  apply (intro strip)
-  apply (simp)
+  apply (clarsimp)
   apply (frule_tac as=brr in sorted_spvec_cons1)
   apply (simp)
   apply (subst sorted_spvec_step)
-  apply (simp split: list.split)
-  apply (clarify, simp)
-  apply (simp add: sorted_add_spvec_helper2)
+  apply (clarsimp simp: sorted_add_spvec_helper2 split: list.split)
   apply (clarify)
   apply (rule conjI)
-  apply (case_tac "a=aa")
-  apply (simp)
   apply (clarify)
   apply (frule_tac as=arr in sorted_spvec_cons1, simp)
   apply (subst sorted_spvec_step)
-  apply (simp split: list.split)
-  apply (clarify, simp)
-  apply (simp add: sorted_add_spvec_helper2 add_spvec_commute)
-  apply (case_tac "a=aa")
-  apply (simp_all)
+  apply (clarsimp simp: sorted_add_spvec_helper2 add_spvec_commute split: list.split)
   apply (clarify)
   apply (frule_tac as=arr in sorted_spvec_cons1)
   apply (frule_tac as=brr in sorted_spvec_cons1)
   apply (simp)
   apply (subst sorted_spvec_step)
   apply (simp split: list.split)
-  apply (clarify, simp)
+  apply (clarsimp)
   apply (drule_tac sorted_add_spvec_helper)
-  apply (auto)
-  apply (case_tac arr)
-  apply (simp_all)
+  apply (auto simp: neq_Nil_conv)
   apply (drule sorted_spvec_cons3)
   apply (simp)
-  apply (case_tac brr)
-  apply (simp_all)
   apply (drule sorted_spvec_cons3)
   apply (simp)
   done
 
-lemma sorted_spvec_add_spmat[rule_format]: "sorted_spvec A \<longrightarrow> sorted_spvec B \<longrightarrow> sorted_spvec (add_spmat (A, B))"
-  apply (rule add_spmat.induct[of _ A B])
+lemma sorted_spvec_add_spmat[rule_format]: "sorted_spvec A \<longrightarrow> sorted_spvec B \<longrightarrow> sorted_spvec (add_spmat A B)"
+  apply (induct A B rule: add_spmat.induct)
   apply (simp_all)
   apply (rule conjI)
   apply (intro strip)
@@ -529,17 +509,11 @@
   apply (simp add: sorted_add_spmat_helper2)
   apply (clarify)
   apply (rule conjI)
-  apply (case_tac "a=aa")
-  apply (simp)
   apply (clarify)
   apply (frule_tac as=as in sorted_spvec_cons1, simp)
   apply (subst sorted_spvec_step)
-  apply (simp split: list.split)
-  apply (clarify, simp)
-  apply (simp add: sorted_add_spmat_helper2 add_spmat_commute)
-  apply (case_tac "a=aa")
-  apply (simp_all)
-  apply (clarify)
+  apply (clarsimp simp: sorted_add_spmat_helper2 add_spmat_commute split: list.split)
+  apply (clarsimp)
   apply (frule_tac as=as in sorted_spvec_cons1)
   apply (frule_tac as=bs in sorted_spvec_cons1)
   apply (simp)
@@ -547,49 +521,37 @@
   apply (simp split: list.split)
   apply (clarify, simp)
   apply (drule_tac sorted_add_spmat_helper)
-  apply (auto)
-  apply (case_tac as)
-  apply (simp_all)
+  apply (auto simp:neq_Nil_conv)
   apply (drule sorted_spvec_cons3)
   apply (simp)
-  apply (case_tac bs)
-  apply (simp_all)
   apply (drule sorted_spvec_cons3)
   apply (simp)
   done
 
-lemma sorted_spmat_add_spmat[rule_format]: "sorted_spmat A \<longrightarrow> sorted_spmat B \<longrightarrow> sorted_spmat (add_spmat (A, B))"
-  apply (rule add_spmat.induct[of _ A B])
+lemma sorted_spmat_add_spmat[rule_format]: "sorted_spmat A \<Longrightarrow> sorted_spmat B \<Longrightarrow> sorted_spmat (add_spmat A B)"
+  apply (induct A B rule: add_spmat.induct)
   apply (simp_all add: sorted_spvec_add_spvec)
   done
 
-consts
-  le_spvec :: "('a::lordered_ab_group_add) spvec * 'a spvec \<Rightarrow> bool" 
-  le_spmat :: "('a::lordered_ab_group_add) spmat * 'a spmat \<Rightarrow> bool" 
+fun le_spvec :: "('a::lordered_ab_group_add) spvec \<Rightarrow> 'a spvec \<Rightarrow> bool" where
+(* "measure (% (a,b). (length a) + (length b))" *)
+  "le_spvec [] [] = True" |
+  "le_spvec ((_,a)#as) [] = (a <= 0 & le_spvec as [])" |
+  "le_spvec [] ((_,b)#bs) = (0 <= b & le_spvec [] bs)" |
+  "le_spvec ((i,a)#as) ((j,b)#bs) = (
+  if (i < j) then a <= 0 & le_spvec as ((j,b)#bs)
+  else if (j < i) then 0 <= b & le_spvec ((i,a)#as) bs
+  else a <= b & le_spvec as bs)"
 
-recdef le_spvec "measure (% (a,b). (length a) + (length b))" 
-  "le_spvec ([], []) = True"
-  "le_spvec (a#as, []) = ((snd a <= 0) & (le_spvec (as, [])))"
-  "le_spvec ([], b#bs) = ((0 <= snd b) & (le_spvec ([], bs)))"
-  "le_spvec (a#as, b#bs) = (
-  if (fst a < fst b) then 
-    ((snd a <= 0) & (le_spvec (as, b#bs)))
-  else (if (fst b < fst a) then
-    ((0 <= snd b) & (le_spvec (a#as, bs)))
-  else 
-    ((snd a <= snd b) & (le_spvec (as, bs)))))"
-
-recdef le_spmat "measure (% (a,b). (length a) + (length b))"
-  "le_spmat ([], []) = True"
-  "le_spmat (a#as, []) = (le_spvec (snd a, []) & (le_spmat (as, [])))"
-  "le_spmat ([], b#bs) = (le_spvec ([], snd b) & (le_spmat ([], bs)))"
-  "le_spmat (a#as, b#bs) = (
-  if fst a < fst b then
-    (le_spvec(snd a,[]) & le_spmat(as, b#bs))
-  else (if (fst b < fst a) then 
-    (le_spvec([], snd b) & le_spmat(a#as, bs))
-  else
-    (le_spvec(snd a, snd b) & le_spmat (as, bs))))"
+fun le_spmat :: "('a::lordered_ab_group_add) spmat \<Rightarrow> 'a spmat \<Rightarrow> bool" where
+(* "measure (% (a,b). (length a) + (length b))" *)
+  "le_spmat [] [] = True" |
+  "le_spmat ((i,a)#as) [] = (le_spvec a [] & le_spmat as [])" |
+  "le_spmat [] ((j,b)#bs) = (le_spvec [] b & le_spmat [] bs)" |
+  "le_spmat ((i,a)#as) ((j,b)#bs) = (
+  if i < j then (le_spvec a [] & le_spmat as ((j,b)#bs))
+  else if j < i then (le_spvec [] b & le_spmat ((i,a)#as) bs)
+  else (le_spvec a b & le_spmat as bs))"
 
 constdefs
   disj_matrices :: "('a::zero) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
@@ -701,8 +663,8 @@
   apply (rule nrows, rule order_trans[of _ 1], simp, drule nrows_notzero, drule less_le_trans[OF _ nrows_spvec], arith)+
   done
 
-lemma le_spvec_iff_sparse_row_le[rule_format]: "(sorted_spvec a) \<longrightarrow> (sorted_spvec b) \<longrightarrow> (le_spvec (a,b)) = (sparse_row_vector a <= sparse_row_vector b)"
-  apply (rule le_spvec.induct)
+lemma le_spvec_iff_sparse_row_le[rule_format]: "(sorted_spvec a) \<longrightarrow> (sorted_spvec b) \<longrightarrow> (le_spvec a b) = (sparse_row_vector a <= sparse_row_vector b)"
+  apply (induct a b rule: le_spvec.induct)
   apply (simp_all add: sorted_spvec_cons1 disj_matrices_add_le_zero disj_matrices_add_zero_le 
     disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
   apply (rule conjI, intro strip)
@@ -718,34 +680,30 @@
   apply (blast)
   apply (intro strip)
   apply (simp add: sorted_spvec_cons1)
-  apply (case_tac "a=aa", simp_all)
+  apply (case_tac "a=b", simp_all)
   apply (subst disj_matrices_add)
   apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
   done
 
-lemma le_spvec_empty2_sparse_row[rule_format]: "(sorted_spvec b) \<longrightarrow> (le_spvec (b,[]) = (sparse_row_vector b <= 0))"
+lemma le_spvec_empty2_sparse_row[rule_format]: "sorted_spvec b \<longrightarrow> le_spvec b [] = (sparse_row_vector b <= 0)"
   apply (induct b)
   apply (simp_all add: sorted_spvec_cons1)
   apply (intro strip)
   apply (subst disj_matrices_add_le_zero)
-  apply (simp add: disj_matrices_commute disj_sparse_row_singleton sorted_spvec_cons1)
-  apply (rule_tac y = "snd a" in disj_sparse_row_singleton[OF order_refl])
-  apply (simp_all)
+  apply (auto simp add: disj_matrices_commute disj_sparse_row_singleton[OF order_refl] sorted_spvec_cons1)
   done
 
-lemma le_spvec_empty1_sparse_row[rule_format]: "(sorted_spvec b) \<longrightarrow> (le_spvec ([],b) = (0 <= sparse_row_vector b))"
+lemma le_spvec_empty1_sparse_row[rule_format]: "(sorted_spvec b) \<longrightarrow> (le_spvec [] b = (0 <= sparse_row_vector b))"
   apply (induct b)
   apply (simp_all add: sorted_spvec_cons1)
   apply (intro strip)
   apply (subst disj_matrices_add_zero_le)
-  apply (simp add: disj_matrices_commute disj_sparse_row_singleton sorted_spvec_cons1)
-  apply (rule_tac y = "snd a" in disj_sparse_row_singleton[OF order_refl])
-  apply (simp_all)
+  apply (auto simp add: disj_matrices_commute disj_sparse_row_singleton[OF order_refl] sorted_spvec_cons1)
   done
 
 lemma le_spmat_iff_sparse_row_le[rule_format]: "(sorted_spvec A) \<longrightarrow> (sorted_spmat A) \<longrightarrow> (sorted_spvec B) \<longrightarrow> (sorted_spmat B) \<longrightarrow> 
-  le_spmat(A, B) = (sparse_row_matrix A <= sparse_row_matrix B)"
-  apply (rule le_spmat.induct)
+  le_spmat A B = (sparse_row_matrix A <= sparse_row_matrix B)"
+  apply (induct A B rule: le_spmat.induct)
   apply (simp add: sparse_row_matrix_cons disj_matrices_add_le_zero disj_matrices_add_zero_le disj_move_sparse_vec_mat[OF order_refl] 
     disj_matrices_commute sorted_spvec_cons1 le_spvec_empty2_sparse_row le_spvec_empty1_sparse_row)+ 
   apply (rule conjI, intro strip)
@@ -765,7 +723,7 @@
   apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute)
   apply (simp, blast)
   apply (intro strip)
-  apply (case_tac "a=aa")
+  apply (case_tac "i=j")
   apply (simp_all)
   apply (subst disj_matrices_add)
   apply (simp_all add: disj_matrices_commute disj_move_sparse_vec_mat[OF order_refl])
@@ -774,16 +732,12 @@
 
 declare [[simp_depth_limit = 999]]
 
-consts 
-   abs_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat"
-   minus_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat"
-
-primrec
-  "abs_spmat [] = []"
+primrec abs_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat" where
+  "abs_spmat [] = []" |
   "abs_spmat (a#as) = (fst a, abs_spvec (snd a))#(abs_spmat as)"
 
-primrec
-  "minus_spmat [] = []"
+primrec minus_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat" where
+  "minus_spmat [] = []" |
   "minus_spmat (a#as) = (fst a, minus_spvec (snd a))#(minus_spmat as)"
 
 lemma sparse_row_matrix_minus:
@@ -851,7 +805,7 @@
 
 constdefs
   diff_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
-  "diff_spmat A B == add_spmat (A, minus_spmat B)"
+  "diff_spmat A B == add_spmat A (minus_spmat B)"
 
 lemma sorted_spmat_diff_spmat: "sorted_spmat A \<Longrightarrow> sorted_spmat B \<Longrightarrow> sorted_spmat (diff_spmat A B)"
   by (simp add: diff_spmat_def sorted_spmat_minus_spmat sorted_spmat_add_spmat)
@@ -1064,8 +1018,8 @@
 constdefs
   mult_est_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
   "mult_est_spmat r1 r2 s1 s2 == 
-  add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2), add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2), 
-  add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1), mult_spmat (nprt_spmat s1) (nprt_spmat r1))))"  
+  add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2)) (add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2)) 
+  (add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1)) (mult_spmat (nprt_spmat s1) (nprt_spmat r1))))"  
 
 lemmas sparse_row_matrix_op_simps =
   sorted_sparse_matrix_imp_spmat sorted_sparse_matrix_imp_spvec
--- a/src/HOL/Matrix/cplex/Cplex.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Matrix/cplex/Cplex.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -19,7 +19,7 @@
   "sorted_sparse_matrix r1"
   "sorted_sparse_matrix r2"
   "sorted_spvec b"
-  "le_spmat ([], y)"
+  "le_spmat [] y"
   "sparse_row_matrix A1 \<le> A"
   "A \<le> sparse_row_matrix A2"
   "sparse_row_matrix c1 \<le> c"
@@ -28,10 +28,10 @@
   "x \<le> sparse_row_matrix r2"
   "A * x \<le> sparse_row_matrix (b::('a::lordered_ring) spmat)"
   shows
-  "c * x \<le> sparse_row_matrix (add_spmat (mult_spmat y b,
+  "c * x \<le> sparse_row_matrix (add_spmat (mult_spmat y b)
   (let s1 = diff_spmat c1 (mult_spmat y A2); s2 = diff_spmat c2 (mult_spmat y A1) in 
-  add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2), add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2), 
-  add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1), mult_spmat (nprt_spmat s1) (nprt_spmat r1)))))))"
+  add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2)) (add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2)) 
+  (add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1)) (mult_spmat (nprt_spmat s1) (nprt_spmat r1))))))"
   apply (simp add: Let_def)
   apply (insert assms)
   apply (simp add: sparse_row_matrix_op_simps algebra_simps)  
@@ -49,7 +49,7 @@
   "sorted_sparse_matrix r1"
   "sorted_sparse_matrix r2"
   "sorted_spvec b"
-  "le_spmat ([], y)"
+  "le_spmat [] y"
   "sparse_row_matrix A1 \<le> A"
   "A \<le> sparse_row_matrix A2"
   "sparse_row_matrix c1 \<le> c"
@@ -58,8 +58,8 @@
   "x \<le> sparse_row_matrix r2"
   "A * x \<le> sparse_row_matrix (b::('a::lordered_ring) spmat)"
   shows
-  "c * x \<le> sparse_row_matrix (add_spmat (mult_spmat y b,
-  mult_est_spmat r1 r2 (diff_spmat c1 (mult_spmat y A2)) (diff_spmat c2 (mult_spmat y A1))))"
+  "c * x \<le> sparse_row_matrix (add_spmat (mult_spmat y b)
+  (mult_est_spmat r1 r2 (diff_spmat c1 (mult_spmat y A2)) (diff_spmat c2 (mult_spmat y A1))))"
   by (simp add: assms mult_est_spmat_def spm_mult_le_dual_prts[where A=A, simplified Let_def])
 
 use "matrixlp.ML"
--- a/src/HOL/MicroJava/BV/BVExample.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/MicroJava/BV/BVExample.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/BV/BVExample.thy
-    ID:         $Id$
     Author:     Gerwin Klein
 *)
 
@@ -94,7 +93,7 @@
 
 text {* Method and field lookup: *}
 lemma method_Object [simp]:
-  "method (E, Object) = empty"
+  "method (E, Object) = Map.empty"
   by (simp add: method_rec_lemma [OF class_Object wf_subcls1_E])
 
 lemma method_append [simp]:
@@ -436,7 +435,7 @@
   "some_elem = (%S. SOME x. x : S)"
 
 consts_code
-  "some_elem" ("hd")
+  "some_elem" ("(case/ _ of/ {*Set*}/ xs/ =>/ hd/ xs)")
 
 text {* This code setup is just a demonstration and \emph{not} sound! *}
 
@@ -451,11 +450,11 @@
 qed
 
 lemma [code]:
-  "iter f step ss w = while (\<lambda>(ss, w). \<not> (is_empty w))
+  "iter f step ss w = while (\<lambda>(ss, w). \<not> is_empty w)
     (\<lambda>(ss, w).
         let p = some_elem w in propa f (step p (ss ! p)) ss (w - {p}))
     (ss, w)"
-  unfolding iter_def is_empty_def some_elem_def ..
+  unfolding iter_def List_Set.is_empty_def some_elem_def ..
 
 lemma JVM_sup_unfold [code]:
  "JVMType.sup S m n = lift2 (Opt.sup
@@ -475,7 +474,6 @@
   test1 = "test_kil E list_name [Class list_name] (PrimT Void) 3 0
     [(Suc 0, 2, 8, Xcpt NullPointer)] append_ins"
   test2 = "test_kil E test_name [] (PrimT Void) 3 2 [] make_list_ins"
-
 ML BV.test1
 ML BV.test2
 
--- a/src/HOL/NewNumberTheory/Cong.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/NewNumberTheory/Cong.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -60,9 +60,7 @@
 lemma nat_1' [simp]: "nat 1 = 1"
 by simp
 
-(* For those annoying moments where Suc reappears *)
-lemma Suc_remove: "Suc n = n + 1"
-by simp
+(* For those annoying moments where Suc reappears, use Suc_eq_plus1 *)
 
 declare nat_1 [simp del]
 declare add_2_eq_Suc [simp del] 
--- a/src/HOL/NewNumberTheory/Fib.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/NewNumberTheory/Fib.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -145,7 +145,7 @@
   apply (subst nat_fib_reduce)
   apply (auto simp add: ring_simps)
   apply (subst (1 3 5) nat_fib_reduce)
-  apply (auto simp add: ring_simps Suc_remove)
+  apply (auto simp add: ring_simps Suc_eq_plus1)
 (* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
   apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
   apply (erule ssubst) back back
@@ -220,7 +220,7 @@
   apply (induct n rule: nat_fib_induct)
   apply auto
   apply (subst (2) nat_fib_reduce)
-  apply (auto simp add: Suc_remove) (* again, natdiff_cancel *)
+  apply (auto simp add: Suc_eq_plus1) (* again, natdiff_cancel *)
   apply (subst add_commute, auto)
   apply (subst nat_gcd_commute, auto simp add: ring_simps)
 done
--- a/src/HOL/NewNumberTheory/Residues.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/NewNumberTheory/Residues.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -202,9 +202,6 @@
   apply (subgoal_tac "a mod m ~= 0")
   apply arith
   apply auto
-  apply (subst (asm) int_gcd_commute)
-  apply (subst (asm) int_gcd_mult)
-  apply auto
   apply (subst (asm) int_gcd_red)
   apply (subst int_gcd_commute, assumption)
 done
--- a/src/HOL/NthRoot.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/NthRoot.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -372,6 +372,41 @@
     using odd_pos [OF n] by (rule isCont_real_root)
 qed
 
+lemma DERIV_even_real_root:
+  assumes n: "0 < n" and "even n"
+  assumes x: "x < 0"
+  shows "DERIV (root n) x :> inverse (- real n * root n x ^ (n - Suc 0))"
+proof (rule DERIV_inverse_function)
+  show "x - 1 < x" by simp
+  show "x < 0" using x .
+next
+  show "\<forall>y. x - 1 < y \<and> y < 0 \<longrightarrow> - (root n y ^ n) = y"
+  proof (rule allI, rule impI, erule conjE)
+    fix y assume "x - 1 < y" and "y < 0"
+    hence "root n (-y) ^ n = -y" using `0 < n` by simp
+    with real_root_minus[OF `0 < n`] and `even n`
+    show "- (root n y ^ n) = y" by simp
+  qed
+next
+  show "DERIV (\<lambda>x. - (x ^ n)) (root n x) :> - real n * root n x ^ (n - Suc 0)"
+    by  (auto intro!: DERIV_intros)
+  show "- real n * root n x ^ (n - Suc 0) \<noteq> 0"
+    using n x by simp
+  show "isCont (root n) x"
+    using n by (rule isCont_real_root)
+qed
+
+lemma DERIV_real_root_generic:
+  assumes "0 < n" and "x \<noteq> 0"
+  and even: "\<lbrakk> even n ; 0 < x \<rbrakk> \<Longrightarrow> D = inverse (real n * root n x ^ (n - Suc 0))"
+  and even: "\<lbrakk> even n ; x < 0 \<rbrakk> \<Longrightarrow> D = - inverse (real n * root n x ^ (n - Suc 0))"
+  and odd: "odd n \<Longrightarrow> D = inverse (real n * root n x ^ (n - Suc 0))"
+  shows "DERIV (root n) x :> D"
+using assms by (cases "even n", cases "0 < x",
+  auto intro: DERIV_real_root[THEN DERIV_cong]
+              DERIV_odd_real_root[THEN DERIV_cong]
+              DERIV_even_real_root[THEN DERIV_cong])
+
 subsection {* Square Root *}
 
 definition
@@ -456,9 +491,21 @@
 lemma isCont_real_sqrt: "isCont sqrt x"
 unfolding sqrt_def by (rule isCont_real_root [OF pos2])
 
+lemma DERIV_real_sqrt_generic:
+  assumes "x \<noteq> 0"
+  assumes "x > 0 \<Longrightarrow> D = inverse (sqrt x) / 2"
+  assumes "x < 0 \<Longrightarrow> D = - inverse (sqrt x) / 2"
+  shows "DERIV sqrt x :> D"
+  using assms unfolding sqrt_def
+  by (auto intro!: DERIV_real_root_generic)
+
 lemma DERIV_real_sqrt:
   "0 < x \<Longrightarrow> DERIV sqrt x :> inverse (sqrt x) / 2"
-unfolding sqrt_def by (rule DERIV_real_root [OF pos2, simplified])
+  using DERIV_real_sqrt_generic by simp
+
+declare
+  DERIV_real_sqrt_generic[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
+  DERIV_real_root_generic[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
 
 lemma not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)"
 apply auto
--- a/src/HOL/README.html	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/README.html	Wed Jul 01 16:19:44 2009 +0100
@@ -96,7 +96,7 @@
 <dt>Real
 <dd>the real numbers, part of Complex
 
-<dt>Real/HahnBanach
+<dt>Hahn_Banach
 <dd>the Hahn-Banach theorem for real vector spaces (in Isabelle/Isar)
 
 <dt>SET-Protocol
--- a/src/HOL/Tools/Datatype/datatype.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/Datatype/datatype.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -18,7 +18,7 @@
   val the_info : theory -> string -> info
   val the_descr : theory -> string list
     -> descr * (string * sort) list * string list
-      * (string list * string list) * (typ list * typ list)
+      * string * (string list * string list) * (typ list * typ list)
   val the_spec : theory -> string -> (string * sort) list * (string * typ list) list
   val get_constrs : theory -> string -> (string * typ) list option
   val get_all : theory -> info Symtab.table
@@ -125,9 +125,10 @@
 
     val names = map Long_Name.base_name (the_default tycos (#alt_names info));
     val (auxnames, _) = Name.make_context names
-      |> fold_map (yield_singleton Name.variants o name_of_typ) Us
+      |> fold_map (yield_singleton Name.variants o name_of_typ) Us;
+    val prefix = space_implode "_" names;
 
-  in (descr, vs, tycos, (names, auxnames), (Ts, Us)) end;
+  in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end;
 
 fun get_constrs thy dtco =
   case try (the_spec thy) dtco
@@ -385,7 +386,7 @@
 fun prove_rep_datatype (config : config) alt_names new_type_names descr sorts induct inject half_distinct thy =
   let
     val ((_, [induct']), _) =
-      Variable.importT_thms [induct] (Variable.thm_context induct);
+      Variable.importT [induct] (Variable.thm_context induct);
 
     fun err t = error ("Ill-formed predicate in induction rule: " ^
       Syntax.string_of_term_global thy t);
--- a/src/HOL/Tools/atp_manager.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/atp_manager.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -17,6 +17,7 @@
   val get_timeout: unit -> int
   val set_timeout: int -> unit
   val get_full_types: unit -> bool
+  val set_full_types: bool -> unit
   val kill: unit -> unit
   val info: unit -> unit
   val messages: int option -> unit
@@ -57,6 +58,7 @@
 fun set_timeout time = CRITICAL (fn () => timeout := time);
 
 fun get_full_types () = CRITICAL (fn () => ! full_types);
+fun set_full_types bool = CRITICAL (fn () => full_types := bool);
 
 val _ =
   ProofGeneralPgip.add_preference Preferences.category_proof
--- a/src/HOL/Tools/atp_wrapper.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/atp_wrapper.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -23,8 +23,9 @@
   val eprover_full: AtpManager.prover
   val spass_opts: int -> bool  -> AtpManager.prover
   val spass: AtpManager.prover
-  val remote_prover_opts: int -> bool -> string -> AtpManager.prover
-  val remote_prover: string -> AtpManager.prover
+  val remote_prover_opts: int -> bool -> string -> string -> AtpManager.prover
+  val remote_prover: string -> string -> AtpManager.prover
+  val refresh_systems: unit -> unit
 end;
 
 structure AtpWrapper: ATP_WRAPPER =
@@ -74,20 +75,16 @@
 
     (* write out problem file and call prover *)
     val probfile = prob_pathname subgoalno
-    val fname = File.platform_path probfile
-    val _ = writer fname clauses
-    val cmdline =
-      if File.exists cmd then "exec " ^ File.shell_path cmd ^ " " ^ args
-      else error ("Bad executable: " ^ Path.implode cmd)
-    val (proof, rc) = system_out (cmdline ^ " " ^ fname)
+    val conj_pos = writer probfile clauses
+    val (proof, rc) = system_out (
+      if File.exists cmd then
+        space_implode " " ["exec", File.shell_path cmd, args, File.platform_path probfile]
+      else error ("Bad executable: " ^ Path.implode cmd))
 
     (* if problemfile has not been exported, delete problemfile; otherwise export proof, too *)
     val _ =
-      if destdir' = "" then OS.FileSys.remove fname
-      else
-        let val out = TextIO.openOut (fname ^ "_proof")
-        val _ = TextIO.output (out, proof)
-        in TextIO.closeOut out end
+      if destdir' = "" then File.rm probfile
+      else File.write (Path.explode (Path.implode probfile ^ "_proof")) proof
     
     (* check for success and print out some information on failure *)
     val failure = find_failure proof
@@ -95,7 +92,8 @@
     val message =
       if is_some failure then "External prover failed."
       else if rc <> 0 then "External prover failed: " ^ proof
-      else "Try this command: " ^ produce_answer name (proof, thm_names, ctxt, th, subgoalno)
+      else "Try this command: " ^
+        produce_answer name (proof, thm_names, conj_pos, ctxt, th, subgoalno)
     val _ = Output.debug (fn () => "Sledgehammer response (rc = " ^ string_of_int rc ^ "):\n" ^ proof)
   in (success, message, proof, thm_names, the_filtered_clauses) end;
 
@@ -112,7 +110,7 @@
   (ResHolClause.tptp_write_file (AtpManager.get_full_types()))
   command
   ResReconstruct.find_failure
-  (if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list_tstp)
+  (if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list false)
   timeout ax_clauses fcls name n goal;
 
 (*arbitrary ATP with TPTP input/output and problemfile as last argument*)
@@ -146,7 +144,7 @@
                ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
   timeout;
 
-val vampire_full = vampire_opts 60 false;
+val vampire_full = vampire_opts_full 60 false;
 
 
 (* E prover *)
@@ -177,7 +175,7 @@
   (Path.explode "$SPASS_HOME/SPASS",
     "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^ string_of_int timeout)
   ResReconstruct.find_failure
-  ResReconstruct.lemma_list_dfg
+  (ResReconstruct.lemma_list true)
   timeout ax_clauses fcls name n goal;
 
 val spass = spass_opts 40 true;
@@ -185,10 +183,32 @@
 
 (* remote prover invocation via SystemOnTPTP *)
 
-fun remote_prover_opts max_new theory_const args timeout =
-  tptp_prover_opts max_new theory_const
-  (Path.explode "$ISABELLE_HOME/lib/scripts/SystemOnTPTP", args ^ " -t " ^ string_of_int timeout)
-  timeout;
+val systems =
+  Synchronized.var "atp_wrapper_systems" ([]: string list);
+
+fun get_systems () =
+  let
+    val (answer, rc) = system_out (("$ISABELLE_HOME/lib/scripts/SystemOnTPTP" |>
+      Path.explode |> File.shell_path) ^ " -w")
+  in
+    if rc <> 0 then error ("Get available systems from SystemOnTPTP:\n" ^ answer)
+    else split_lines answer
+  end;
+
+fun refresh_systems () = Synchronized.change systems (fn _ =>
+  get_systems());
+
+fun get_system prefix = Synchronized.change_result systems (fn systems =>
+  let val systems = if null systems then get_systems() else systems
+  in (find_first (String.isPrefix prefix) systems, systems) end);
+
+fun remote_prover_opts max_new theory_const args prover_prefix timeout =
+  let val sys = case get_system prover_prefix of
+      NONE => error ("No system like " ^ quote prover_prefix ^ " at SystemOnTPTP")
+    | SOME sys => sys
+  in tptp_prover_opts max_new theory_const
+    (Path.explode "$ISABELLE_HOME/lib/scripts/SystemOnTPTP",
+      args ^ " -t " ^ string_of_int timeout ^ " -s " ^ sys) timeout end;
 
 val remote_prover = remote_prover_opts 60 false;
 
--- a/src/HOL/Tools/dseq.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/dseq.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/dseq.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Sequences with recursion depth limit.
--- a/src/HOL/Tools/inductive_codegen.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/inductive_codegen.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -378,7 +378,7 @@
           end
       | compile_prems out_ts vs names ps gr =
           let
-            val vs' = distinct (op =) (List.concat (vs :: map term_vs out_ts));
+            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
             val SOME (p, mode as SOME (Mode (_, js, _))) =
               select_mode_prem thy modes' vs' ps;
             val ps' = filter_out (equal p) ps;
@@ -404,7 +404,9 @@
                          (compile_expr thy defs dep module false modes
                            (mode, t) gr2)
                      else
-                       apfst (fn p => [str "DSeq.of_list", Pretty.brk 1, p])
+                       apfst (fn p => Pretty.breaks [str "DSeq.of_list", str "(case", p,
+                         str "of", str "Set", str "xs", str "=>", str "xs)"])
+                         (*this is a very strong assumption about the generated code!*)
                            (invoke_codegen thy defs dep module true t gr2);
                    val (rest, gr4) = compile_prems out_ts''' vs' (fst nvs) ps' gr3;
                  in
@@ -661,8 +663,10 @@
                     let val (call_p, gr') = mk_ind_call thy defs dep module true
                       s T (ts1 @ ts2') names thyname k intrs gr 
                     in SOME ((if brack then parens else I) (Pretty.block
-                      [str "DSeq.list_of", Pretty.brk 1, str "(",
-                       conv_ntuple fs ots call_p, str ")"]), gr')
+                      [str "Set", Pretty.brk 1, str "(DSeq.list_of", Pretty.brk 1, str "(",
+                       conv_ntuple fs ots call_p, str "))"]),
+                       (*this is a very strong assumption about the generated code!*)
+                       gr')
                     end
                   else NONE
                 end
--- a/src/HOL/Tools/quickcheck_generators.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/quickcheck_generators.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -11,10 +11,10 @@
     -> (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed)
     -> seed -> (('a -> 'b) * (unit -> Term.term)) * seed
   val ensure_random_typecopy: string -> theory -> theory
-  val random_aux_specification: string -> term list -> local_theory -> local_theory
+  val random_aux_specification: string -> string -> term list -> local_theory -> local_theory
   val mk_random_aux_eqs: theory -> Datatype.descr -> (string * sort) list
     -> string list -> string list * string list -> typ list * typ list
-    -> string * (term list * (term * term) list)
+    -> term list * (term * term) list
   val ensure_random_datatype: Datatype.config -> string list -> theory -> theory
   val eval_ref: (unit -> int -> seed -> term list option * seed) option ref
   val setup: theory -> theory
@@ -184,18 +184,18 @@
 
 end;
 
-fun random_aux_primrec_multi prefix [eq] lthy =
+fun random_aux_primrec_multi auxname [eq] lthy =
       lthy
       |> random_aux_primrec eq
       |>> (fn simp => [simp])
-  | random_aux_primrec_multi prefix (eqs as _ :: _ :: _) lthy =
+  | random_aux_primrec_multi auxname (eqs as _ :: _ :: _) lthy =
       let
         val thy = ProofContext.theory_of lthy;
         val (lhss, rhss) = map_split (HOLogic.dest_eq o HOLogic.dest_Trueprop) eqs;
         val (vs, (arg as Free (v, _)) :: _) = map_split (fn (t1 $ t2) => (t1, t2)) lhss;
         val Ts = map fastype_of lhss;
         val tupleT = foldr1 HOLogic.mk_prodT Ts;
-        val aux_lhs = Free ("mutual_" ^ prefix, fastype_of arg --> tupleT) $ arg;
+        val aux_lhs = Free ("mutual_" ^ auxname, fastype_of arg --> tupleT) $ arg;
         val aux_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
           (aux_lhs, foldr1 HOLogic.mk_prod rhss);
         fun mk_proj t [T] = [t]
@@ -223,7 +223,7 @@
         |-> (fn (aux_simp, proj_defs) => prove_eqs aux_simp proj_defs)
       end;
 
-fun random_aux_specification prefix eqs lthy =
+fun random_aux_specification prfx name eqs lthy =
   let
     val vs = fold Term.add_free_names ((snd o strip_comb o fst o HOLogic.dest_eq
       o HOLogic.dest_Trueprop o hd) eqs) [];
@@ -237,10 +237,10 @@
         val ext_simps = map (fn thm => fun_cong OF [fun_cong OF [thm]]) proto_simps;
         val tac = ALLGOALS (ProofContext.fact_tac ext_simps);
       in (map (fn prop => SkipProof.prove lthy vs [] prop (K tac)) eqs, lthy) end;
-    val b = Binding.qualify true prefix (Binding.name "simps");
+    val b = Binding.qualify true prfx (Binding.qualify true name (Binding.name "simps"));
   in
     lthy
-    |> random_aux_primrec_multi prefix proto_eqs
+    |> random_aux_primrec_multi (name ^ prfx) proto_eqs
     |-> (fn proto_simps => prove_simps proto_simps)
     |-> (fn simps => LocalTheory.note Thm.generatedK ((b,
            Code.add_default_eqn_attrib :: map (Attrib.internal o K)
@@ -252,6 +252,8 @@
 
 (* constructing random instances on datatypes *)
 
+val random_auxN = "random_aux";
+
 fun mk_random_aux_eqs thy descr vs tycos (names, auxnames) (Ts, Us) =
   let
     val mk_const = curry (Sign.mk_const thy);
@@ -259,7 +261,6 @@
     val i1 = @{term "(i\<Colon>code_numeral) - 1"};
     val j = @{term "j\<Colon>code_numeral"};
     val seed = @{term "s\<Colon>Random.seed"};
-    val random_auxN = "random_aux";
     val random_auxsN = map (prefix (random_auxN ^ "_")) (names @ auxnames);
     fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit \<Rightarrow> term"});
     val rTs = Ts @ Us;
@@ -316,10 +317,9 @@
           $ seed;
     val auxs_lhss = map (fn t => t $ i $ j $ seed) random_auxs;
     val auxs_rhss = map mk_select gen_exprss;
-    val prefix = space_implode "_" (random_auxN :: names);
-  in (prefix, (random_auxs, auxs_lhss ~~ auxs_rhss)) end;
+  in (random_auxs, auxs_lhss ~~ auxs_rhss) end;
 
-fun mk_random_datatype config descr vs tycos (names, auxnames) (Ts, Us) thy =
+fun mk_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
   let
     val _ = DatatypeAux.message config "Creating quickcheck generators ...";
     val i = @{term "i\<Colon>code_numeral"};
@@ -329,14 +329,14 @@
           else @{term "max :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"}
             $ HOLogic.mk_number @{typ code_numeral} l $ i
       | NONE => i;
-    val (prefix, (random_auxs, auxs_eqs)) = (apsnd o apsnd o map) mk_prop_eq
+    val (random_auxs, auxs_eqs) = (apsnd o map) mk_prop_eq
       (mk_random_aux_eqs thy descr vs tycos (names, auxnames) (Ts, Us));
     val random_defs = map_index (fn (k, T) => mk_prop_eq
       (HOLogic.mk_random T i, nth random_auxs k $ mk_size_arg k $ i)) Ts;
   in
     thy
     |> TheoryTarget.instantiation (tycos, vs, @{sort random})
-    |> random_aux_specification prefix auxs_eqs
+    |> random_aux_specification prfx random_auxN auxs_eqs
     |> `(fn lthy => map (Syntax.check_term lthy) random_defs)
     |-> (fn random_defs' => fold_map (fn random_def =>
           Specification.definition (NONE, (Attrib.empty_binding,
@@ -359,7 +359,7 @@
   let
     val pp = Syntax.pp_global thy;
     val algebra = Sign.classes_of thy;
-    val (descr, raw_vs, tycos, (names, auxnames), raw_TUs) =
+    val (descr, raw_vs, tycos, prfx, (names, auxnames), raw_TUs) =
       Datatype.the_descr thy raw_tycos;
     val typrep_vs = (map o apsnd)
       (curry (Sorts.inter_sort algebra) @{sort typerep}) raw_vs;
@@ -374,7 +374,7 @@
   in if has_inst then thy
     else case perhaps_constrain thy (random_insts @ term_of_insts) typrep_vs
      of SOME constrain => mk_random_datatype config descr
-          (map constrain typrep_vs) tycos (names, auxnames)
+          (map constrain typrep_vs) tycos prfx (names, auxnames)
             ((pairself o map o map_atyps) (fn TFree v => TFree (constrain v)) raw_TUs) thy
       | NONE => thy
   end;
--- a/src/HOL/Tools/res_atp.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/res_atp.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -11,8 +11,9 @@
   val prepare_clauses : bool -> thm list -> thm list ->
     (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list ->
     (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list -> theory ->
-    ResHolClause.axiom_name vector * (ResHolClause.clause list * ResHolClause.clause list *
-    ResHolClause.clause list * ResClause.classrelClause list * ResClause.arityClause list)
+    ResHolClause.axiom_name vector *
+      (ResHolClause.clause list * ResHolClause.clause list * ResHolClause.clause list *
+      ResHolClause.clause list * ResClause.classrelClause list * ResClause.arityClause list)
 end;
 
 structure ResAtp: RES_ATP =
@@ -296,7 +297,11 @@
 
 (*The rule subsetI is frequently omitted by the relevance filter. This could be theory data
   or identified with ATPset (which however is too big currently)*)
-val whitelist = [subsetI];
+val whitelist_fo = [subsetI];
+(* ext has been a 'helper clause', always given to the atps.
+  As such it was not passed to metis, but metis does not include ext (in contrast
+  to the other helper_clauses *)
+val whitelist_ho = [ResHolClause.ext];
 
 (*** retrieve lemmas from clasimpset and atpset, may filter them ***)
 
@@ -531,10 +536,12 @@
    create additional clauses based on the information from extra_cls *)
 fun prepare_clauses dfg goal_cls chain_ths axcls extra_cls thy =
   let
-    val white_thms =
-      filter check_named (map ResAxioms.pairname (if null chain_ths then whitelist else chain_ths))
+    val isFO = isFO thy goal_cls
+    val white_thms = filter check_named (map ResAxioms.pairname
+      (whitelist_fo @ (if isFO then [] else whitelist_ho) @ chain_ths))
     val white_cls = ResAxioms.cnf_rules_pairs thy white_thms
     val extra_cls = white_cls @ extra_cls
+    val axcls = white_cls @ axcls
     val ccls = subtract_cls goal_cls extra_cls
     val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) ccls
     val ccltms = map prop_of ccls
@@ -544,13 +551,14 @@
     and tycons = type_consts_of_terms thy (ccltms@axtms)
     (*TFrees in conjecture clauses; TVars in axiom clauses*)
     val conjectures = ResHolClause.make_conjecture_clauses dfg thy ccls
+    val (_, extra_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy extra_cls)
     val (clnames,axiom_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy axcls)
-    val helper_clauses = ResHolClause.get_helper_clauses dfg thy (isFO thy goal_cls) (conjectures, extra_cls, [])
+    val helper_clauses = ResHolClause.get_helper_clauses dfg thy isFO (conjectures, extra_cls, [])
     val (supers',arity_clauses) = ResClause.make_arity_clauses_dfg dfg thy tycons supers
     val classrel_clauses = ResClause.make_classrel_clauses thy subs supers'
   in
     (Vector.fromList clnames,
-      (conjectures, axiom_clauses, helper_clauses, classrel_clauses, arity_clauses))
+      (conjectures, axiom_clauses, extra_clauses, helper_clauses, classrel_clauses, arity_clauses))
   end
 
 end;
--- a/src/HOL/Tools/res_axioms.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/res_axioms.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -269,7 +269,7 @@
 (*Converts an Isabelle theorem (intro, elim or simp format, even higher-order) into NNF.*)
 fun to_nnf th ctxt0 =
   let val th1 = th |> transform_elim |> zero_var_indexes
-      val ((_,[th2]),ctxt) = Variable.import_thms true [th1] ctxt0
+      val ((_,[th2]),ctxt) = Variable.import true [th1] ctxt0
       val th3 = th2 |> Conv.fconv_rule ObjectLogic.atomize |> Meson.make_nnf |> strip_lambdas ~1
   in  (th3, ctxt)  end;
 
--- a/src/HOL/Tools/res_clause.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/res_clause.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -57,7 +57,6 @@
   val add_foltype_funcs: fol_type * int Symtab.table -> int Symtab.table
   val add_arityClause_funcs: arityClause * int Symtab.table -> int Symtab.table
   val init_functab: int Symtab.table
-  val writeln_strs: TextIO.outstream -> string list -> unit
   val dfg_sign: bool -> string -> string
   val dfg_of_typeLit: bool -> type_literal -> string
   val gen_dfg_cls: int * string * kind * string list * string list * string list -> string
@@ -441,8 +440,6 @@
 fun string_of_type_clsname (cls_id,ax_name,idx) =
     string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
 
-fun writeln_strs os = List.app (fn s => TextIO.output (os, s));
-
 
 (**** Producing DFG files ****)
 
--- a/src/HOL/Tools/res_hol_clause.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/res_hol_clause.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -12,8 +12,6 @@
   val comb_B: thm
   val comb_C: thm
   val comb_S: thm
-  datatype type_level = T_FULL | T_CONST
-  val typ_level: type_level
   val minimize_applies: bool
   type axiom_name = string
   type polarity = bool
@@ -28,19 +26,23 @@
   val strip_comb: combterm -> combterm * combterm list
   val literals_of_term: theory -> term -> literal list * typ list
   exception TOO_TRIVIAL
-  val make_conjecture_clauses:  bool -> theory -> thm list -> clause list (* dfg thy ccls *)
+  val make_conjecture_clauses:  bool -> theory -> thm list -> clause list
   val make_axiom_clauses: bool ->
        theory ->
-       (thm * (axiom_name * clause_id)) list -> (axiom_name * clause) list (* dfg thy axcls *)
+       (thm * (axiom_name * clause_id)) list -> (axiom_name * clause) list
   val get_helper_clauses: bool ->
        theory ->
        bool ->
        clause list * (thm * (axiom_name * clause_id)) list * string list ->
        clause list
-  val tptp_write_file: bool -> string ->
-    clause list * clause list * clause list * ResClause.classrelClause list * ResClause.arityClause list -> unit
-  val dfg_write_file: bool -> string -> 
-    clause list * clause list * clause list * ResClause.classrelClause list * ResClause.arityClause list -> unit
+  val tptp_write_file: bool -> Path.T ->
+    clause list * clause list * clause list * clause list *
+    ResClause.classrelClause list * ResClause.arityClause list ->
+    int * int
+  val dfg_write_file: bool -> Path.T ->
+    clause list * clause list * clause list * clause list *
+    ResClause.classrelClause list * ResClause.arityClause list ->
+    int * int
 end
 
 structure ResHolClause: RES_HOL_CLAUSE =
@@ -59,10 +61,8 @@
 val equal_imp_fequal = thm "ATP_Linkup.equal_imp_fequal";
 
 
-(*The different translations of types*)
-datatype type_level = T_FULL | T_CONST;
-
-val typ_level = T_CONST;
+(* Parameter t_full below indicates that full type information is to be
+exported *)
 
 (*If true, each function will be directly applied to as many arguments as possible, avoiding
   use of the "apply" operator. Use of hBOOL is also minimized.*)
@@ -260,26 +260,26 @@
 fun wrap_type_if t_full cnh (head, s, tp) =
   if head_needs_hBOOL cnh head then wrap_type t_full (s, tp) else s;
 
-fun string_of_combterm t_full cma cnh t =
+fun string_of_combterm (params as (t_full, cma, cnh)) t =
   let val (head, args) = strip_comb t
   in  wrap_type_if t_full cnh (head,
-                    string_of_applic t_full cma (head, map (string_of_combterm t_full cma cnh) args),
+                    string_of_applic t_full cma (head, map (string_of_combterm (params)) args),
                     type_of_combterm t)
   end;
 
 (*Boolean-valued terms are here converted to literals.*)
-fun boolify t_full cma cnh t =
-  "hBOOL" ^ RC.paren_pack [string_of_combterm t_full cma cnh t];
+fun boolify params t =
+  "hBOOL" ^ RC.paren_pack [string_of_combterm params t];
 
-fun string_of_predicate t_full cma cnh t =
+fun string_of_predicate (params as (_,_,cnh)) t =
   case t of
       (CombApp(CombApp(CombConst("equal",_,_), t1), t2)) =>
           (*DFG only: new TPTP prefers infix equality*)
-          ("equal" ^ RC.paren_pack [string_of_combterm t_full cma cnh t1, string_of_combterm t_full cma cnh t2])
+          ("equal" ^ RC.paren_pack [string_of_combterm params t1, string_of_combterm params t2])
     | _ =>
           case #1 (strip_comb t) of
-              CombConst(c,_,_) => if needs_hBOOL cnh c then boolify t_full cma cnh t else string_of_combterm t_full cma cnh t
-            | _ => boolify t_full cma cnh t;
+              CombConst(c,_,_) => if needs_hBOOL cnh c then boolify params t else string_of_combterm params t
+            | _ => boolify params t;
 
 fun string_of_clausename (cls_id,ax_name) =
     RC.clause_prefix ^ RC.ascii_of ax_name ^ "_" ^ Int.toString cls_id;
@@ -290,23 +290,23 @@
 
 (*** tptp format ***)
 
-fun tptp_of_equality t_full cma cnh pol (t1,t2) =
+fun tptp_of_equality params pol (t1,t2) =
   let val eqop = if pol then " = " else " != "
-  in  string_of_combterm t_full cma cnh t1 ^ eqop ^ string_of_combterm t_full cma cnh t2  end;
+  in  string_of_combterm params t1 ^ eqop ^ string_of_combterm params t2  end;
 
-fun tptp_literal t_full cma cnh (Literal(pol, CombApp(CombApp(CombConst("equal",_,_), t1), t2))) =
-      tptp_of_equality t_full cma cnh pol (t1,t2)
-  | tptp_literal t_full cma cnh (Literal(pol,pred)) =
-      RC.tptp_sign pol (string_of_predicate t_full cma cnh pred);
+fun tptp_literal params (Literal(pol, CombApp(CombApp(CombConst("equal",_,_), t1), t2))) =
+      tptp_of_equality params pol (t1,t2)
+  | tptp_literal params (Literal(pol,pred)) =
+      RC.tptp_sign pol (string_of_predicate params pred);
 
 (*Given a clause, returns its literals paired with a list of literals concerning TFrees;
   the latter should only occur in conjecture clauses.*)
-fun tptp_type_lits t_full cma cnh pos (Clause{literals, ctypes_sorts, ...}) =
-      (map (tptp_literal t_full cma cnh) literals, 
+fun tptp_type_lits params pos (Clause{literals, ctypes_sorts, ...}) =
+      (map (tptp_literal params) literals, 
        map (RC.tptp_of_typeLit pos) (RC.add_typs ctypes_sorts));
 
-fun clause2tptp (t_full, cma, cnh) (cls as Clause{axiom_name,clause_id,kind,ctypes_sorts,...}) =
-  let val (lits,tylits) = tptp_type_lits t_full cma cnh (kind = RC.Conjecture) cls
+fun clause2tptp params (cls as Clause{axiom_name,clause_id,kind,ctypes_sorts,...}) =
+  let val (lits,tylits) = tptp_type_lits params (kind = RC.Conjecture) cls
   in
       (RC.gen_tptp_cls(clause_id,axiom_name,kind,lits,tylits), tylits)
   end;
@@ -314,10 +314,10 @@
 
 (*** dfg format ***)
 
-fun dfg_literal t_full cma cnh (Literal(pol,pred)) = RC.dfg_sign pol (string_of_predicate t_full cma cnh pred);
+fun dfg_literal params (Literal(pol,pred)) = RC.dfg_sign pol (string_of_predicate params pred);
 
-fun dfg_type_lits t_full cma cnh pos (Clause{literals, ctypes_sorts, ...}) =
-      (map (dfg_literal t_full cma cnh) literals, 
+fun dfg_type_lits params pos (Clause{literals, ctypes_sorts, ...}) =
+      (map (dfg_literal params) literals, 
        map (RC.dfg_of_typeLit pos) (RC.add_typs ctypes_sorts));
 
 fun get_uvars (CombConst _) vars = vars
@@ -328,8 +328,8 @@
 
 fun dfg_vars (Clause {literals,...}) = RC.union_all (map get_uvars_l literals);
 
-fun clause2dfg (t_full, cma, cnh) (cls as Clause{axiom_name,clause_id,kind,ctypes_sorts,...}) =
-  let val (lits,tylits) = dfg_type_lits t_full cma cnh (kind = RC.Conjecture) cls
+fun clause2dfg params (cls as Clause{axiom_name,clause_id,kind,ctypes_sorts,...}) =
+  let val (lits,tylits) = dfg_type_lits params (kind = RC.Conjecture) cls
       val vars = dfg_vars cls
       val tvars = RC.get_tvar_strs ctypes_sorts
   in
@@ -341,7 +341,7 @@
 
 fun addtypes tvars tab = List.foldl RC.add_foltype_funcs tab tvars;
 
-fun add_decls t_full cma cnh (CombConst(c,tp,tvars), (funcs,preds)) =
+fun add_decls (t_full, cma, cnh) (CombConst(c,tp,tvars), (funcs,preds)) =
       if c = "equal" then (addtypes tvars funcs, preds)
       else
 	let val arity = min_arity_of cma c
@@ -351,20 +351,20 @@
 	    if needs_hBOOL cnh c then (addtypes tvars (addit funcs), preds)
 	    else (addtypes tvars funcs, addit preds)
 	end
-  | add_decls _ _ _ (CombVar(_,ctp), (funcs,preds)) =
+  | add_decls _ (CombVar(_,ctp), (funcs,preds)) =
       (RC.add_foltype_funcs (ctp,funcs), preds)
-  | add_decls t_full cma cnh (CombApp(P,Q),decls) = add_decls t_full cma cnh (P,add_decls t_full cma cnh (Q,decls));
+  | add_decls params (CombApp(P,Q),decls) = add_decls params (P,add_decls params (Q,decls));
 
-fun add_literal_decls t_full cma cnh (Literal(_,c), decls) = add_decls t_full cma cnh (c,decls);
+fun add_literal_decls params (Literal(_,c), decls) = add_decls params (c,decls);
 
-fun add_clause_decls t_full cma cnh (Clause {literals, ...}, decls) =
-    List.foldl (add_literal_decls t_full cma cnh) decls literals
+fun add_clause_decls params (Clause {literals, ...}, decls) =
+    List.foldl (add_literal_decls params) decls literals
     handle Symtab.DUP a => error ("function " ^ a ^ " has multiple arities")
 
-fun decls_of_clauses (t_full, cma, cnh) clauses arity_clauses =
+fun decls_of_clauses params clauses arity_clauses =
   let val init_functab = Symtab.update (type_wrapper,2) (Symtab.update ("hAPP",2) RC.init_functab)
       val init_predtab = Symtab.update ("hBOOL",1) Symtab.empty
-      val (functab,predtab) = (List.foldl (add_clause_decls t_full cma cnh) (init_functab, init_predtab) clauses)
+      val (functab,predtab) = (List.foldl (add_clause_decls params) (init_functab, init_predtab) clauses)
   in
       (Symtab.dest (List.foldl RC.add_arityClause_funcs functab arity_clauses),
        Symtab.dest predtab)
@@ -427,7 +427,7 @@
         val S = if needed "c_COMBS"
                 then (Output.debug (fn () => "Include combinator S"); cnf_helper_thms thy [comb_S])
                 else []
-        val other = cnf_helper_thms thy [ext,fequal_imp_equal,equal_imp_fequal]
+        val other = cnf_helper_thms thy [fequal_imp_equal,equal_imp_fequal]
     in
         map #2 (make_axiom_clauses dfg thy (other @ IK @ BC @ S))
     end;
@@ -461,11 +461,11 @@
   Output.debug (fn () => "Constant: " ^ c ^ " arity:\t" ^ Int.toString n ^
                 (if needs_hBOOL const_needs_hBOOL c then " needs hBOOL" else ""));
 
-fun count_constants (conjectures, axclauses, helper_clauses, _, _) =
+fun count_constants (conjectures, _, extra_clauses, helper_clauses, _, _) =
   if minimize_applies then
      let val (const_min_arity, const_needs_hBOOL) =
           fold count_constants_clause conjectures (Symtab.empty, Symtab.empty)
-       |> fold count_constants_clause axclauses
+       |> fold count_constants_clause extra_clauses
        |> fold count_constants_clause helper_clauses
      val _ = List.app (display_arity const_needs_hBOOL) (Symtab.dest (const_min_arity))
      in (const_min_arity, const_needs_hBOOL) end
@@ -473,59 +473,62 @@
 
 (* tptp format *)
 
-fun tptp_write_file t_full filename clauses =
+fun tptp_write_file t_full file clauses =
   let
-    val (conjectures, axclauses, helper_clauses, classrel_clauses, arity_clauses) = clauses
+    val (conjectures, axclauses, _, helper_clauses,
+      classrel_clauses, arity_clauses) = clauses
     val (cma, cnh) = count_constants clauses
     val params = (t_full, cma, cnh)
     val (tptp_clss,tfree_litss) = ListPair.unzip (map (clause2tptp params) conjectures)
     val tfree_clss = map RC.tptp_tfree_clause (List.foldl (op union_string) [] tfree_litss)
-    val out = TextIO.openOut filename
-  in
-    List.app (curry TextIO.output out o #1 o (clause2tptp params)) axclauses;
-    RC.writeln_strs out tfree_clss;
-    RC.writeln_strs out tptp_clss;
-    List.app (curry TextIO.output out o RC.tptp_classrelClause) classrel_clauses;
-    List.app (curry TextIO.output out o RC.tptp_arity_clause) arity_clauses;
-    List.app (curry TextIO.output out o #1 o (clause2tptp params)) helper_clauses;
-    TextIO.closeOut out
+    val _ =
+      File.write_list file (
+        map (#1 o (clause2tptp params)) axclauses @
+        tfree_clss @
+        tptp_clss @
+        map RC.tptp_classrelClause classrel_clauses @
+        map RC.tptp_arity_clause arity_clauses @
+        map (#1 o (clause2tptp params)) helper_clauses)
+    in (length axclauses + 1, length tfree_clss + length tptp_clss)
   end;
 
 
 (* dfg format *)
 
-fun dfg_write_file t_full filename clauses =
+fun dfg_write_file t_full file clauses =
   let
-    val (conjectures, axclauses, helper_clauses, classrel_clauses, arity_clauses) = clauses
+    val (conjectures, axclauses, _, helper_clauses,
+      classrel_clauses, arity_clauses) = clauses
     val (cma, cnh) = count_constants clauses
     val params = (t_full, cma, cnh)
     val (dfg_clss, tfree_litss) = ListPair.unzip (map (clause2dfg params) conjectures)
-    and probname = Path.implode (Path.base (Path.explode filename))
+    and probname = Path.implode (Path.base file)
     val axstrs = map (#1 o (clause2dfg params)) axclauses
     val tfree_clss = map RC.dfg_tfree_clause (RC.union_all tfree_litss)
-    val out = TextIO.openOut filename
     val helper_clauses_strs = map (#1 o (clause2dfg params)) helper_clauses
     val (funcs,cl_preds) = decls_of_clauses params (helper_clauses @ conjectures @ axclauses) arity_clauses
     and ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses
-  in
-    TextIO.output (out, RC.string_of_start probname);
-    TextIO.output (out, RC.string_of_descrip probname);
-    TextIO.output (out, RC.string_of_symbols
-                          (RC.string_of_funcs funcs)
-                          (RC.string_of_preds (cl_preds @ ty_preds)));
-    TextIO.output (out, "list_of_clauses(axioms,cnf).\n");
-    RC.writeln_strs out axstrs;
-    List.app (curry TextIO.output out o RC.dfg_classrelClause) classrel_clauses;
-    List.app (curry TextIO.output out o RC.dfg_arity_clause) arity_clauses;
-    RC.writeln_strs out helper_clauses_strs;
-    TextIO.output (out, "end_of_list.\n\nlist_of_clauses(conjectures,cnf).\n");
-    RC.writeln_strs out tfree_clss;
-    RC.writeln_strs out dfg_clss;
-    TextIO.output (out, "end_of_list.\n\n");
-    (*VarWeight=3 helps the HO problems, probably by counteracting the presence of hAPP*)
-    TextIO.output (out, "list_of_settings(SPASS).\n{*\nset_flag(VarWeight,3).\n*}\nend_of_list.\n\n");
-    TextIO.output (out, "end_problem.\n");
-    TextIO.closeOut out
+    val _ =
+      File.write_list file (
+        RC.string_of_start probname ::
+        RC.string_of_descrip probname ::
+        RC.string_of_symbols (RC.string_of_funcs funcs)
+          (RC.string_of_preds (cl_preds @ ty_preds)) ::
+        "list_of_clauses(axioms,cnf).\n" ::
+        axstrs @
+        map RC.dfg_classrelClause classrel_clauses @
+        map RC.dfg_arity_clause arity_clauses @
+        helper_clauses_strs @
+        ["end_of_list.\n\nlist_of_clauses(conjectures,cnf).\n"] @
+        tfree_clss @
+        dfg_clss @
+        ["end_of_list.\n\n",
+        (*VarWeight=3 helps the HO problems, probably by counteracting the presence of hAPP*)
+         "list_of_settings(SPASS).\n{*\nset_flag(VarWeight,3).\n*}\nend_of_list.\n\n",
+         "end_problem.\n"])
+
+    in (length axclauses + length classrel_clauses + length arity_clauses +
+      length helper_clauses + 1, length tfree_clss + length dfg_clss)
   end;
 
 end
--- a/src/HOL/Tools/res_reconstruct.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/res_reconstruct.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -16,10 +16,11 @@
   val setup: Context.theory -> Context.theory
   (* extracting lemma list*)
   val find_failure: string -> string option
-  val lemma_list_dfg: string -> string * string vector * Proof.context * Thm.thm * int -> string
-  val lemma_list_tstp: string -> string * string vector * Proof.context * Thm.thm * int -> string
+  val lemma_list: bool -> string ->
+    string * string vector * (int * int) * Proof.context * Thm.thm * int -> string
   (* structured proofs *)
-  val structured_proof: string -> string * string vector * Proof.context * Thm.thm * int -> string
+  val structured_proof: string ->
+    string * string vector * (int * int) * Proof.context * Thm.thm * int -> string
 end;
 
 structure ResReconstruct : RES_RECONSTRUCT =
@@ -456,23 +457,8 @@
     in  trace msg; msg  end;
 
 
-    
-  (* ==== CHECK IF PROOF OF E OR VAMPIRE WAS SUCCESSFUL === *)
-  
-  val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable",
-    "SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"];
-  val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"];
-  val failure_strings_SPASS = ["SPASS beiseite: Completion found.",
-    "SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."];
-  val failure_strings_remote = ["Remote-script could not extract proof"];
-  fun find_failure proof =
-    let val failures =
-      map_filter (fn s => if String.isSubstring s proof then SOME s else NONE)
-         (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote)
-    in if null failures then NONE else SOME (hd failures) end;
-    
   (*=== EXTRACTING PROOF-TEXT === *)
-  
+
   val begin_proof_strings = ["# SZS output start CNFRefutation.",
       "=========== Refutation ==========",
   "Here is a proof"];
@@ -482,31 +468,49 @@
   fun get_proof_extract proof =
     let
     (*splits to_split by the first possible of a list of splitters*)
-    fun first_field_any [] to_split = NONE
-      | first_field_any (splitter::splitters) to_split =
-      let
-      val result = (first_field splitter to_split)
-      in
-        if isSome result then result else first_field_any splitters to_split
-      end
-    val (a:string, b:string) = valOf (first_field_any begin_proof_strings proof)
-    val (proofextract:string, c:string) = valOf (first_field_any end_proof_strings b)
-    in proofextract end;
-  
+    val (begin_string, end_string) =
+      (find_first (fn s => String.isSubstring s proof) begin_proof_strings,
+      find_first (fn s => String.isSubstring s proof) end_proof_strings)
+    in
+      if is_none begin_string orelse is_none end_string
+      then error "Could not extract proof (no substring indicating a proof)"
+      else proof |> first_field (the begin_string) |> the |> snd
+                 |> first_field (the end_string) |> the |> fst end;
+
+(* ==== CHECK IF PROOF OF E OR VAMPIRE WAS SUCCESSFUL === *)
+
+  val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable",
+    "SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"];
+  val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"];
+  val failure_strings_SPASS = ["SPASS beiseite: Completion found.",
+    "SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."];
+  val failure_strings_remote = ["Remote-script could not extract proof"];
+  fun find_failure proof =
+    let val failures =
+      map_filter (fn s => if String.isSubstring s proof then SOME s else NONE)
+        (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote)
+    val correct = null failures andalso
+      exists (fn s => String.isSubstring s proof) begin_proof_strings andalso
+      exists (fn s => String.isSubstring s proof) end_proof_strings
+    in
+      if correct then NONE
+      else if null failures then SOME "Output of ATP not in proper format"
+      else SOME (hd failures) end;
+
   (* === EXTRACTING LEMMAS === *)
   (* lines have the form "cnf(108, axiom, ...",
   the number (108) has to be extracted)*)
-  fun get_step_nums_tstp proofextract =
+  fun get_step_nums false proofextract =
     let val toks = String.tokens (not o Char.isAlphaNum)
     fun inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok
+      | inputno ("cnf"::ntok::"negated"::"conjecture"::_) = Int.fromString ntok
       | inputno _ = NONE
     val lines = split_lines proofextract
     in  List.mapPartial (inputno o toks) lines  end
-    
-    (*String contains multiple lines. We want those of the form
+  (*String contains multiple lines. We want those of the form
     "253[0:Inp] et cetera..."
     A list consisting of the first number in each line is returned. *)
-  fun get_step_nums_dfg proofextract =
+  |  get_step_nums true proofextract =
     let val toks = String.tokens (not o Char.isAlphaNum)
     fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok
       | inputno _ = NONE
@@ -514,15 +518,19 @@
     in  List.mapPartial (inputno o toks) lines  end
     
   (*extracting lemmas from tstp-output between the lines from above*)
-  fun extract_lemmas get_step_nums (proof, thm_names, _, _, _) =
+  fun extract_lemmas get_step_nums (proof, thm_names, conj_count, _, _, _) =
     let
     (* get the names of axioms from their numbers*)
     fun get_axiom_names thm_names step_nums =
       let
-      fun is_axiom n = n <= Vector.length thm_names
+      val last_axiom = Vector.length thm_names
+      fun is_axiom n = n <= last_axiom
+      fun is_conj n = n >= #1 conj_count andalso n < #1 conj_count + #2 conj_count
       fun getname i = Vector.sub(thm_names, i-1)
       in
-        sort_distinct string_ord (filter (fn x => x <> "??.unknown") (map getname (filter is_axiom step_nums)))
+        (sort_distinct string_ord (filter (fn x => x <> "??.unknown")
+          (map getname (filter is_axiom step_nums))),
+        exists is_conj step_nums)
       end
     val proofextract = get_proof_extract proof
     in
@@ -545,22 +553,23 @@
 
   fun sendback_metis_nochained lemmas =
     (Markup.markup Markup.sendback o metis_line) (nochained lemmas)
-  fun lemma_list_tstp name result =
-    let val lemmas = extract_lemmas get_step_nums_tstp result
-    in sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas end;
-  fun lemma_list_dfg name result =
-    let val lemmas = extract_lemmas get_step_nums_dfg result
-    in sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas end;
+
+  fun lemma_list dfg name result =
+    let val (lemmas, used_conj) = extract_lemmas (get_step_nums dfg) result
+    in sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^
+      (if used_conj then ""
+       else "\nWarning: Goal is provable because context is inconsistent.")
+    end;
 
   (* === Extracting structured Isar-proof === *)
-  fun structured_proof name (result as (proof, thm_names, ctxt, goal, subgoalno)) =
+  fun structured_proof name (result as (proof, thm_names, conj_count, ctxt, goal, subgoalno)) =
     let
     (*Could use split_lines, but it can return blank lines...*)
     val lines = String.tokens (equal #"\n");
     val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c)
     val proofextract = get_proof_extract proof
     val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
-    val one_line_proof = lemma_list_tstp name result
+    val one_line_proof = lemma_list false name result
     val structured = if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then ""
                 else decode_tstp_file cnfs ctxt goal subgoalno thm_names
     in
--- a/src/HOL/Tools/transfer_data.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Tools/transfer_data.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -54,7 +54,7 @@
 
 fun basic_transfer_rule injonly a0 D0 e leave ctxt0 th = 
  let 
-  val ([a,D], ctxt) = apfst (map Drule.dest_term o snd) (Variable.import_thms true (map Drule.mk_term [a0, D0]) ctxt0)
+  val ([a,D], ctxt) = apfst (map Drule.dest_term o snd) (Variable.import true (map Drule.mk_term [a0, D0]) ctxt0)
   val (aT,bT) = 
      let val T = typ_of (ctyp_of_term a) 
      in (Term.range_type T, Term.domain_type T)
--- a/src/HOL/Transcendental.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/Transcendental.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -784,9 +784,8 @@
 	  also have "\<dots> = \<bar>f n * real (Suc n) * R' ^ n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult real_mult_assoc[symmetric] by algebra
 	  finally show ?thesis .
 	qed }
-      { fix n
-	from DERIV_pow[of "Suc n" x0, THEN DERIV_cmult[where c="f n"]]
-	show "DERIV (\<lambda> x. ?f x n) x0 :> (?f' x0 n)" unfolding real_mult_assoc by auto }
+      { fix n show "DERIV (\<lambda> x. ?f x n) x0 :> (?f' x0 n)"
+	  by (auto intro!: DERIV_intros simp del: power_Suc) }
       { fix x assume "x \<in> {-R' <..< R'}" hence "R' \<in> {-R <..< R}" and "norm x < norm R'" using assms `R' < R` by auto
 	have "summable (\<lambda> n. f n * x^n)"
 	proof (rule summable_le2[THEN conjunct1, OF _ powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`]], rule allI)
@@ -1362,6 +1361,12 @@
 by (rule DERIV_cos [THEN DERIV_isCont])
 
 
+declare
+  DERIV_exp[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
+  DERIV_ln[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
+  DERIV_sin[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
+  DERIV_cos[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
+
 subsection {* Properties of Sine and Cosine *}
 
 lemma sin_zero [simp]: "sin 0 = 0"
@@ -1410,24 +1415,17 @@
 by auto
 
 lemma DERIV_cos_realpow2b: "DERIV (%x. (cos x)\<twosuperior>) x :> -(2 * cos(x) * sin(x))"
-apply (rule lemma_DERIV_subst)
-apply (rule DERIV_cos_realpow2a, auto)
-done
+  by (auto intro!: DERIV_intros)
 
 (* most useful *)
 lemma DERIV_cos_cos_mult3 [simp]:
      "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))"
-apply (rule lemma_DERIV_subst)
-apply (rule DERIV_cos_cos_mult2, auto)
-done
+  by (auto intro!: DERIV_intros)
 
 lemma DERIV_sin_circle_all: 
      "\<forall>x. DERIV (%x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :>  
              (2*cos(x)*sin(x) - 2*cos(x)*sin(x))"
-apply (simp only: diff_minus, safe)
-apply (rule DERIV_add) 
-apply (auto simp add: numeral_2_eq_2)
-done
+  by (auto intro!: DERIV_intros)
 
 lemma DERIV_sin_circle_all_zero [simp]:
      "\<forall>x. DERIV (%x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :> 0"
@@ -1513,22 +1511,12 @@
 apply (rule DERIV_cos, auto)
 done
 
-lemmas DERIV_intros = DERIV_ident DERIV_const DERIV_cos DERIV_cmult 
-                    DERIV_sin  DERIV_exp  DERIV_inverse DERIV_pow 
-                    DERIV_add  DERIV_diff  DERIV_mult  DERIV_minus 
-                    DERIV_inverse_fun DERIV_quotient DERIV_fun_pow 
-                    DERIV_fun_exp DERIV_fun_sin DERIV_fun_cos 
-
 (* lemma *)
 lemma lemma_DERIV_sin_cos_add:
      "\<forall>x.  
          DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +  
                (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0"
-apply (safe, rule lemma_DERIV_subst)
-apply (best intro!: DERIV_intros intro: DERIV_chain2) 
-  --{*replaces the old @{text DERIV_tac}*}
-apply (auto simp add: algebra_simps)
-done
+  by (auto intro!: DERIV_intros simp add: algebra_simps)
 
 lemma sin_cos_add [simp]:
      "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +  
@@ -1550,10 +1538,8 @@
 
 lemma lemma_DERIV_sin_cos_minus:
     "\<forall>x. DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0"
-apply (safe, rule lemma_DERIV_subst)
-apply (best intro!: DERIV_intros intro: DERIV_chain2)
-apply (simp add: algebra_simps)
-done
+  by (auto intro!: DERIV_intros simp add: algebra_simps)
+
 
 lemma sin_cos_minus: 
     "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0"
@@ -1722,7 +1708,7 @@
 apply (assumption, rule_tac y=y in order_less_le_trans, simp_all) 
 apply (drule_tac y1 = y in order_le_less_trans [THEN sin_gt_zero], assumption, simp_all) 
 done
-    
+
 lemma pi_half: "pi/2 = (THE x. 0 \<le> x & x \<le> 2 & cos x = 0)"
 by (simp add: pi_def)
 
@@ -2121,10 +2107,7 @@
 
 lemma lemma_DERIV_tan:
      "cos x \<noteq> 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\<twosuperior>)"
-apply (rule lemma_DERIV_subst)
-apply (best intro!: DERIV_intros intro: DERIV_chain2) 
-apply (auto simp add: divide_inverse numeral_2_eq_2)
-done
+  by (auto intro!: DERIV_intros simp add: field_simps numeral_2_eq_2)
 
 lemma DERIV_tan [simp]: "cos x \<noteq> 0 ==> DERIV tan x :> inverse((cos x)\<twosuperior>)"
 by (auto dest: lemma_DERIV_tan simp add: tan_def [symmetric])
@@ -2500,6 +2483,11 @@
 apply (simp, simp, simp, rule isCont_arctan)
 done
 
+declare
+  DERIV_arcsin[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
+  DERIV_arccos[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
+  DERIV_arctan[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
+
 subsection {* More Theorems about Sin and Cos *}
 
 lemma cos_45: "cos (pi / 4) = sqrt 2 / 2"
@@ -2589,11 +2577,7 @@
 by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto)
 
 lemma DERIV_sin_add [simp]: "DERIV (%x. sin (x + k)) xa :> cos (xa + k)"
-apply (rule lemma_DERIV_subst)
-apply (rule_tac f = sin and g = "%x. x + k" in DERIV_chain2)
-apply (best intro!: DERIV_intros intro: DERIV_chain2)+
-apply (simp (no_asm))
-done
+  by (auto intro!: DERIV_intros)
 
 lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n"
 proof -
@@ -2634,11 +2618,7 @@
 by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib add_divide_distrib, auto)
 
 lemma DERIV_cos_add [simp]: "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)"
-apply (rule lemma_DERIV_subst)
-apply (rule_tac f = cos and g = "%x. x + k" in DERIV_chain2)
-apply (best intro!: DERIV_intros intro: DERIV_chain2)+
-apply (simp (no_asm))
-done
+  by (auto intro!: DERIV_intros)
 
 lemma sin_zero_abs_cos_one: "sin x = 0 ==> \<bar>cos x\<bar> = 1"
 by (auto simp add: sin_zero_iff even_mult_two_ex)
--- a/src/HOL/ex/Codegenerator_Candidates.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/ex/Codegenerator_Candidates.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -8,6 +8,7 @@
   Complex_Main
   AssocList
   Binomial
+  Fset
   Commutative_Ring
   Enum
   List_Prefix
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Wed Jul 01 16:19:44 2009 +0100
@@ -58,7 +58,7 @@
 
 lemma [code_pred_intros]:
 "r a b ==> tranclp r a b"
-"r a b ==> tranclp r b c ==> tranclp r a c" 
+"r a b ==> tranclp r b c ==> tranclp r a c"
 by auto
 
 (* Setup requires quick and dirty proof *)
@@ -71,7 +71,6 @@
 
 thm tranclp.equation
 *)
-
 inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
     "succ 0 1"
   | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
--- a/src/HOL/ex/predicate_compile.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/HOL/ex/predicate_compile.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -9,6 +9,8 @@
   type mode = int list option list * int list
   val add_equations_of: string list -> theory -> theory
   val register_predicate : (thm list * thm * int) -> theory -> theory
+  val is_registered : theory -> string -> bool
+  val fetch_pred_data : theory -> string -> (thm list * thm * int)  
   val predfun_intro_of: theory -> string -> mode -> thm
   val predfun_elim_of: theory -> string -> mode -> thm
   val strip_intro_concl: int -> term -> term * (term list * term list)
@@ -17,14 +19,18 @@
   val modes_of: theory -> string -> mode list
   val intros_of: theory -> string -> thm list
   val nparams_of: theory -> string -> int
+  val add_intro: thm -> theory -> theory
+  val set_elim: thm -> theory -> theory
   val setup: theory -> theory
   val code_pred: string -> Proof.context -> Proof.state
   val code_pred_cmd: string -> Proof.context -> Proof.state
-(*  val print_alternative_rules: theory -> theory (*FIXME diagnostic command?*) *)
+  val print_stored_rules: theory -> unit
   val do_proofs: bool ref
+  val mk_casesrule : Proof.context -> int -> thm list -> term
   val analyze_compr: theory -> term -> term
   val eval_ref: (unit -> term Predicate.pred) option ref
   val add_equations : string -> theory -> theory
+  val code_pred_intros_attrib : attribute
 end;
 
 structure Predicate_Compile : PREDICATE_COMPILE =
@@ -187,7 +193,7 @@
  of NONE => error ("No such predicate " ^ quote name)  
   | SOME data => data;
 
-val is_pred = is_some oo lookup_pred_data 
+val is_registered = is_some oo lookup_pred_data 
 
 val all_preds_of = Graph.keys o PredData.get
 
@@ -223,25 +229,26 @@
 
 val predfun_elim_of = #elim ooo the_predfun_data
 
+fun print_stored_rules thy =
+  let
+    val preds = (Graph.keys o PredData.get) thy
+    fun print pred () = let
+      val _ = writeln ("predicate: " ^ pred)
+      val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred))
+      val _ = writeln ("introrules: ")
+      val _ = fold (fn thm => fn u =>  writeln (Display.string_of_thm thm))
+        (rev (intros_of thy pred)) ()
+    in
+      if (has_elim thy pred) then
+        writeln ("elimrule: " ^ Display.string_of_thm (the_elim_of thy pred))
+      else
+        writeln ("no elimrule defined")
+    end
+  in
+    fold print preds ()
+  end;
 
-(* replaces print_alternative_rules *)
-(* TODO:
-fun print_alternative_rules thy = let
-    val d = IndCodegenData.get thy
-    val preds = (Symtab.keys (#intro_rules d)) union (Symtab.keys (#elim_rules d))
-    val _ = tracing ("preds: " ^ (makestring preds))
-    fun print pred = let
-      val _ = tracing ("predicate: " ^ pred)
-      val _ = tracing ("introrules: ")
-      val _ = fold (fn thm => fn u => tracing (makestring thm))
-        (rev (Symtab.lookup_list (#intro_rules d) pred)) ()
-      val _ = tracing ("casesrule: ")
-      val _ = tracing (makestring (Symtab.lookup (#elim_rules d) pred))
-    in () end
-    val _ = map print preds
- in thy end;
-*)
-
+(** preprocessing rules **)  
 
 fun imp_prems_conv cv ct =
   case Thm.term_of ct of
@@ -298,6 +305,23 @@
     val add = apsnd (cons (mode, mk_predfun_data data))
   in PredData.map (Graph.map_node name (map_pred_data add)) end
 
+fun is_inductive_predicate thy name =
+  is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
+
+fun depending_preds_of thy intros = fold Term.add_consts (map Thm.prop_of intros) [] |> map fst
+    |> filter (fn c => is_inductive_predicate thy c orelse is_registered thy c)
+
+(* code dependency graph *)    
+fun dependencies_of thy name =
+  let
+    val (intros, elim, nparams) = fetch_pred_data thy name 
+    val data = mk_pred_data ((intros, SOME elim, nparams), [])
+    val keys = depending_preds_of thy intros
+  in
+    (data, keys)
+  end;
+
+(* TODO: add_edges - by analysing dependencies *)
 fun add_intro thm thy = let
    val (name, _) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
    fun cons_intro gr =
@@ -320,10 +344,13 @@
     fun set (intros, elim, _ ) = (intros, elim, nparams) 
   in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
 
-fun register_predicate (intros, elim, nparams) = let
+fun register_predicate (intros, elim, nparams) thy = let
     val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd intros))))
     fun set _ = (intros, SOME elim, nparams)
-  in PredData.map (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), []))) end
+  in
+    PredData.map (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), []))
+      #> fold Graph.add_edge (map (pair name) (depending_preds_of thy intros))) thy
+  end
 
   
 (* Mode analysis *)
@@ -625,10 +652,11 @@
 
 and compile_param thy modes (NONE, t) = t
  | compile_param thy modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
-   (case t of
-     Abs _ => compile_param_ext thy modes (m, t)
-   |  _ => let
-     val (f, args) = strip_comb t
+   (* (case t of
+     Abs _ => error "compile_param: Invalid term" *) (* compile_param_ext thy modes (m, t) *)
+   (*  |  _ => let *)
+   let  
+     val (f, args) = strip_comb (Envir.eta_contract t)
      val (params, args') = chop (length ms) args
      val params' = map (compile_param thy modes) (ms ~~ params)
      val f' = case f of
@@ -637,7 +665,7 @@
              Const (predfun_name_of thy name (iss, is'), funT'_of (iss, is') T)
           else error "compile param: Not an inductive predicate with correct mode"
       | Free (name, T) => Free (name, funT_of T (SOME is'))
-   in list_comb (f', params' @ args') end)
+   in list_comb (f', params' @ args') end
  | compile_param _ _ _ = error "compile params"
   
   
@@ -1116,7 +1144,7 @@
 (* VERY LARGE SIMILIRATIY to function prove_param 
 -- join both functions
 *)
-(*
+
 fun prove_param2 thy modes (NONE, t) = all_tac 
   | prove_param2 thy modes (m as SOME (Mode (mode, is, ms)), t) = let
     val  (f, args) = strip_comb t
@@ -1132,7 +1160,7 @@
     THEN print_tac "after simplification in prove_args"
     THEN (EVERY (map (prove_param2 thy modes) (ms ~~ params)))
   end
-*)
+
 
 fun prove_expr2 thy modes (SOME (Mode (mode, is, ms)), t) = 
   (case strip_comb t of
@@ -1140,14 +1168,16 @@
       if AList.defined op = modes name then
         etac @{thm bindE} 1
         THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
+        THEN new_print_tac "prove_expr2-before"
         THEN (debug_tac (Syntax.string_of_term_global thy
           (prop_of (predfun_elim_of thy name mode))))
         THEN (etac (predfun_elim_of thy name mode) 1)
         THEN new_print_tac "prove_expr2"
         (* TODO -- FIXME: replace remove_last_goal*)
-        THEN (EVERY (replicate (length args) (remove_last_goal thy)))
+        (* THEN (EVERY (replicate (length args) (remove_last_goal thy))) *)
+        THEN (EVERY (map (prove_param thy modes) (ms ~~ args)))
         THEN new_print_tac "finished prove_expr2"
-        (* THEN (EVERY (map (prove_param thy modes) (ms ~~ args))) *)
+      
       else error "Prove expr2 if case not implemented"
     | _ => etac @{thm bindE} 1)
   | prove_expr2 _ _ _ = error "Prove expr2 not implemented"
@@ -1243,7 +1273,7 @@
     end;
   val prems_tac = prove_prems2 in_ts' param_vs ps 
 in
-  print_tac "starting prove_clause2"
+  new_print_tac "starting prove_clause2"
   THEN etac @{thm bindE} 1
   THEN (etac @{thm singleE'} 1)
   THEN (TRY (etac @{thm Pair_inject} 1))
@@ -1259,6 +1289,7 @@
   (DETERM (TRY (rtac @{thm unit.induct} 1)))
    THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
    THEN (rtac (predfun_intro_of thy pred mode) 1)
+   THEN (REPEAT_DETERM (rtac @{thm refl} 2))
    THEN (EVERY (map prove_clause (clauses ~~ (1 upto (length clauses)))))
 end;
 
@@ -1321,7 +1352,7 @@
         | Negprem _ => error ("Double negation not allowed in premise: " ^ (Syntax.string_of_term_global thy (c $ t))) 
         | Sidecond t => Sidecond (c $ t))
       | (c as Const (s, _), ts) =>
-        if is_pred thy s then
+        if is_registered thy s then
           let val (ts1, ts2) = chop (nparams_of thy s) ts
           in Prem (ts2, list_comb (c, ts1)) end
         else Sidecond t
@@ -1373,7 +1404,7 @@
   val _ = map (Output.tracing o (Syntax.string_of_term_global thy')) (flat ts)
   val pred_mode =
     maps (fn (s, (T, _)) => map (pair (s, T)) ((the o AList.lookup (op =) modes) s)) clauses' 
-  val _ = tracing "Proving equations..."
+  val _ = Output.tracing "Proving equations..."
   val result_thms =
     prove_preds thy' all_vs param_vs (extra_modes @ modes) clauses (pred_mode ~~ (flat ts))
   val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
@@ -1409,21 +1440,6 @@
     val cases = map mk_case intros
   in Logic.list_implies (assm :: cases, prop) end;
 
-(* code dependency graph *)
-
-fun dependencies_of thy name =
-  let
-    fun is_inductive_predicate thy name =
-      is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
-    val (intro, elim, nparams) = fetch_pred_data thy name 
-    val data = mk_pred_data ((intro, SOME elim, nparams), [])
-    val intros = map Thm.prop_of (#intros (rep_pred_data data))
-    val keys = fold Term.add_consts intros [] |> map fst
-    |> filter (is_inductive_predicate thy)
-  in
-    (data, keys)
-  end;
-
 fun add_equations name thy =
   let
     val thy' = PredData.map (Graph.extend (dependencies_of thy) name) thy |> Theory.checkpoint;
@@ -1437,17 +1453,15 @@
       scc thy' |> Theory.checkpoint
   in thy'' end
 
+  
+fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
+
+val code_pred_intros_attrib = attrib add_intro;
+
 (** user interface **)
 
 local
 
-fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
-
-(*
-val add_elim_attrib = attrib set_elim;
-*)
-
-
 (* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
 (* TODO: must create state to prove multiple cases *)
 fun generic_code_pred prep_const raw_const lthy =
--- a/src/Pure/General/file.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/General/file.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -12,7 +12,7 @@
   val pwd: unit -> Path.T
   val full_path: Path.T -> Path.T
   val tmp_path: Path.T -> Path.T
-  val isabelle_tool: string -> int
+  val isabelle_tool: string -> string -> int
   val system_command: string -> unit
   eqtype ident
   val rep_ident: ident -> string
@@ -65,7 +65,17 @@
 
 (* system commands *)
 
-fun isabelle_tool cmd = system ("\"$ISABELLE_TOOL\" " ^ cmd);
+fun isabelle_tool name args =
+  (case space_explode ":" (getenv "ISABELLE_TOOLS") |> get_first (fn dir =>
+      let val path = dir ^ "/" ^ name in
+        if can OS.FileSys.modTime path andalso
+          not (OS.FileSys.isDir path) andalso
+          OS.FileSys.access (path, [OS.FileSys.A_READ, OS.FileSys.A_EXEC])
+        then SOME path
+        else NONE
+      end handle OS.SysErr _ => NONE) of
+    SOME path => system (shell_quote path ^ " " ^ args)
+  | NONE => (writeln ("Unknown Isabelle tool: " ^ name); 2));
 
 fun system_command cmd =
   if system cmd <> 0 then error ("System command failed: " ^ cmd)
--- a/src/Pure/General/swing.scala	Wed Jun 24 16:28:30 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-/*  Title:      Pure/General/swing.scala
-    Author:     Makarius
-
-Swing utilities.
-*/
-
-package isabelle
-
-import javax.swing.SwingUtilities
-
-object Swing
-{
-  def now[A](body: => A): A = {
-    var result: Option[A] = None
-    if (SwingUtilities.isEventDispatchThread) { result = Some(body) }
-    else SwingUtilities.invokeAndWait(new Runnable { def run = { result = Some(body) } })
-    result.get
-  }
-
-  def later(body: => Unit) {
-    if (SwingUtilities.isEventDispatchThread) body
-    else SwingUtilities.invokeLater(new Runnable { def run = body })
-  }
-}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/swing_thread.scala	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,24 @@
+/*  Title:      Pure/General/swing_thread.scala
+    Author:     Makarius
+
+Evaluation within the AWT/Swing thread.
+*/
+
+package isabelle
+
+import javax.swing.SwingUtilities
+
+object Swing_Thread
+{
+  def now[A](body: => A): A = {
+    var result: Option[A] = None
+    if (SwingUtilities.isEventDispatchThread) { result = Some(body) }
+    else SwingUtilities.invokeAndWait(new Runnable { def run = { result = Some(body) } })
+    result.get
+  }
+
+  def later(body: => Unit) {
+    if (SwingUtilities.isEventDispatchThread) body
+    else SwingUtilities.invokeLater(new Runnable { def run = body })
+  }
+}
--- a/src/Pure/IsaMakefile	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/IsaMakefile	Wed Jul 01 16:19:44 2009 +0100
@@ -118,26 +118,32 @@
 ## Scala material
 
 SCALA_FILES = General/event_bus.scala General/markup.scala		\
-  General/position.scala General/scan.scala General/swing.scala		\
+  General/position.scala General/scan.scala General/swing_thread.scala	\
   General/symbol.scala General/xml.scala General/yxml.scala		\
   Isar/isar.scala Isar/isar_document.scala Isar/outer_keyword.scala	\
-  System/cygwin.scala System/isabelle_process.scala			\
-  System/isabelle_system.scala Thy/completion.scala			\
-  Thy/thy_header.scala Tools/isabelle_syntax.scala
+  System/cygwin.scala System/gui_setup.scala				\
+  System/isabelle_process.scala System/isabelle_system.scala		\
+  System/platform.scala Thy/completion.scala Thy/thy_header.scala	\
+  Tools/isabelle_syntax.scala
 
 
-SCALA_TARGET = $(ISABELLE_HOME)/lib/classes/Pure.jar
+JAR_DIR = $(ISABELLE_HOME)/lib/classes
+PURE_JAR = $(JAR_DIR)/Pure.jar
+FULL_JAR = $(JAR_DIR)/isabelle-scala.jar
 
-jar: $(SCALA_TARGET)
+jars: $(FULL_JAR)
 
-$(SCALA_TARGET): $(SCALA_FILES)
+$(FULL_JAR): $(SCALA_FILES)
 	@rm -rf classes && mkdir classes
-	scalac -deprecation -d classes -target jvm-1.5 $(SCALA_FILES)
-	scaladoc -d classes $(SCALA_FILES)
+	"$(SCALA_HOME)/bin/scalac" -deprecation -d classes -target jvm-1.5 $(SCALA_FILES)
+	"$(SCALA_HOME)/bin/scaladoc" -d classes $(SCALA_FILES)
 	@cp $(SCALA_FILES) classes/isabelle
-	@mkdir -p `dirname $@`
-	@cd classes; jar cf `jvmpath $@` isabelle
+	@mkdir -p "$(JAR_DIR)"
+	@cd classes; jar cfe `jvmpath "$(PURE_JAR)"` isabelle.GUI_Setup isabelle
+	@cd classes; cp "$(SCALA_HOME)/lib/scala-swing.jar" .; jar xf scala-swing.jar; \
+          cp "$(SCALA_HOME)/lib/scala-library.jar" "$(FULL_JAR)"; \
+          jar ufe `jvmpath $(FULL_JAR)` isabelle.GUI_Setup isabelle scala
 	@rm -rf classes
 
-clean-jar:
-	@rm -f $(SCALA_TARGET)
+clean-jars:
+	@rm -f "$(PURE_JAR)" "$(FULL_JAR)"
--- a/src/Pure/Isar/attrib.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/attrib.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -265,7 +265,7 @@
 val no_vars = Thm.rule_attribute (fn context => fn th =>
   let
     val ctxt = Variable.set_body false (Context.proof_of context);
-    val ((_, [th']), _) = Variable.import_thms true [th] ctxt;
+    val ((_, [th']), _) = Variable.import true [th] ctxt;
   in th' end);
 
 val eta_long =
--- a/src/Pure/Isar/class.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/class.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -75,7 +75,7 @@
     (* assm_intro *)
     fun prove_assm_intro thm =
       let
-        val ((_, [thm']), _) = Variable.import_thms true [thm] empty_ctxt;
+        val ((_, [thm']), _) = Variable.import true [thm] empty_ctxt;
         val thm'' = Morphism.thm (const_morph $> eq_morph) thm';
         val tac = ALLGOALS (ProofContext.fact_tac [thm'']);
       in SkipProof.prove_global thy [] [] (Thm.prop_of thm'') (K tac) end;
--- a/src/Pure/Isar/class_target.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/class_target.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -32,6 +32,7 @@
   (*instances*)
   val init_instantiation: string list * (string * sort) list * sort
     -> theory -> local_theory
+  val instance_arity_cmd: xstring list * xstring list * xstring -> theory -> Proof.state
   val instantiation_instance: (local_theory -> local_theory)
     -> local_theory -> Proof.state
   val prove_instantiation_instance: (Proof.context -> tactic)
@@ -44,7 +45,8 @@
   val instantiation_param: local_theory -> binding -> string option
   val confirm_declaration: binding -> local_theory -> local_theory
   val pretty_instantiation: local_theory -> Pretty.T
-  val instance_arity_cmd: xstring * xstring list * xstring -> theory -> Proof.state
+  val read_multi_arity: theory -> xstring list * xstring list * xstring
+    -> string list * (string * sort) list * sort
   val type_name: string -> string
 
   (*subclasses*)
@@ -419,6 +421,15 @@
   |> find_first (fn (_, (v, _)) => Binding.name_of b = v)
   |> Option.map (fst o fst);
 
+fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
+  let
+    val all_arities = map (fn raw_tyco => Sign.read_arity thy
+      (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
+    val tycos = map #1 all_arities;
+    val (_, sorts, sort) = hd all_arities;
+    val vs = Name.names Name.context Name.aT sorts;
+  in (tycos, vs, sort) end;
+
 
 (* syntax *)
 
@@ -578,15 +589,17 @@
 
 (* simplified instantiation interface with no class parameter *)
 
-fun instance_arity_cmd arities thy =
+fun instance_arity_cmd raw_arities thy =
   let
+    val (tycos, vs, sort) = read_multi_arity thy raw_arities;
+    val sorts = map snd vs;
+    val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos;
     fun after_qed results = ProofContext.theory
       ((fold o fold) AxClass.add_arity results);
   in
     thy
     |> ProofContext.init
-    |> Proof.theorem_i NONE after_qed ((map (fn t => [(t, [])])
-        o Logic.mk_arities o Sign.read_arity thy) arities)
+    |> Proof.theorem_i NONE after_qed (map (fn t => [(t, [])]) arities)
   end;
 
 
--- a/src/Pure/Isar/code.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/code.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -75,7 +75,7 @@
   val these_eqns: theory -> string -> (thm * bool) list
   val all_eqns: theory -> (thm * bool) list
   val get_case_scheme: theory -> string -> (int * (int * string list)) option
-  val is_undefined: theory -> string -> bool
+  val undefineds: theory -> string list
   val print_codesetup: theory -> unit
 end;
 
@@ -898,7 +898,7 @@
 
 fun get_case_scheme thy = Symtab.lookup ((fst o the_cases o the_exec) thy);
 
-val is_undefined = Symtab.defined o snd o the_cases o the_exec;
+val undefineds = Symtab.keys o snd o the_cases o the_exec;
 
 structure TypeInterpretation = InterpretationFun(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
 
--- a/src/Pure/Isar/element.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/element.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -225,7 +225,7 @@
     val th = MetaSimplifier.norm_hhf raw_th;
     val is_elim = ObjectLogic.is_elim th;
 
-    val ((_, [th']), ctxt') = Variable.import_thms true [th] (Variable.set_body false ctxt);
+    val ((_, [th']), ctxt') = Variable.import true [th] (Variable.set_body false ctxt);
     val prop = Thm.prop_of th';
     val (prems, concl) = Logic.strip_horn prop;
     val concl_term = ObjectLogic.drop_judgment thy concl;
--- a/src/Pure/Isar/isar.scala	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/isar.scala	Wed Jul 01 16:19:44 2009 +0100
@@ -7,8 +7,9 @@
 package isabelle
 
 
-class Isar(isabelle_system: IsabelleSystem, results: EventBus[IsabelleProcess.Result], args: String*)
-  extends IsabelleProcess(isabelle_system, results, args: _*)
+class Isar(isabelle_system: Isabelle_System,
+    results: EventBus[Isabelle_Process.Result], args: String*)
+  extends Isabelle_Process(isabelle_system, results, args: _*)
 {
   /* basic editor commands */
 
--- a/src/Pure/Isar/isar_cmd.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/isar_cmd.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -306,11 +306,11 @@
 
 fun display_drafts files = Toplevel.imperative (fn () =>
   let val outfile = File.shell_path (Present.drafts (getenv "ISABELLE_DOC_FORMAT") files)
-  in File.isabelle_tool ("display -c " ^ outfile ^ " &"); () end);
+  in File.isabelle_tool "display" ("-c " ^ outfile ^ " &"); () end);
 
 fun print_drafts files = Toplevel.imperative (fn () =>
   let val outfile = File.shell_path (Present.drafts "ps" files)
-  in File.isabelle_tool ("print -c " ^ outfile); () end);
+  in File.isabelle_tool "print" ("-c " ^ outfile); () end);
 
 
 (* pretty_setmargin *)
--- a/src/Pure/Isar/isar_document.scala	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/isar_document.scala	Wed Jul 01 16:19:44 2009 +0100
@@ -14,7 +14,7 @@
   type Document_ID = String
 }
 
-trait IsarDocument extends IsabelleProcess
+trait IsarDocument extends Isabelle_Process
 {
   import IsarDocument._
 
--- a/src/Pure/Isar/isar_syn.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -465,7 +465,7 @@
 val _ =
   OuterSyntax.command "instance" "prove type arity or subclass relation" K.thy_goal
   ((P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname) >> Class.classrel_cmd ||
-    P.arity >> Class.instance_arity_cmd)
+    P.multi_arity >> Class.instance_arity_cmd)
     >> (Toplevel.print oo Toplevel.theory_to_proof)
   || Scan.succeed (Toplevel.print o Toplevel.local_theory_to_proof NONE (Class.instantiation_instance I)));
 
--- a/src/Pure/Isar/local_defs.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/local_defs.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -170,7 +170,7 @@
       (case filter_out is_trivial raw_eqs of
         [] => th
       | eqs =>
-          let val ((_, th' :: eqs'), ctxt') = Variable.import_thms true (th :: eqs) ctxt
+          let val ((_, th' :: eqs'), ctxt') = Variable.import true (th :: eqs) ctxt
           in singleton (Variable.export ctxt' ctxt) (fold trans eqs' th') end);
 
 in
--- a/src/Pure/Isar/obtain.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/obtain.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -79,7 +79,7 @@
     val _ = ObjectLogic.is_judgment thy (Thm.concl_of thm) orelse
       error "Conclusion in obtained context must be object-logic judgment";
 
-    val ((_, [thm']), ctxt') = Variable.import_thms true [thm] fix_ctxt;
+    val ((_, [thm']), ctxt') = Variable.import true [thm] fix_ctxt;
     val prems = Drule.strip_imp_prems (#prop (Thm.crep_thm thm'));
   in
     ((Drule.implies_elim_list thm' (map Thm.assume prems)
@@ -196,7 +196,7 @@
       | SOME th => check_result ctxt thesis (MetaSimplifier.norm_hhf (Goal.conclude th)));
 
     val closed_rule = Thm.forall_intr (cert (Free thesis_var)) rule;
-    val ((_, [rule']), ctxt') = Variable.import_thms false [closed_rule] ctxt;
+    val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt;
     val obtain_rule = Thm.forall_elim (cert (Logic.varify (Free thesis_var))) rule';
     val ((params, stmt), fix_ctxt) = Variable.focus (Thm.cprem_of obtain_rule 1) ctxt';
     val (prems, ctxt'') =
@@ -249,7 +249,7 @@
       |> Thm.forall_intr (cert (Free thesis_var))
       |> Thm.instantiate (instT, []);
 
-    val ((_, rule' :: terms'), ctxt') = Variable.import_thms false (closed_rule :: terms) ctxt;
+    val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt;
     val vars' =
       map (dest_Free o Thm.term_of o Drule.dest_term) terms' ~~
       (map snd vars @ replicate (length ys) NoSyn);
--- a/src/Pure/Isar/rule_cases.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/rule_cases.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -333,7 +333,7 @@
   | mutual_rule _ [th] = SOME ([0], th)
   | mutual_rule ctxt (ths as th :: _) =
       let
-        val ((_, ths'), ctxt') = Variable.import_thms true ths ctxt;
+        val ((_, ths'), ctxt') = Variable.import true ths ctxt;
         val rules as (prems, _) :: _ = map (prep_rule (get_consumes th)) ths';
         val (ns, rls) = split_list (map #2 rules);
       in
--- a/src/Pure/Isar/theory_target.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Isar/theory_target.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -330,15 +330,6 @@
        else I)}
 and init_lthy_ctxt ta = init_lthy ta o init_ctxt ta;
 
-fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
-  let
-    val all_arities = map (fn raw_tyco => Sign.read_arity thy
-      (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
-    val tycos = map #1 all_arities;
-    val (_, sorts, sort) = hd all_arities;
-    val vs = Name.names Name.context Name.aT sorts;
-  in (tycos, vs, sort) end;
-
 fun gen_overloading prep_const raw_ops thy =
   let
     val ctxt = ProofContext.init thy;
@@ -356,7 +347,7 @@
 
 fun instantiation arities = init_lthy_ctxt (make_target "" false false arities []);
 fun instantiation_cmd raw_arities thy =
-  instantiation (read_multi_arity thy raw_arities) thy;
+  instantiation (Class_Target.read_multi_arity thy raw_arities) thy;
 
 val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const);
 val overloading_cmd = gen_overloading Syntax.read_term;
--- a/src/Pure/ML-Systems/multithreading_polyml.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/ML-Systems/multithreading_polyml.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -38,7 +38,7 @@
 
 val available = true;
 
-val max_threads = ref 1;
+val max_threads = ref 0;
 
 val tested_platform =
   let val ml_platform = getenv "ML_PLATFORM"
--- a/src/Pure/System/cygwin.scala	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/System/cygwin.scala	Wed Jul 01 16:19:44 2009 +0100
@@ -7,6 +7,7 @@
 package isabelle
 
 import java.lang.reflect.Method
+import java.io.File
 
 
 object Cygwin
@@ -75,10 +76,31 @@
 
   /* Cygwin installation */
 
+  // old-style mount points (Cygwin 1.5)
   private val CYGWIN_MOUNTS = "Software\\Cygnus Solutions\\Cygwin\\mounts v2"
 
-  def cygdrive(): Option[String] = query_registry(CYGWIN_MOUNTS, "cygdrive prefix")
-  def root(): Option[String] = query_registry(CYGWIN_MOUNTS + "\\/", "native")
+  // new-style setup (Cygwin 1.7)
+  private val CYGWIN_SETUP1 = "Software\\Cygwin\\setup"
+  private val CYGWIN_SETUP2 = "Software\\Wow6432Node\\Cygwin\\setup"  // !?
 
+  def config(): (String, String) =
+  {
+    query_registry(CYGWIN_SETUP1, "rootdir") match {
+      case Some(root) => (root, "/cygdrive")
+      case None =>
+        val root =
+          query_registry(CYGWIN_MOUNTS + "\\/", "native") getOrElse "C:\\cygwin"
+        val cygdrive =
+          query_registry(CYGWIN_MOUNTS, "cygdrive prefix") getOrElse "cygdrive"
+        (root, cygdrive)
+    }
+  }
+
+
+  /* basic sanity check */
+
+  def check(root: String): Boolean =
+    new File(root + "\\bin\\bash.exe").isFile &&
+    new File(root + "\\bin\\env.exe").isFile
 }
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/System/gui_setup.scala	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,65 @@
+/*  Title:      Pure/System/gui_setup.scala
+    Author:     Makarius
+
+GUI for basic system setup.
+*/
+
+package isabelle
+
+import javax.swing.UIManager
+
+import scala.swing._
+import scala.swing.event._
+
+
+object GUI_Setup extends GUIApplication
+{
+  def main(args: Array[String]) =
+  {
+    Swing_Thread.later {
+      UIManager.setLookAndFeel(Platform.look_and_feel)
+      top.pack()
+      top.visible = true
+    }
+  }
+
+  def top = new MainFrame {
+    title = "Isabelle setup"
+
+    // components
+    val text = new TextArea {
+      editable = false
+      columns = 80
+      rows = 20
+      xLayoutAlignment = 0.5
+    }
+    val ok = new Button {
+      text = "OK"
+      xLayoutAlignment = 0.5
+    }
+    contents = new BoxPanel(Orientation.Vertical) {
+      contents += text
+      contents += ok
+    }
+
+    // values
+    if (Platform.is_windows) {
+      text.append("Cygwin root: " + Cygwin.config()._1 + "\n")
+    }
+    Platform.defaults match {
+      case None =>
+      case Some((name, None)) => text.append("Platform: " + name + "\n")
+      case Some((name1, Some(name2))) =>
+        text.append("Main platform: " + name1 + "\n")
+        text.append("Alternative platform: " + name2 + "\n")
+    }
+    text.append("Isabelle home: " + java.lang.System.getProperty("isabelle.home"))
+
+    // reactions
+    listenTo(ok)
+    reactions += {
+      case ButtonClicked(`ok`) => System.exit(0)
+    }
+  }
+}
+
--- a/src/Pure/System/isabelle_process.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/System/isabelle_process.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -27,7 +27,7 @@
   val init: string -> unit
 end;
 
-structure IsabelleProcess: ISABELLE_PROCESS =
+structure Isabelle_Process: ISABELLE_PROCESS =
 struct
 
 (* print modes *)
--- a/src/Pure/System/isabelle_process.scala	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/System/isabelle_process.scala	Wed Jul 01 16:19:44 2009 +0100
@@ -12,7 +12,7 @@
   InputStream, OutputStream, IOException}
 
 
-object IsabelleProcess {
+object Isabelle_Process {
 
   /* results */
 
@@ -96,7 +96,7 @@
     def is_system = Kind.is_system(kind)
   }
 
-  def parse_message(isabelle_system: IsabelleSystem, result: Result) =
+  def parse_message(isabelle_system: Isabelle_System, result: Result) =
   {
     XML.Elem(Markup.MESSAGE, (Markup.CLASS, Kind.markup(result.kind)) :: result.props,
       YXML.parse_body_failsafe(isabelle_system.symbols.decode(result.result)))
@@ -104,16 +104,16 @@
 }
 
 
-class IsabelleProcess(isabelle_system: IsabelleSystem,
-  results: EventBus[IsabelleProcess.Result], args: String*)
+class Isabelle_Process(isabelle_system: Isabelle_System,
+  results: EventBus[Isabelle_Process.Result], args: String*)
 {
-  import IsabelleProcess._
+  import Isabelle_Process._
 
 
   /* demo constructor */
 
   def this(args: String*) =
-    this(new IsabelleSystem, new EventBus[IsabelleProcess.Result] + Console.println, args: _*)
+    this(new Isabelle_System, new EventBus[Isabelle_Process.Result] + Console.println, args: _*)
 
 
   /* process information */
@@ -128,7 +128,7 @@
   /* results */
 
   def parse_message(result: Result): XML.Tree =
-    IsabelleProcess.parse_message(isabelle_system, result)
+    Isabelle_Process.parse_message(isabelle_system, result)
 
   private val result_queue = new LinkedBlockingQueue[Result]
 
@@ -230,7 +230,7 @@
 
   private class StdinThread(out_stream: OutputStream) extends Thread("isabelle: stdin") {
     override def run() = {
-      val writer = new BufferedWriter(new OutputStreamWriter(out_stream, IsabelleSystem.charset))
+      val writer = new BufferedWriter(new OutputStreamWriter(out_stream, Isabelle_System.charset))
       var finished = false
       while (!finished) {
         try {
@@ -260,7 +260,7 @@
 
   private class StdoutThread(in_stream: InputStream) extends Thread("isabelle: stdout") {
     override def run() = {
-      val reader = new BufferedReader(new InputStreamReader(in_stream, IsabelleSystem.charset))
+      val reader = new BufferedReader(new InputStreamReader(in_stream, Isabelle_System.charset))
       var result = new StringBuilder(100)
 
       var finished = false
--- a/src/Pure/System/isabelle_system.scala	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/System/isabelle_system.scala	Wed Jul 01 16:19:44 2009 +0100
@@ -1,25 +1,23 @@
 /*  Title:      Pure/System/isabelle_system.scala
     Author:     Makarius
 
-Isabelle system support -- basic Cygwin/Posix compatibility.
+Isabelle system support, with basic Cygwin/Posix compatibility.
 */
 
 package isabelle
 
 import java.util.regex.Pattern
+import java.util.Locale
 import java.io.{BufferedReader, InputStreamReader, FileInputStream, File, IOException}
 
 import scala.io.Source
 import scala.util.matching.Regex
 
 
-object IsabelleSystem
+object Isabelle_System
 {
-
   val charset = "UTF-8"
 
-  val is_cygwin = System.getProperty("os.name").startsWith("Windows")
-
 
   /* shell processes */
 
@@ -48,28 +46,34 @@
 }
 
 
-class IsabelleSystem
+class Isabelle_System
 {
 
-  /* unique ids */
+  /** unique ids **/
 
   private var id_count: BigInt = 0
   def id(): String = synchronized { id_count += 1; "j" + id_count }
 
 
-  /* Isabelle settings environment */
+
+  /** Isabelle environment **/
+
+  /* platform prefixes */
 
   private val (platform_root, drive_prefix, shell_prefix) =
   {
-    if (IsabelleSystem.is_cygwin) {
-      val root = Cygwin.root() getOrElse "C:\\cygwin"
-      val drive = Cygwin.cygdrive() getOrElse "/cygdrive"
+    if (Platform.is_windows) {
+      val (root, drive) = Cygwin.config()
+      if (!Cygwin.check(root)) error("Bad Cygwin installation: " + root)
       val shell = List(root + "\\bin\\bash", "-l")
       (root, drive, shell)
     }
     else ("/", "", Nil)
   }
 
+
+  /* bash environment */
+
   private val environment: Map[String, String] =
   {
     import scala.collection.jcl.Conversions._
@@ -88,8 +92,8 @@
     val dump = File.createTempFile("isabelle", null)
     try {
       val cmdline = shell_prefix ::: List(isabelle, "getenv", "-d", dump.toString)
-      val proc = IsabelleSystem.raw_execute(env0, true, cmdline: _*)
-      val (output, rc) = IsabelleSystem.process_output(proc)
+      val proc = Isabelle_System.raw_execute(env0, true, cmdline: _*)
+      val (output, rc) = Isabelle_System.process_output(proc)
       if (rc != 0) error(output)
 
       val entries =
@@ -105,10 +109,11 @@
     finally { dump.delete }
   }
 
+
+  /* getenv */
+
   def getenv(name: String): String =
-  {
     environment.getOrElse(if (name == "HOME") "HOME_JVM" else name, "")
-  }
 
   def getenv_strict(name: String): String =
   {
@@ -119,44 +124,47 @@
   override def toString = getenv("ISABELLE_HOME")
 
 
-  /* file path specifications */
+
+  /** file path specifications **/
 
-  private val Cygdrive = new Regex(
-    if (drive_prefix == "") "\0"
-    else Pattern.quote(drive_prefix) + "/([a-zA-Z])($|/.*)")
+  /* expand_path */
 
-  def platform_path(source_path: String): String =
+  def expand_path(isabelle_path: String): String =
   {
     val result_path = new StringBuilder
-
-    def init(path: String): String =
+    def init(path: String)
     {
-      path match {
-        case Cygdrive(drive, rest) =>
-          result_path.length = 0
-          result_path.append(drive)
-          result_path.append(":")
-          result_path.append(File.separator)
-          rest
-        case _ if (path.startsWith("/")) =>
-          result_path.length = 0
-          result_path.append(platform_root)
-          path.substring(1)
-        case _ => path
+      if (path.startsWith("/")) {
+        result_path.clear
+        result_path += '/'
       }
     }
     def append(path: String)
     {
-      for (p <- init(path) split "/") {
-        if (p != "") {
+      init(path)
+      for (p <- path.split("/") if p != "" && p != ".") {
+        if (p == "..") {
+          val result = result_path.toString
+          val i = result.lastIndexOf("/")
+          if (result == "")
+            result_path ++= ".."
+          else if (result.substring(i + 1) == "..")
+            result_path ++= "/.."
+          else if (i < 1)
+            result_path.length = i + 1
+          else
+            result_path.length = i
+        }
+        else {
           val len = result_path.length
-          if (len > 0 && result_path(len - 1) != File.separatorChar)
-            result_path.append(File.separator)
-          result_path.append(p)
+          if (len > 0 && result_path(len - 1) != '/')
+            result_path += '/'
+          result_path ++= p
         }
       }
     }
-    for (p <- init(source_path) split "/") {
+    init(isabelle_path)
+    for (p <- isabelle_path.split("/")) {
       if (p.startsWith("$")) append(getenv_strict(p.substring(1)))
       else if (p == "~") append(getenv_strict("HOME"))
       else if (p == "~~") append(getenv_strict("ISABELLE_HOME"))
@@ -165,9 +173,58 @@
     result_path.toString
   }
 
+
+  /* platform_path */
+
+  private val Cygdrive =
+    new Regex(Pattern.quote(drive_prefix) + "/([a-zA-Z])($|/.*)")
+
+  def platform_path(isabelle_path: String): String =
+  {
+    val result_path = new StringBuilder
+    val rest =
+      expand_path(isabelle_path) match {
+        case Cygdrive(drive, rest) if Platform.is_windows =>
+          result_path ++= (drive + ":" + File.separator)
+          rest
+        case path if path.startsWith("/") =>
+          result_path ++= platform_root
+          path
+        case path => path
+      }
+    for (p <- rest.split("/") if p != "") {
+      val len = result_path.length
+      if (len > 0 && result_path(len - 1) != File.separatorChar)
+        result_path += File.separatorChar
+      result_path ++= p
+    }
+    result_path.toString
+  }
+
   def platform_file(path: String) = new File(platform_path(path))
 
 
+  /* isabelle_path */
+
+  private val Platform_Root = new Regex("(?i)" +
+    Pattern.quote(platform_root) + """(?:\\+|\z)(.*)""")
+  private val Drive = new Regex("""([a-zA-Z]):\\*(.*)""")
+
+  def isabelle_path(platform_path: String): String =
+  {
+    if (Platform.is_windows) {
+      platform_path.replace('/', '\\') match {
+        case Platform_Root(rest) => "/" + rest.replace('\\', '/')
+        case Drive(letter, rest) =>
+          drive_prefix + "/" + letter.toLowerCase(Locale.ENGLISH) +
+            (if (rest == "") "" else "/" + rest.replace('\\', '/'))
+        case path => path.replace('\\', '/')
+      }
+    }
+    else platform_path
+  }
+
+
   /* source files */
 
   private def try_file(file: File) = if (file.isFile) Some(file) else None
@@ -186,20 +243,30 @@
   }
 
 
+  /** system tools **/
+
   /* external processes */
 
   def execute(redirect: Boolean, args: String*): Process =
   {
     val cmdline =
-      if (IsabelleSystem.is_cygwin) List(platform_path("/bin/env")) ++ args
+      if (Platform.is_windows) List(platform_path("/bin/env")) ++ args
       else args
-    IsabelleSystem.raw_execute(environment, redirect, cmdline: _*)
+    Isabelle_System.raw_execute(environment, redirect, cmdline: _*)
   }
 
-  def isabelle_tool(args: String*): (String, Int) =
+  def isabelle_tool(name: String, args: String*): (String, Int) =
   {
-    val proc = execute(true, (List(getenv_strict("ISABELLE_TOOL")) ++ args): _*)
-    IsabelleSystem.process_output(proc)
+    getenv_strict("ISABELLE_TOOLS").split(":").find(dir => {
+      val file = platform_file(dir + "/" + name)
+      try { file.isFile && file.canRead } //  file.canExecute requires Java 1.6
+      catch { case _: SecurityException => false }
+    }) match {
+      case Some(dir) =>
+        val proc = execute(true, (List(platform_path(dir + "/" + name)) ++ args): _*)
+        Isabelle_System.process_output(proc)
+      case None => ("Unknown Isabelle tool: " + name, 2)
+    }
   }
 
 
@@ -222,12 +289,14 @@
   {
     // blocks until writer is ready
     val stream =
-      if (IsabelleSystem.is_cygwin) execute(false, "cat", fifo).getInputStream
+      if (Platform.is_windows) execute(false, "cat", fifo).getInputStream
       else new FileInputStream(fifo)
-    new BufferedReader(new InputStreamReader(stream, IsabelleSystem.charset))
+    new BufferedReader(new InputStreamReader(stream, Isabelle_System.charset))
   }
 
 
+  /** Isabelle resources **/
+
   /* find logics */
 
   def find_logics(): List[String] =
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/System/platform.scala	Wed Jul 01 16:19:44 2009 +0100
@@ -0,0 +1,69 @@
+/*  Title:      Pure/System/platform.scala
+    Author:     Makarius
+
+Raw platform identification.
+*/
+
+package isabelle
+
+import javax.swing.UIManager
+
+import scala.util.matching.Regex
+
+
+object Platform
+{
+  /* main OS variants */
+
+  val is_macos = System.getProperty("os.name") == "Mac OS X"
+  val is_windows = System.getProperty("os.name").startsWith("Windows")
+
+
+  /* Isabelle platform identifiers */
+
+  private val Solaris = new Regex("SunOS|Solaris")
+  private val Linux = new Regex("Linux")
+  private val Darwin = new Regex("Mac OS X")
+  private val Cygwin = new Regex("Windows.*")
+
+  private val X86 = new Regex("i.86|x86")
+  private val X86_64 = new Regex("amd64|x86_64")
+  private val Sparc = new Regex("sparc")
+  private val PPC = new Regex("PowerPC|ppc")
+
+  // main default, optional 64bit variant
+  val defaults: Option[(String, Option[String])] =
+  {
+    (java.lang.System.getProperty("os.name") match {
+      case Solaris() => Some("solaris")
+      case Linux() => Some("linux")
+      case Darwin() => Some("darwin")
+      case Cygwin() => Some("cygwin")
+      case _ => None
+    }) match {
+      case Some(name) =>
+        java.lang.System.getProperty("os.arch") match {
+          case X86() => Some(("x86-" + name, None))
+          case X86_64() => Some(("x86-" + name, if (is_windows) None else Some("x86_64-" + name)))
+          case Sparc() => Some(("sparc-" + name, None))
+          case PPC() => Some(("ppc-" + name, None))
+        }
+      case None => None
+    }
+  }
+
+
+  /* Swing look-and-feel */
+
+  def look_and_feel(): String =
+  {
+    if (is_windows || is_macos) UIManager.getSystemLookAndFeelClassName()
+    else {
+      UIManager.getInstalledLookAndFeels().find(laf => laf.getName == "Nimbus") match {
+        case None => UIManager.getCrossPlatformLookAndFeelClassName()
+        case Some(laf) => laf.getClassName
+      }
+    }
+  }
+}
+
--- a/src/Pure/System/session.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/System/session.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -103,7 +103,10 @@
           val start = start_timing ();
           val _ = use root;
           val stop = end_timing start;
-          val _ = Output.std_error ("Timing " ^ item ^ " (" ^ #message stop ^ ")\n");
+          val _ =
+            Output.std_error ("Timing " ^ item ^ " (" ^
+              string_of_int (Multithreading.max_threads_value ()) ^ " threads, " ^
+              #message stop ^ ")\n");
         in () end
       else use root;
       finish ()))
--- a/src/Pure/Thy/present.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/Thy/present.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -95,7 +95,7 @@
     val _ = writeln "Displaying graph ...";
     val path = File.tmp_path (Path.explode "tmp.graph");
     val _ = write_graph gr path;
-    val _ = File.isabelle_tool ("browser -c " ^ File.shell_path path ^ " &");
+    val _ = File.isabelle_tool "browser" ("-c " ^ File.shell_path path ^ " &");
   in () end;
 
 
@@ -341,10 +341,11 @@
     val pdf_path = File.tmp_path graph_pdf_path;
     val eps_path = File.tmp_path graph_eps_path;
     val gr_path = File.tmp_path graph_path;
-    val s = "browser -o " ^ File.shell_path pdf_path ^ " " ^ File.shell_path gr_path;
+    val args = "-o " ^ File.shell_path pdf_path ^ " " ^ File.shell_path gr_path;
   in
     write_graph graph gr_path;
-    if File.isabelle_tool s <> 0 orelse not (File.exists eps_path) orelse not (File.exists pdf_path)
+    if File.isabelle_tool "browser" args <> 0 orelse
+      not (File.exists eps_path) orelse not (File.exists pdf_path)
     then error "Failed to prepare dependency graph"
     else
       let val pdf = File.read pdf_path and eps = File.read eps_path
@@ -385,7 +386,8 @@
     fun prepare_sources cp path =
      (File.mkdir path;
       if cp then File.copy_dir document_path path else ();
-      File.isabelle_tool ("latex -o sty " ^ File.shell_path (Path.append path (Path.basic "root.tex")));
+      File.isabelle_tool "latex"
+        ("-o sty " ^ File.shell_path (Path.append path (Path.basic "root.tex")));
       (case opt_graphs of NONE => () | SOME (pdf, eps) =>
         (File.write (Path.append path graph_pdf_path) pdf;
           File.write (Path.append path graph_eps_path) eps));
@@ -513,8 +515,8 @@
     val _ = File.mkdir doc_path;
     val root_path = Path.append doc_path (Path.basic "root.tex");
     val _ = File.copy (Path.explode "~~/lib/texinputs/draft.tex") root_path;
-    val _ = File.isabelle_tool ("latex -o sty " ^ File.shell_path root_path);
-    val _ = File.isabelle_tool ("latex -o syms " ^ File.shell_path root_path);
+    val _ = File.isabelle_tool "latex" ("-o sty " ^ File.shell_path root_path);
+    val _ = File.isabelle_tool "latex" ("-o syms " ^ File.shell_path root_path);
 
     fun known name =
       let val ss = split_lines (File.read (Path.append doc_path (Path.basic name)))
--- a/src/Pure/subgoal.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/subgoal.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -29,7 +29,7 @@
 fun focus ctxt i st =
   let
     val ((schematics, [st']), ctxt') =
-      Variable.import_thms true [Simplifier.norm_hhf_protect st] ctxt;
+      Variable.import true [Simplifier.norm_hhf_protect st] ctxt;
     val ((params, goal), ctxt'') = Variable.focus_subgoal i st' ctxt';
     val asms = Drule.strip_imp_prems goal;
     val concl = Drule.strip_imp_concl goal;
--- a/src/Pure/tactic.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/tactic.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -91,7 +91,7 @@
 fun rule_by_tactic tac rl =
   let
     val ctxt = Variable.thm_context rl;
-    val ((_, [st]), ctxt') = Variable.import_thms true [rl] ctxt;
+    val ((_, [st]), ctxt') = Variable.import true [rl] ctxt;
   in
     (case Seq.pull (tac st) of
       NONE => raise THM ("rule_by_tactic", 0, [rl])
--- a/src/Pure/variable.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Pure/variable.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -51,9 +51,9 @@
     (((indexname * sort) * typ) list * ((indexname * typ) * term) list) * Proof.context
   val importT_terms: term list -> Proof.context -> term list * Proof.context
   val import_terms: bool -> term list -> Proof.context -> term list * Proof.context
-  val importT_thms: thm list -> Proof.context -> (ctyp list * thm list) * Proof.context
+  val importT: thm list -> Proof.context -> (ctyp list * thm list) * Proof.context
   val import_prf: bool -> Proofterm.proof -> Proof.context -> Proofterm.proof * Proof.context
-  val import_thms: bool -> thm list -> Proof.context ->
+  val import: bool -> thm list -> Proof.context ->
     ((ctyp list * cterm list) * thm list) * Proof.context
   val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
   val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
@@ -424,7 +424,7 @@
   let val (inst, ctxt') = import_inst is_open ts ctxt
   in (map (TermSubst.instantiate inst) ts, ctxt') end;
 
-fun importT_thms ths ctxt =
+fun importT ths ctxt =
   let
     val thy = ProofContext.theory_of ctxt;
     val certT = Thm.ctyp_of thy;
@@ -439,7 +439,7 @@
     val (insts, ctxt') = import_inst is_open ts ctxt;
   in (Proofterm.instantiate insts prf, ctxt') end;
 
-fun import_thms is_open ths ctxt =
+fun import is_open ths ctxt =
   let
     val thy = ProofContext.theory_of ctxt;
     val cert = Thm.cterm_of thy;
@@ -457,8 +457,8 @@
   let val ((_, ths'), ctxt') = imp ths ctxt
   in exp ctxt' ctxt (f ctxt' ths') end;
 
-val tradeT = gen_trade importT_thms exportT;
-val trade = gen_trade (import_thms true) export;
+val tradeT = gen_trade importT exportT;
+val trade = gen_trade (import true) export;
 
 
 (* focus on outermost parameters *)
--- a/src/Tools/Code/code_haskell.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Tools/Code/code_haskell.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -23,14 +23,6 @@
 
 (** Haskell serializer **)
 
-fun pr_haskell_bind pr_term =
-  let
-    fun pr_bind ((NONE, NONE), _) = str "_"
-      | pr_bind ((SOME v, NONE), _) = str v
-      | pr_bind ((NONE, SOME p), _) = p
-      | pr_bind ((SOME v, SOME p), _) = brackets [str v, str "@", p];
-  in gen_pr_bind pr_bind pr_term end;
-
 fun pr_haskell_stmt labelled_name syntax_class syntax_tyco syntax_const
     init_syms deresolve is_cons contr_classparam_typs deriving_show =
   let
@@ -68,13 +60,14 @@
                   pr_term tyvars thm vars NOBR t1,
                   pr_term tyvars thm vars BR t2
                 ])
-      | pr_term tyvars thm vars fxy (IVar v) =
+      | pr_term tyvars thm vars fxy (IVar NONE) =
+          str "_"
+      | pr_term tyvars thm vars fxy (IVar (SOME v)) =
           (str o Code_Printer.lookup_var vars) v
       | pr_term tyvars thm vars fxy (t as _ `|=> _) =
           let
-            val (binds, t') = Code_Thingol.unfold_abs t;
-            fun pr ((v, pat), ty) = pr_bind tyvars thm BR ((SOME v, pat), ty);
-            val (ps, vars') = fold_map pr binds vars;
+            val (binds, t') = Code_Thingol.unfold_pat_abs t;
+            val (ps, vars') = fold_map (pr_bind tyvars thm BR o fst) binds vars;
           in brackets (str "\\" :: ps @ str "->" @@ pr_term tyvars thm vars' NOBR t') end
       | pr_term tyvars thm vars fxy (ICase (cases as (_, t0))) =
           (case Code_Thingol.unfold_const_app t0
@@ -97,13 +90,13 @@
           else (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts
         end
     and pr_app tyvars = gen_pr_app (pr_app' tyvars) (pr_term tyvars) syntax_const
-    and pr_bind tyvars = pr_haskell_bind (pr_term tyvars)
+    and pr_bind tyvars thm fxy p = gen_pr_bind (pr_term tyvars) thm fxy p
     and pr_case tyvars thm vars fxy (cases as ((_, [_]), _)) =
           let
             val (binds, body) = Code_Thingol.unfold_let (ICase cases);
             fun pr ((pat, ty), t) vars =
               vars
-              |> pr_bind tyvars thm BR ((NONE, SOME pat), ty)
+              |> pr_bind tyvars thm BR pat
               |>> (fn p => semicolon [p, str "=", pr_term tyvars thm vars NOBR t])
             val (ps, vars') = fold_map pr binds vars;
           in brackify_block fxy (str "let {")
@@ -114,7 +107,7 @@
           let
             fun pr (pat, body) =
               let
-                val (p, vars') = pr_bind tyvars thm NOBR ((NONE, SOME pat), ty) vars;
+                val (p, vars') = pr_bind tyvars thm NOBR pat vars;
               in semicolon [p, str "->", pr_term tyvars thm vars' NOBR body] end;
           in brackify_block fxy
             (concat [str "case", pr_term tyvars thm vars NOBR t, str "of", str "{"])
@@ -240,8 +233,6 @@
           end
       | pr_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, classparam_insts))) =
           let
-            val split_abs_pure = (fn (v, _) `|=> t => SOME (v, t) | _ => NONE);
-            val unfold_abs_pure = Code_Thingol.unfoldr split_abs_pure;
             val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
             fun pr_instdef ((classparam, c_inst), (thm, _)) = case syntax_const classparam
              of NONE => semicolon [
@@ -255,10 +246,10 @@
                     val const = if (is_some o syntax_const) c_inst_name
                       then NONE else (SOME o Long_Name.base_name o deresolve) c_inst_name;
                     val proto_rhs = Code_Thingol.eta_expand k (c_inst, []);
-                    val (vs, rhs) = unfold_abs_pure proto_rhs;
+                    val (vs, rhs) = (apfst o map) fst (Code_Thingol.unfold_abs proto_rhs);
                     val vars = init_syms
                       |> Code_Printer.intro_vars (the_list const)
-                      |> Code_Printer.intro_vars vs;
+                      |> Code_Printer.intro_vars (map_filter I vs);
                     val lhs = IConst (classparam, (([], []), tys)) `$$ map IVar vs;
                       (*dictionaries are not relevant at this late stage*)
                   in
@@ -447,30 +438,30 @@
 
 fun pretty_haskell_monad c_bind =
   let
-    fun dest_bind t1 t2 = case Code_Thingol.split_abs t2
-     of SOME (((v, pat), ty), t') =>
-          SOME ((SOME (((SOME v, pat), ty), true), t1), t')
+    fun dest_bind t1 t2 = case Code_Thingol.split_pat_abs t2
+     of SOME ((pat, ty), t') =>
+          SOME ((SOME ((pat, ty), true), t1), t')
       | NONE => NONE;
     fun dest_monad c_bind_name (IConst (c, _) `$ t1 `$ t2) =
           if c = c_bind_name then dest_bind t1 t2
           else NONE
       | dest_monad _ t = case Code_Thingol.split_let t
          of SOME (((pat, ty), tbind), t') =>
-              SOME ((SOME (((NONE, SOME pat), ty), false), tbind), t')
+              SOME ((SOME ((pat, ty), false), tbind), t')
           | NONE => NONE;
     fun implode_monad c_bind_name = Code_Thingol.unfoldr (dest_monad c_bind_name);
     fun pr_monad pr_bind pr (NONE, t) vars =
           (semicolon [pr vars NOBR t], vars)
-      | pr_monad pr_bind pr (SOME (bind, true), t) vars = vars
+      | pr_monad pr_bind pr (SOME ((bind, _), true), t) vars = vars
           |> pr_bind NOBR bind
           |>> (fn p => semicolon [p, str "<-", pr vars NOBR t])
-      | pr_monad pr_bind pr (SOME (bind, false), t) vars = vars
+      | pr_monad pr_bind pr (SOME ((bind, _), false), t) vars = vars
           |> pr_bind NOBR bind
           |>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]);
     fun pretty _ [c_bind'] pr thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2
      of SOME (bind, t') => let
           val (binds, t'') = implode_monad c_bind' t'
-          val (ps, vars') = fold_map (pr_monad (pr_haskell_bind (K pr) thm) pr) (bind :: binds) vars;
+          val (ps, vars') = fold_map (pr_monad (gen_pr_bind (K pr) thm) pr) (bind :: binds) vars;
         in (brackify fxy o single o Pretty.enclose "do {" "}" o Pretty.breaks) (ps @| pr vars' NOBR t'') end
       | NONE => brackify_infix (1, L) fxy
           [pr vars (INFX (1, L)) t1, str ">>=", pr vars (INFX (1, X)) t2]
--- a/src/Tools/Code/code_ml.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Tools/Code/code_ml.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -85,7 +85,9 @@
       | pr_typ fxy (ITyVar v) = str ("'" ^ v);
     fun pr_term is_closure thm vars fxy (IConst c) =
           pr_app is_closure thm vars fxy (c, [])
-      | pr_term is_closure thm vars fxy (IVar v) =
+      | pr_term is_closure thm vars fxy (IVar NONE) =
+          str "_"
+      | pr_term is_closure thm vars fxy (IVar (SOME v)) =
           str (Code_Printer.lookup_var vars v)
       | pr_term is_closure thm vars fxy (t as t1 `$ t2) =
           (case Code_Thingol.unfold_const_app t
@@ -94,9 +96,9 @@
                [pr_term is_closure thm vars NOBR t1, pr_term is_closure thm vars BR t2])
       | pr_term is_closure thm vars fxy (t as _ `|=> _) =
           let
-            val (binds, t') = Code_Thingol.unfold_abs t;
-            fun pr ((v, pat), ty) =
-              pr_bind is_closure thm NOBR ((SOME v, pat), ty)
+            val (binds, t') = Code_Thingol.unfold_pat_abs t;
+            fun pr (pat, ty) =
+              pr_bind is_closure thm NOBR pat
               #>> (fn p => concat [str "fn", p, str "=>"]);
             val (ps, vars') = fold_map pr binds vars;
           in brackets (ps @ [pr_term is_closure thm vars' NOBR t']) end
@@ -122,17 +124,13 @@
           :: (map (pr_dicts BR) o filter_out null) iss @ map (pr_term is_closure thm vars BR) ts
     and pr_app is_closure thm vars = gen_pr_app (pr_app' is_closure) (pr_term is_closure)
       syntax_const thm vars
-    and pr_bind' ((NONE, NONE), _) = str "_"
-      | pr_bind' ((SOME v, NONE), _) = str v
-      | pr_bind' ((NONE, SOME p), _) = p
-      | pr_bind' ((SOME v, SOME p), _) = concat [str v, str "as", p]
-    and pr_bind is_closure = gen_pr_bind pr_bind' (pr_term is_closure)
+    and pr_bind is_closure = gen_pr_bind (pr_term is_closure)
     and pr_case is_closure thm vars fxy (cases as ((_, [_]), _)) =
           let
             val (binds, body) = Code_Thingol.unfold_let (ICase cases);
             fun pr ((pat, ty), t) vars =
               vars
-              |> pr_bind is_closure thm NOBR ((NONE, SOME pat), ty)
+              |> pr_bind is_closure thm NOBR pat
               |>> (fn p => semicolon [str "val", p, str "=", pr_term is_closure thm vars NOBR t])
             val (ps, vars') = fold_map pr binds vars;
           in
@@ -146,7 +144,7 @@
           let
             fun pr delim (pat, body) =
               let
-                val (p, vars') = pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) vars;
+                val (p, vars') = pr_bind is_closure thm NOBR pat vars;
               in
                 concat [str delim, p, str "=>", pr_term is_closure thm vars' NOBR body]
               end;
@@ -394,7 +392,9 @@
       | pr_typ fxy (ITyVar v) = str ("'" ^ v);
     fun pr_term is_closure thm vars fxy (IConst c) =
           pr_app is_closure thm vars fxy (c, [])
-      | pr_term is_closure thm vars fxy (IVar v) =
+      | pr_term is_closure thm vars fxy (IVar NONE) =
+          str "_"
+      | pr_term is_closure thm vars fxy (IVar (SOME v)) =
           str (Code_Printer.lookup_var vars v)
       | pr_term is_closure thm vars fxy (t as t1 `$ t2) =
           (case Code_Thingol.unfold_const_app t
@@ -403,9 +403,8 @@
                 brackify fxy [pr_term is_closure thm vars NOBR t1, pr_term is_closure thm vars BR t2])
       | pr_term is_closure thm vars fxy (t as _ `|=> _) =
           let
-            val (binds, t') = Code_Thingol.unfold_abs t;
-            fun pr ((v, pat), ty) = pr_bind is_closure thm BR ((SOME v, pat), ty);
-            val (ps, vars') = fold_map pr binds vars;
+            val (binds, t') = Code_Thingol.unfold_pat_abs t;
+            val (ps, vars') = fold_map (pr_bind is_closure thm BR o fst) binds vars;
           in brackets (str "fun" :: ps @ str "->" @@ pr_term is_closure thm vars' NOBR t') end
       | pr_term is_closure thm vars fxy (ICase (cases as (_, t0))) = (case Code_Thingol.unfold_const_app t0
            of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
@@ -427,17 +426,13 @@
         :: ((map (pr_dicts BR) o filter_out null) iss @ map (pr_term is_closure thm vars BR) ts)
     and pr_app is_closure = gen_pr_app (pr_app' is_closure) (pr_term is_closure)
       syntax_const
-    and pr_bind' ((NONE, NONE), _) = str "_"
-      | pr_bind' ((SOME v, NONE), _) = str v
-      | pr_bind' ((NONE, SOME p), _) = p
-      | pr_bind' ((SOME v, SOME p), _) = brackets [p, str "as", str v]
-    and pr_bind is_closure = gen_pr_bind pr_bind' (pr_term is_closure)
+    and pr_bind is_closure = gen_pr_bind (pr_term is_closure)
     and pr_case is_closure thm vars fxy (cases as ((_, [_]), _)) =
           let
             val (binds, body) = Code_Thingol.unfold_let (ICase cases);
             fun pr ((pat, ty), t) vars =
               vars
-              |> pr_bind is_closure thm NOBR ((NONE, SOME pat), ty)
+              |> pr_bind is_closure thm NOBR pat
               |>> (fn p => concat
                   [str "let", p, str "=", pr_term is_closure thm vars NOBR t, str "in"])
             val (ps, vars') = fold_map pr binds vars;
@@ -449,7 +444,7 @@
           let
             fun pr delim (pat, body) =
               let
-                val (p, vars') = pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) vars;
+                val (p, vars') = pr_bind is_closure thm NOBR pat vars;
               in concat [str delim, p, str "->", pr_term is_closure thm vars' NOBR body] end;
           in
             brackets (
@@ -464,7 +459,7 @@
     fun fish_params vars eqs =
       let
         fun fish_param _ (w as SOME _) = w
-          | fish_param (IVar v) NONE = SOME v
+          | fish_param (IVar (SOME v)) NONE = SOME v
           | fish_param _ NONE = NONE;
         fun fillup_param _ (_, SOME v) = v
           | fillup_param x (i, NONE) = x ^ string_of_int i;
@@ -781,7 +776,7 @@
         val eqs = filter (snd o snd) raw_eqs;
         val (eqs', is_value) = if null (filter_out (null o snd) vs) then case eqs
          of [(([], t), thm)] => if (not o null o fst o Code_Thingol.unfold_fun) ty
-            then ([(([IVar "x"], t `$ IVar "x"), thm)], false)
+            then ([(([IVar (SOME "x")], t `$ IVar (SOME "x")), thm)], false)
             else (eqs, not (Code_Thingol.fold_constnames
               (fn name' => fn b => b orelse name = name') t false))
           | _ => (eqs, false)
--- a/src/Tools/Code/code_printer.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Tools/Code/code_printer.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -68,10 +68,9 @@
     -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
     -> (string -> const_syntax option)
     -> thm -> var_ctxt -> fixity -> const * iterm list -> Pretty.T
-  val gen_pr_bind: ((string option * Pretty.T option) * itype -> Pretty.T)
-    -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
+  val gen_pr_bind: (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
     -> thm -> fixity
-    -> (string option * iterm option) * itype -> var_ctxt -> Pretty.T * var_ctxt
+    -> iterm -> var_ctxt -> Pretty.T * var_ctxt
 
   val mk_name_module: Name.context -> string option -> (string -> string option)
     -> 'a Graph.T -> string -> string
@@ -216,16 +215,11 @@
           else pr_term thm vars fxy (Code_Thingol.eta_expand k app)
         end;
 
-fun gen_pr_bind pr_bind pr_term thm (fxy : fixity) ((v, pat), ty : itype) vars =
+fun gen_pr_bind pr_term thm (fxy : fixity) pat vars =
   let
-    val vs = case pat
-     of SOME pat => Code_Thingol.fold_varnames (insert (op =)) pat []
-      | NONE => [];
-    val vars' = intro_vars (the_list v) vars;
-    val vars'' = intro_vars vs vars';
-    val v' = Option.map (lookup_var vars') v;
-    val pat' = Option.map (pr_term thm vars'' fxy) pat;
-  in (pr_bind ((v', pat'), ty), vars'') end;
+    val vs = Code_Thingol.fold_varnames (insert (op =)) pat [];
+    val vars' = intro_vars vs vars;
+  in (pr_term thm vars' fxy pat, vars') end;
 
 
 (* mixfix syntax *)
--- a/src/Tools/Code/code_thingol.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Tools/Code/code_thingol.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -23,13 +23,13 @@
   type const = string * ((itype list * dict list list) * itype list (*types of arguments*))
   datatype iterm =
       IConst of const
-    | IVar of vname
+    | IVar of vname option
     | `$ of iterm * iterm
-    | `|=> of (vname * itype) * iterm
+    | `|=> of (vname option * itype) * iterm
     | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
         (*((term, type), [(selector pattern, body term )]), primitive term)*)
   val `$$ : iterm * iterm list -> iterm;
-  val `|==> : (vname * itype) list * iterm -> iterm;
+  val `|==> : (vname option * itype) list * iterm -> iterm;
   type typscheme = (vname * sort) list * itype;
 end;
 
@@ -40,13 +40,12 @@
   val unfoldr: ('a -> ('b * 'a) option) -> 'a -> 'b list * 'a
   val unfold_fun: itype -> itype list * itype
   val unfold_app: iterm -> iterm * iterm list
-  val split_abs: iterm -> (((vname * iterm option) * itype) * iterm) option
-  val unfold_abs: iterm -> ((vname * iterm option) * itype) list * iterm
+  val unfold_abs: iterm -> (vname option * itype) list * iterm
   val split_let: iterm -> (((iterm * itype) * iterm) * iterm) option
   val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm
+  val split_pat_abs: iterm -> ((iterm * itype) * iterm) option
+  val unfold_pat_abs: iterm -> (iterm * itype) list * iterm
   val unfold_const_app: iterm -> (const * iterm list) option
-  val collapse_let: ((vname * itype) * iterm) * iterm
-    -> (iterm * itype) * (iterm * iterm) list
   val eta_expand: int -> const * iterm list -> iterm
   val contains_dictvar: iterm -> bool
   val locally_monomorphic: iterm -> bool
@@ -126,9 +125,9 @@
 
 datatype iterm =
     IConst of const
-  | IVar of vname
+  | IVar of vname option
   | `$ of iterm * iterm
-  | `|=> of (vname * itype) * iterm
+  | `|=> of (vname option * itype) * iterm
   | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
     (*see also signature*)
 
@@ -139,14 +138,10 @@
   (fn op `$ t => SOME t
     | _ => NONE);
 
-val split_abs =
-  (fn (v, ty) `|=> (t as ICase (((IVar w, _), [(p, t')]), _)) =>
-        if v = w then SOME (((v, SOME p), ty), t') else SOME (((v, NONE), ty), t)
-    | (v, ty) `|=> t => SOME (((v, NONE), ty), t)
+val unfold_abs = unfoldr
+  (fn op `|=> t => SOME t
     | _ => NONE);
 
-val unfold_abs = unfoldr split_abs;
-
 val split_let = 
   (fn ICase (((td, ty), [(p, t)]), _) => SOME (((p, ty), td), t)
     | _ => NONE);
@@ -172,39 +167,55 @@
 
 fun fold_varnames f =
   let
-    fun add (IVar v) = f v
-      | add ((v, _) `|=> _) = f v
+    fun add (IVar (SOME v)) = f v
+      | add ((SOME v, _) `|=> _) = f v
       | add _ = I;
   in fold_aiterms add end;
 
 fun fold_unbound_varnames f =
   let
     fun add _ (IConst _) = I
-      | add vs (IVar v) = if not (member (op =) vs v) then f v else I
+      | add vs (IVar (SOME v)) = if not (member (op =) vs v) then f v else I
+      | add _ (IVar NONE) = I
       | add vs (t1 `$ t2) = add vs t1 #> add vs t2
-      | add vs ((v, _) `|=> t) = add (insert (op =) v vs) t
-      | add vs (ICase (_, t)) = add vs t;
+      | add vs ((SOME v, _) `|=> t) = add (insert (op =) v vs) t
+      | add vs ((NONE, _) `|=> t) = add vs t
+      | add vs (ICase (((t, _), ds), _)) = add vs t #> fold (add_case vs) ds
+    and add_case vs (p, t) = add (fold_varnames (insert (op =)) p vs) t;
   in add [] end;
 
-fun collapse_let (((v, ty), se), be as ICase (((IVar w, _), ds), _)) =
+fun exists_var t v = fold_unbound_varnames (fn w => fn b => v = w orelse b) t false;
+
+fun split_pat_abs ((NONE, ty) `|=> t) = SOME ((IVar NONE, ty), t)
+  | split_pat_abs ((SOME v, ty) `|=> t) = SOME (case t
+     of ICase (((IVar (SOME w), _), [(p, t')]), _) =>
+          if v = w andalso (exists_var p v orelse not (exists_var t' v))
+          then ((p, ty), t')
+          else ((IVar (SOME v), ty), t)
+      | _ => ((IVar (SOME v), ty), t))
+  | split_pat_abs _ = NONE;
+
+val unfold_pat_abs = unfoldr split_pat_abs;
+
+fun unfold_abs_eta [] t = ([], t)
+  | unfold_abs_eta (_ :: tys) (v_ty `|=> t) =
       let
-        fun exists_v t = fold_unbound_varnames (fn w => fn b =>
-          b orelse v = w) t false;
-      in if v = w andalso forall (fn (t1, t2) =>
-        exists_v t1 orelse not (exists_v t2)) ds
-        then ((se, ty), ds)
-        else ((se, ty), [(IVar v, be)])
-      end
-  | collapse_let (((v, ty), se), be) =
-      ((se, ty), [(IVar v, be)])
+        val (vs_tys, t') = unfold_abs_eta tys t;
+      in (v_ty :: vs_tys, t') end
+  | unfold_abs_eta tys t =
+      let
+        val ctxt = fold_varnames Name.declare t Name.context;
+        val vs_tys = (map o apfst) SOME (Name.names ctxt "a" tys);
+      in (vs_tys, t `$$ map (IVar o fst) vs_tys) end;
 
 fun eta_expand k (c as (_, (_, tys)), ts) =
   let
     val j = length ts;
     val l = k - j;
     val ctxt = (fold o fold_varnames) Name.declare ts Name.context;
-    val vs_tys = Name.names ctxt "a" ((curry Library.take l o curry Library.drop j) tys);
-  in vs_tys `|==> IConst c `$$ ts @ map (fn (v, _) => IVar v) vs_tys end;
+    val vs_tys = (map o apfst) SOME
+      (Name.names ctxt "a" ((curry Library.take l o curry Library.drop j) tys));
+  in vs_tys `|==> IConst c `$$ ts @ map (IVar o fst) vs_tys end;
 
 fun contains_dictvar t =
   let
@@ -574,14 +585,16 @@
 and translate_term thy algbr funcgr thm (Const (c, ty)) =
       translate_app thy algbr funcgr thm ((c, ty), [])
   | translate_term thy algbr funcgr thm (Free (v, _)) =
-      pair (IVar v)
+      pair (IVar (SOME v))
   | translate_term thy algbr funcgr thm (Abs (abs as (_, ty, _))) =
       let
         val (v, t) = Syntax.variant_abs abs;
+        val v' = if member (op =) (Term.add_free_names t []) v
+          then SOME v else NONE
       in
         translate_typ thy algbr funcgr ty
         ##>> translate_term thy algbr funcgr thm t
-        #>> (fn (ty, t) => (v, ty) `|=> t)
+        #>> (fn (ty, t) => (v', ty) `|=> t)
       end
   | translate_term thy algbr funcgr thm (t as _ $ _) =
       case strip_comb t
@@ -618,45 +631,37 @@
   #>> (fn (t, ts) => t `$$ ts)
 and translate_case thy algbr funcgr thm (num_args, (t_pos, case_pats)) (c_ty, ts) =
   let
-    val (tys, _) = (chop num_args o fst o strip_type o snd) c_ty;
-    val t = nth ts t_pos;
+    fun arg_types num_args ty = (fst o chop num_args o fst o strip_type) ty;
+    val tys = arg_types num_args (snd c_ty);
     val ty = nth tys t_pos;
-    val ts_clause = nth_drop t_pos ts;
-    fun mk_clause (co, num_co_args) t =
+    fun mk_constr c t = let val n = Code.no_args thy c
+      in ((c, arg_types (Code.no_args thy c) (fastype_of t) ---> ty), n) end;
+    val constrs = if null case_pats then []
+      else map2 mk_constr case_pats (nth_drop t_pos ts);
+    fun casify naming constrs ty ts =
       let
-        val (vs, body) = Term.strip_abs_eta num_co_args t;
-        val not_undefined = case body
-         of (Const (c, _)) => not (Code.is_undefined thy c)
-          | _ => true;
-        val pat = list_comb (Const (co, map snd vs ---> ty), map Free vs);
-      in (not_undefined, (pat, body)) end;
-    val clauses = if null case_pats then let val ([v_ty], body) =
-        Term.strip_abs_eta 1 (the_single ts_clause)
-      in [(true, (Free v_ty, body))] end
-      else map (uncurry mk_clause)
-        (AList.make (Code.no_args thy) case_pats ~~ ts_clause);
-    fun retermify ty (_, (IVar x, body)) =
-          (x, ty) `|=> body
-      | retermify _ (_, (pat, body)) =
+        val t = nth ts t_pos;
+        val ts_clause = nth_drop t_pos ts;
+        val undefineds = map_filter (lookup_const naming) (Code.undefineds thy);
+        fun mk_clause ((constr as IConst (_, (_, tys)), n), t) =
           let
-            val (IConst (_, (_, tys)), ts) = unfold_app pat;
-            val vs = map2 (fn IVar x => fn ty => (x, ty)) ts tys;
-          in vs `|==> body end;
-    fun mk_icase const t ty clauses =
-      let
-        val (ts1, ts2) = chop t_pos (map (retermify ty) clauses);
-      in
-        ICase (((t, ty), map_filter (fn (b, d) => if b then SOME d else NONE) clauses),
-          const `$$ (ts1 @ t :: ts2))
-      end;
+            val (vs, t') = unfold_abs_eta (curry Library.take n tys) t;
+            val is_undefined = case t
+             of IConst (c, _) => member (op =) undefineds c
+              | _ => false;
+          in if is_undefined then NONE else SOME (constr `$$ map (IVar o fst) vs, t') end;
+        val clauses = if null case_pats then let val ([(v, _)], t) =
+            unfold_abs_eta [ty] (the_single ts_clause)
+          in [(IVar v, t)] end
+          else map_filter mk_clause (constrs ~~ ts_clause);
+      in ((t, ty), clauses) end;
   in
     translate_const thy algbr funcgr thm c_ty
-    ##>> translate_term thy algbr funcgr thm t
+    ##>> fold_map (fn (constr, n) => translate_const thy algbr funcgr thm constr #>> rpair n) constrs
     ##>> translate_typ thy algbr funcgr ty
-    ##>> fold_map (fn (b, (pat, body)) => translate_term thy algbr funcgr thm pat
-      ##>> translate_term thy algbr funcgr thm body
-      #>> pair b) clauses
-    #>> (fn (((const, t), ty), ds) => mk_icase const t ty ds)
+    ##>> fold_map (translate_term thy algbr funcgr thm) ts
+    #-> (fn (((t, constrs), ty), ts) =>
+      `(fn (_, (naming, _)) => ICase (casify naming constrs ty ts, t `$$ ts)))
   end
 and translate_app_case thy algbr funcgr thm (case_scheme as (num_args, _)) ((c, ty), ts) =
   if length ts < num_args then
@@ -668,7 +673,7 @@
     in
       fold_map (translate_typ thy algbr funcgr) tys
       ##>> translate_case thy algbr funcgr thm case_scheme ((c, ty), ts @ map Free vs)
-      #>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|==> t)
+      #>> (fn (tys, t) => map2 (fn (v, _) => pair (SOME v)) vs tys `|==> t)
     end
   else if length ts > num_args then
     translate_case thy algbr funcgr thm case_scheme ((c, ty), Library.take (num_args, ts))
--- a/src/Tools/coherent.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Tools/coherent.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -110,9 +110,9 @@
 (* Check whether disjunction is valid in given state *)
 fun is_valid_disj ctxt facts [] = false
   | is_valid_disj ctxt facts ((Ts, ts) :: ds) =
-      let val vs = rev (map_index (fn (i, T) => Var (("x", i), T)) Ts)
+      let val vs = map_index (fn (i, T) => Var (("x", i), T)) Ts
       in case Seq.pull (valid_conj ctxt facts empty_env
-        (map (fn t => subst_bounds (vs, t)) ts)) of
+        (map (fn t => subst_bounds (rev vs, t)) ts)) of
           SOME _ => true
         | NONE => is_valid_disj ctxt facts ds
       end;
@@ -153,10 +153,10 @@
   | valid_cases ctxt rules goal dom facts nfacts nparams ((Ts, ts) :: ds) =
       let
         val _ = message (fn () => "case " ^ commas (map (Syntax.string_of_term ctxt) ts));
-        val params = rev (map_index (fn (i, T) =>
-          Free ("par" ^ string_of_int (nparams + i), T)) Ts);
+        val params = map_index (fn (i, T) =>
+          Free ("par" ^ string_of_int (nparams + i), T)) Ts;
         val ts' = map_index (fn (i, t) =>
-          (subst_bounds (params, t), nfacts + i)) ts;
+          (subst_bounds (rev params, t), nfacts + i)) ts;
         val dom' = fold (fn (T, p) =>
           Typtab.map_default (T, []) (fn ps => ps @ [p]))
             (Ts ~~ params) dom;
--- a/src/Tools/nbe.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Tools/nbe.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -135,6 +135,8 @@
   | nbe_fun i c = "c_" ^ translate_string (fn "." => "_" | c => c) c ^ "_" ^ string_of_int i;
 fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n;
 fun nbe_bound v = "v_" ^ v;
+fun nbe_bound_optional NONE = "_"
+  | nbe_bound_optional  (SOME v) = nbe_bound v;
 fun nbe_default v = "w_" ^ v;
 
 (*note: these three are the "turning spots" where proper argument order is established!*)
@@ -191,9 +193,9 @@
             val (t', ts) = Code_Thingol.unfold_app t
           in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end
         and of_iapp match_cont (IConst (c, ((_, dss), _))) ts = constapp c dss ts
-          | of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound v) ts
+          | of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound_optional v) ts
           | of_iapp match_cont ((v, _) `|=> t) ts =
-              nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound v]) (of_iterm NONE t))) ts
+              nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (of_iterm NONE t))) ts
           | of_iapp match_cont (ICase (((t, _), cs), t0)) ts =
               nbe_apps (ml_cases (of_iterm NONE t)
                 (map (fn (p, t) => (of_iterm NONE p, of_iterm match_cont t)) cs
@@ -212,8 +214,9 @@
           else pair (v, [])) vs names;
         val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames;
         fun subst_vars (t as IConst _) samepairs = (t, samepairs)
-          | subst_vars (t as IVar v) samepairs = (case AList.lookup (op =) samepairs v
-             of SOME v' => (IVar v', AList.delete (op =) v samepairs)
+          | subst_vars (t as IVar NONE) samepairs = (t, samepairs)
+          | subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v
+             of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs)
               | NONE => (t, samepairs))
           | subst_vars (t1 `$ t2) samepairs = samepairs
               |> subst_vars t1
@@ -295,7 +298,8 @@
         val params = Name.invent_list [] "d" (length names);
         fun mk (k, name) =
           (name, ([(v, [])],
-            [([IConst (class, (([], []), [])) `$$ map IVar params], IVar (nth params k))]));
+            [([IConst (class, (([], []), [])) `$$ map (IVar o SOME) params],
+              IVar (SOME (nth params k)))]));
       in map_index mk names end
   | eqns_of_stmt (_, Code_Thingol.Classrel _) =
       []
--- a/src/Tools/project_rule.ML	Wed Jun 24 16:28:30 2009 +0100
+++ b/src/Tools/project_rule.ML	Wed Jul 01 16:19:44 2009 +0100
@@ -39,7 +39,7 @@
       (case try imp th of
         NONE => (k, th)
       | SOME th' => prems (k + 1) th');
-    val ((_, [rule]), ctxt') = Variable.import_thms true [raw_rule] ctxt;
+    val ((_, [rule]), ctxt') = Variable.import true [raw_rule] ctxt;
     fun result i =
       rule
       |> proj i
@@ -59,7 +59,7 @@
       (case try conj2 th of
         NONE => k
       | SOME th' => projs (k + 1) th');
-    val ((_, [rule]), _) = Variable.import_thms true [raw_rule] ctxt;
+    val ((_, [rule]), _) = Variable.import true [raw_rule] ctxt;
   in projects ctxt (1 upto projs 1 rule) raw_rule end;
 
 end;