merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
authorimmler
Thu, 03 May 2018 15:07:14 +0200
changeset 68073 fad29d2a17a5
parent 68072 493b818e8e10 (current diff)
parent 68070 8dc792d440b9 (diff)
child 68074 8d50467f7555
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
CONTRIBUTORS
NEWS
src/HOL/Algebra/Order.thy
src/HOL/Analysis/Bochner_Integration.thy
src/HOL/Analysis/Cartesian_Euclidean_Space.thy
src/HOL/Analysis/Cartesian_Space.thy
src/HOL/Analysis/Change_Of_Vars.thy
src/HOL/Analysis/Convex_Euclidean_Space.thy
src/HOL/Analysis/Derivative.thy
src/HOL/Analysis/Determinants.thy
src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
src/HOL/Analysis/Finite_Cartesian_Product.thy
src/HOL/Analysis/Henstock_Kurzweil_Integration.thy
src/HOL/Analysis/Homeomorphism.thy
src/HOL/Analysis/Inner_Product.thy
src/HOL/Analysis/Linear_Algebra.thy
src/HOL/Analysis/Polytope.thy
src/HOL/Analysis/Sigma_Algebra.thy
src/HOL/Analysis/Starlike.thy
src/HOL/Codegenerator_Test/Generate_Pretty_Char.thy
src/HOL/Hull.thy
src/HOL/Library/Code_Char.thy
src/HOL/Library/Library.thy
src/HOL/Modules.thy
src/HOL/Tools/string_code.ML
src/HOL/Vector_Spaces.thy
src/Tools/Adhoc_Overloading.thy
src/Tools/adhoc_overloading.ML
--- a/Admin/PLATFORMS	Wed May 02 13:49:38 2018 +0200
+++ b/Admin/PLATFORMS	Thu May 03 15:07:14 2018 +0200
@@ -5,8 +5,8 @@
 --------
 
 The general programming model is that of a stylized ML + Scala + POSIX
-environment, with as little system-specific code in user-space tools
-as possible.
+environment, with a minimum of system-specific code in user-space
+tools.
 
 The Isabelle system infrastructure provides some facilities to make
 this work, e.g. see the ML and Scala modules File and Path, or
@@ -19,8 +19,8 @@
 When producing add-on tools, it is important to stay within this clean
 room of Isabelle, and refrain from non-portable access to operating
 system functions. The Isabelle environment uses peculiar scripts for
-GNU bash and perl to get the plumbing right. This style should be
-imitated as far as possible.
+GNU bash and perl as system glue: this style should be observed as far
+as possible.
 
 
 Supported platforms
@@ -36,6 +36,7 @@
   x86_64-darwin     Mac OS X 10.10 Yosemite (macbroy31 MacBookPro6,2)
                     Mac OS X 10.11 El Capitan (macbroy2 MacPro4,1)
                     macOS 10.12 Sierra (macbroy30 MacBookPro6,2)
+                    macOS 10.13 High Sierra
 
   x86_64-windows    Windows 7
   x86_64-cygwin     Cygwin 2.8 http://isabelle.in.tum.de/cygwin_2017 (x86_64/release)
@@ -43,7 +44,7 @@
 All of the above platforms are 100% supported by Isabelle -- end-users
 should not have to care about the differences (at least in theory).
 
-Fringe platforms like BSD or Solaris are not supported.
+Exotic platforms like BSD, Solaris, NixOS are not supported.
 
 
 64 bit vs. 32 bit platform personality
@@ -52,42 +53,41 @@
 Isabelle requires 64 bit hardware running a 64 bit operating
 system. Windows and Mac OS X allow x86 executables as well, but for
 Linux this requires separate installation of 32 bit shared
-libraries. The POSIX emulation on Windows via Cygwin64 is exclusively
-for x86_64.
+libraries. The POSIX emulation on Windows via Cygwin64 works
+exclusively for x86_64.
 
-ML works both for x86_64 and x86, and the latter is preferred for space
-and performance reasons. Java is always for x86_64 on all platforms.
+Poly/ML supports both for x86_64 and x86, and the latter is preferred
+for space and performance reasons. Java is always the x86_64 version
+on all platforms.
 
 Add-on executables are expected to work without manual user
 configuration. Each component settings script needs to determine the
 platform details appropriately.
 
 
-The Isabelle settings environment provides the following variables to
-help configuring platform-dependent tools:
+The Isabelle settings environment provides the following important
+variables to help configuring platform-dependent tools:
 
   ISABELLE_PLATFORM64  (potentially empty)
   ISABELLE_PLATFORM32  (potentially empty)
-  ISABELLE_PLATFORM
 
-The ISABELLE_PLATFORM setting variable prefers the 32 bit personality of
-the platform, if possible. Using regular bash notation, tools may
-express their preference for 64 bit with a fall-back for 32 bit as
-follows:
+Each can be empty, but not both at the same time. Using GNU bash
+notation, tools may express their platform preference, e.g. first 64
+bit and second 32 bit, or the opposite:
 
   "${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}"
+  "${ISABELLE_PLATFORM32:-$ISABELLE_PLATFORM64}"
 
 
-There is a second set of settings for native Windows (instead of the
+There is a another set of settings for native Windows (instead of the
 POSIX emulation of Cygwin used before):
 
   ISABELLE_WINDOWS_PLATFORM64
   ISABELLE_WINDOWS_PLATFORM32
-  ISABELLE_WINDOWS_PLATFORM
 
-It can be used like this:
-
-  "${ISABELLE_WINDOWS_PLATFORM:-$ISABELLE_PLATFORM}"
+These are always empty on Linux and Mac OS X, and non-empty on
+Windows. They can be used like this to prefer native Windows and then
+Unix (first 64 bit second 32 bit):
 
   "${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"
 
@@ -97,13 +97,13 @@
 
 The following portable system tools can be taken for granted:
 
-* Scala on top of Java 8.  Isabelle/Scala irons out many oddities and
+* Scala on top of Java.  Isabelle/Scala irons out many oddities and
   portability issues of the Java platform.
 
-* GNU bash as uniform shell on all platforms. The POSIX "standard" shell
-  /bin/sh does *not* work -- there are too many non-standard
-  implementations of it. On Debian and Ubuntu /bin/sh is actually
-  /bin/dash and thus introduces many oddities.
+* GNU bash as uniform shell on all platforms. The POSIX "standard"
+  shell /bin/sh does *not* work portably -- there are too many
+  non-standard implementations of it. On Debian and Ubuntu /bin/sh is
+  actually /bin/dash and introduces many oddities.
 
 * Perl as largely portable system programming language, with its
   fairly robust support for processes, signals, sockets etc.
@@ -123,12 +123,6 @@
   Such add-ons are usually included in Apple's /usr/bin/perl by
   default.
 
-* The Java runtime has its own idea about the underlying platform, which
-  affects Java native libraries in particular. In Isabelle/Scala the
-  function isabelle.Platform.jvm_platform identifies the JVM platform.
-  In the settings environment, ISABELLE_JAVA_PLATFORM provides the same
-  information without running the JVM.
-
 * Common Unix tools like /bin/sh, /bin/kill, sed, ulimit are
   notoriously non-portable an should be avoided.
 
--- a/Admin/components/components.sha1	Wed May 02 13:49:38 2018 +0200
+++ b/Admin/components/components.sha1	Thu May 03 15:07:14 2018 +0200
@@ -88,6 +88,7 @@
 e45edcf184f608d6f4a7b966d65a5d3289462693  jdk-8u144.tar.gz
 264e806b9300a4fb3b6e15ba0e2c664d4ea698c8  jdk-8u152.tar.gz
 84b04d877a2ea3a4e2082297b540e14f76722bc5  jdk-8u162.tar.gz
+87303a0de3fd595aa3857c8f7cececa036d6ed18  jdk-8u172.tar.gz
 cfecb1383faaf027ffbabfcd77a0b6a6521e0969  jdk-8u20.tar.gz
 44ffeeae219782d40ce6822b580e608e72fd4c76  jdk-8u31.tar.gz
 4132cf52d5025bf330d53b96a5c6466fef432377  jdk-8u51.tar.gz
@@ -127,6 +128,7 @@
 7bcb202e13358dd750e964b2f747664428b5d8b3  jedit_build-20180417.tar.gz
 0bd2bc2d9a491ba5fc8dd99df27c04f11a72e8fa  jfreechart-1.0.14-1.tar.gz
 8122526f1fc362ddae1a328bdbc2152853186fee  jfreechart-1.0.14.tar.gz
+d911f63a5c9b4c7335bb73f805cb1711ce017a84  jfreechart-1.5.0.tar.gz
 c8a19a36adf6cefa779d85f22ded2f4654e68ea5  jortho-1.0-1.tar.gz
 2155e0bdbd29cd3d2905454de2e7203b9661d239  jortho-1.0-2.tar.gz
 ffe179867cf5ffaabbb6bb096db9bdc0d7110065  jortho-1.0.tar.gz
@@ -175,6 +177,7 @@
 a619177143fea42a464f49bb864665407c07a16c  polyml-test-fb4f42af00fa.tar.gz
 53123dc011b2d4b4e8fe307f3c9fa355718ad01a  postgresql-42.1.1.tar.gz
 3a5d31377ec07a5069957f5477a4848cfc89a594  postgresql-42.1.4.tar.gz
+e7cd5c7955e9eb5ce8cd07feb97230b23d2eec40  postgresql-42.2.2.tar.gz
 f132329ca1045858ef456cc08b197c9eeea6881b  postgresql-9.4.1212.tar.gz
 8ee375cfc38972f080dbc78f07b68dac03efe968  ProofGeneral-3.7.1.1.tar.gz
 847b52c0676b5eb0fbf0476f64fc08c2d72afd0c  ProofGeneral-4.1.tar.gz
--- a/Admin/components/main	Wed May 02 13:49:38 2018 +0200
+++ b/Admin/components/main	Thu May 03 15:07:14 2018 +0200
@@ -5,14 +5,14 @@
 cvc4-1.5-3
 e-2.0-1
 isabelle_fonts-20180113
-jdk-8u162
+jdk-8u172
 jedit_build-20180417
-jfreechart-1.0.14-1
+jfreechart-1.5.0
 jortho-1.0-2
 kodkodi-1.5.2
 nunchaku-0.5
 polyml-5.7.1-5
-postgresql-42.1.4
+postgresql-42.2.2
 scala-2.12.5
 smbc-0.4.1
 ssh-java-20161009
--- a/CONTRIBUTORS	Wed May 02 13:49:38 2018 +0200
+++ b/CONTRIBUTORS	Thu May 03 15:07:14 2018 +0200
@@ -6,14 +6,16 @@
 Contributions to this Isabelle version
 --------------------------------------
 
-* April 2018: Jose Divasón (Universidad de la Rioja),
+* May 2018: Jose Divasón (Universidad de la Rioja),
   Jesús Aransay (Universidad de la Rioja), Johannes Hölzl (VU Amsterdam),
   Fabian Immler (TUM)
   Generalizations in the formalization of linear algebra.
 
+* May 2018: Florian Haftmann
+  Consolidation of string-like types in HOL.
 
 * March 2018: Florian Haftmann
-  Abstract bit operations push_bit, push_take, push_drop, alongside
+  Abstract bit operations push_bit, take_bit, drop_bit, alongside
   with an algebraic foundation for bit strings and word types in
   HOL-ex.
 
--- a/NEWS	Wed May 02 13:49:38 2018 +0200
+++ b/NEWS	Thu May 03 15:07:14 2018 +0200
@@ -110,7 +110,8 @@
 notably bibtex database files and ML files.
 
 * Action "isabelle.draft" is similar to "isabelle.preview", but shows a
-plain-text document draft.
+plain-text document draft. Both are available via the menu "Plugins /
+Isabelle".
 
 * When loading text files, the Isabelle symbols encoding UTF-8-Isabelle
 is only used if there is no conflict with existing Unicode sequences in
@@ -196,8 +197,38 @@
 
 *** HOL ***
 
-* Abstract bit operations as part of Main: push_bit, push_take,
-push_drop.
+* Clarified relationship of characters, strings and code generation:
+
+  * Type "char" is now a proper datatype of 8-bit values.
+
+  * Conversions "nat_of_char" and "char_of_nat" are gone; use more
+    general conversions "of_char" and "char_of" with suitable
+    type constraints instead.
+
+  * The zero character is just written "CHR 0x00", not
+    "0" any longer.
+
+  * Type "String.literal" (for code generation) is now isomorphic
+    to lists of 7-bit (ASCII) values; concrete values can be written
+    as "STR ''...''" for sequences of printable characters and
+    "STR 0x..." for one single ASCII code point given
+    as hexadecimal numeral.
+
+  * Type "String.literal" supports concatenation "... + ..."
+    for all standard target languages.
+
+  * Theory Library/Code_Char is gone; study the explanations concerning
+    "String.literal" in the tutorial on code generation to get an idea
+    how target-language string literals can be converted to HOL string
+    values and vice versa.
+
+  * Imperative-HOL: operation "raise" directly takes a value of type
+    "String.literal" as argument, not type "string".
+
+INCOMPATIBILITY.
+
+* Abstract bit operations as part of Main: push_bit, take_bit,
+drop_bit.
 
 * New, more general, axiomatization of complete_distrib_lattice. 
 The former axioms:
@@ -317,6 +348,20 @@
 been renamed to ISABELLE_TOOL_JAVA_OPTIONS and JEDIT_JAVA_OPTIONS,
 instead of former 32/64 variants. INCOMPATIBILITY.
 
+* Old settings ISABELLE_PLATFORM and ISABELLE_WINDOWS_PLATFORM should be
+phased out due to unclear preference of 32bit vs. 64bit architecture.
+Explicit GNU bash expressions are now preferred, for example (with
+quotes):
+
+  #Posix executables (Unix or Cygwin), with preference for 64bit
+  "${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}"
+
+  #native Windows or Unix executables, with preference for 64bit
+  "${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"
+
+  #native Windows (32bit) or Unix executables (preference for 64bit)
+  "${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}"
+
 * Command-line tool "isabelle build" supports new options:
   - option -B NAME: include session NAME and all descendants
   - option -S: only observe changes of sources, not heap images
--- a/lib/browser/build	Wed May 02 13:49:38 2018 +0200
+++ b/lib/browser/build	Thu May 03 15:07:14 2018 +0200
@@ -63,7 +63,7 @@
 
   rm -rf classes && mkdir classes
 
-  isabelle_jdk javac -d classes -source 1.4 "${SOURCES[@]}" || \
+  isabelle_jdk javac -d classes -source 1.6 "${SOURCES[@]}" || \
     fail "Failed to compile sources"
   isabelle_jdk jar cf "$(platform_path "$TARGET")" -C classes . ||
     fail "Failed to produce $TARGET"
--- a/lib/scripts/getsettings	Wed May 02 13:49:38 2018 +0200
+++ b/lib/scripts/getsettings	Thu May 03 15:07:14 2018 +0200
@@ -102,7 +102,12 @@
 ISABELLE_OUTPUT="$ISABELLE_OUTPUT/$ML_IDENTIFIER"
 
 #enforce JAVA_HOME
-export JAVA_HOME="$ISABELLE_JDK_HOME/jre"
+if [ -d "$ISABELLE_JDK_HOME/jre" ]
+then
+  export JAVA_HOME="$ISABELLE_JDK_HOME/jre"
+else
+  export JAVA_HOME="$ISABELLE_JDK_HOME"
+fi
 
 set +o allexport
 
--- a/src/Doc/Codegen/Adaptation.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/Doc/Codegen/Adaptation.thy	Thu May 03 15:07:14 2018 +0200
@@ -168,6 +168,35 @@
        Useful for code setups which involve e.g.~indexing
        of target-language arrays.  Part of @{text "HOL-Main"}.
 
+    \item[@{theory "String"}] provides an additional datatype @{typ
+       String.literal} which is isomorphic to lists of 7-bit (ASCII) characters;
+       @{typ String.literal}s are mapped to target-language strings.
+
+       Literal values of type @{typ String.literal} can be written
+       as @{text "STR ''\<dots>''"} for sequences of printable characters and
+       @{text "STR 0x\<dots>"} for one single ASCII code point given
+       as hexadecimal numeral; @{typ String.literal} supports concatenation
+       @{text "\<dots> + \<dots>"} for all standard target languages.
+
+       Note that the particular notion of \qt{string} is target-language
+       specific (sequence of 8-bit units, sequence of unicode code points, \ldots);
+       hence ASCII is the only reliable common base e.g.~for
+       printing (error) messages; more sophisticated applications
+       like verifying parsing algorithms require a dedicated
+       target-language specific model.
+
+       Nevertheless @{typ String.literal}s can be analyzed; the core operations
+       for this are @{term_type String.asciis_of_literal} and
+       @{term_type String.literal_of_asciis} which are implemented
+       in a target-language-specific way; particularly @{const String.asciis_of_literal}
+       checks its argument at runtime to make sure that it does
+       not contain non-ASCII-characters, to safeguard consistency.
+       On top of these, more abstract conversions like @{term_type
+       String.explode} and @{term_type String.implode}
+       are implemented.
+       
+       Part of @{text "HOL-Main"}.
+
     \item[@{text "Code_Target_Int"}] implements type @{typ int}
        by @{typ integer} and thus by target-language built-in integers.
 
@@ -186,17 +215,6 @@
        containing both @{text "Code_Target_Nat"} and
        @{text "Code_Target_Int"}.
 
-    \item[@{theory "String"}] provides an additional datatype @{typ
-       String.literal} which is isomorphic to strings; @{typ
-       String.literal}s are mapped to target-language strings.  Useful
-       for code setups which involve e.g.~printing (error) messages.
-       Part of @{text "HOL-Main"}.
-
-    \item[@{text "Code_Char"}] represents @{text HOL} characters by
-       character literals in target languages.  \emph{Warning:} This
-       modifies adaptation in a non-conservative manner and thus
-       should always be imported \emph{last} in a theory header.
-
     \item[@{theory "IArray"}] provides a type @{typ "'a iarray"}
        isomorphic to lists but implemented by (effectively immutable)
        arrays \emph{in SML only}.
--- a/src/Doc/Codegen/Computations.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/Doc/Codegen/Computations.thy	Thu May 03 15:07:14 2018 +0200
@@ -472,20 +472,20 @@
   check_int @{context} @{cprop "even ((0::int) + 1 + 2 + 3 + -1 + -2 + -3)"}
 \<close>
   
-paragraph \<open>An example for @{typ char}\<close>
+paragraph \<open>An example for @{typ String.literal}\<close>
 
-definition %quote is_cap_letter :: "char \<Rightarrow> bool"
-  where "is_cap_letter c \<longleftrightarrow> (let n = nat_of_char c in 65 \<le> n \<and> n \<le> 90)" (*<*)
+definition %quote is_cap_letter :: "String.literal \<Rightarrow> bool"
+  where "is_cap_letter s \<longleftrightarrow> (case String.asciis_of_literal s
+    of [] \<Rightarrow> False | k # _ \<Rightarrow> 65 \<le> k \<and> k \<le> 90)" (*<*)
 
 (*>*) ML %quotetypewriter \<open>
-  val check_char = @{computation_check terms:
-    Trueprop is_cap_letter
-    Char datatypes: num
+  val check_literal = @{computation_check terms:
+    Trueprop is_cap_letter datatypes: bool String.literal
   }
 \<close>
 
 ML_val %quotetypewriter \<open>
-  check_char @{context} @{cprop "is_cap_letter (CHR ''Y'')"}
+  check_literal @{context} @{cprop "is_cap_letter (STR ''Q'')"}
 \<close>
 
   
--- a/src/Doc/JEdit/JEdit.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/Doc/JEdit/JEdit.thy	Thu May 03 15:07:14 2018 +0200
@@ -73,9 +73,9 @@
   Isabelle/jEdit (\figref{fig:isabelle-jedit}) consists of some plugins for
   the jEdit text editor, while preserving its general look-and-feel as far as
   possible. The main plugin is called ``Isabelle'' and has its own menu
-  \<^emph>\<open>Plugins~/ Isabelle\<close> with access to several panels (see also
-  \secref{sec:dockables}), as well as \<^emph>\<open>Plugins~/ Plugin Options~/ Isabelle\<close>
-  (see also \secref{sec:options}).
+  \<^emph>\<open>Plugins~/ Isabelle\<close> with access to several actions and add-on panels (see
+  also \secref{sec:dockables}), as well as \<^emph>\<open>Plugins~/ Plugin Options~/
+  Isabelle\<close> (see also \secref{sec:options}).
 
   The options allow to specify a logic session name, but the same selector is
   also accessible in the \<^emph>\<open>Theories\<close> panel (\secref{sec:theories}). After
@@ -1154,7 +1154,7 @@
 text \<open>
   The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Find Theorems\<close> mode retrieves facts from the theory
   or proof context matching all of given criteria in the \<^emph>\<open>Find\<close> text field. A
-  single criterium has the following syntax:
+  single criterion has the following syntax:
 
   @{rail \<open>
     ('-'?) ('name' ':' @{syntax name} | 'intro' | 'elim' | 'dest' |
@@ -1171,7 +1171,7 @@
 text \<open>
   The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Find Constants\<close> mode prints all constants whose type
   meets all of the given criteria in the \<^emph>\<open>Find\<close> text field. A single
-  criterium has the following syntax:
+  criterion has the following syntax:
 
   @{rail \<open>
     ('-'?)
--- a/src/Doc/System/Environment.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/Doc/System/Environment.thy	Thu May 03 15:07:14 2018 +0200
@@ -118,38 +118,37 @@
   \<^descr>[@{setting_def ISABELLE_PLATFORM_FAMILY}\<open>\<^sup>*\<close>] is automatically set to the
   general platform family: \<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>macos\<close>, \<^verbatim>\<open>windows\<close>. Note that
   platform-dependent tools usually need to refer to the more specific
-  identification according to @{setting ISABELLE_PLATFORM} etc.
+  identification according to @{setting ISABELLE_PLATFORM64}, @{setting
+  ISABELLE_PLATFORM32}, @{setting ISABELLE_WINDOWS_PLATFORM64}, @{setting
+  ISABELLE_WINDOWS_PLATFORM32}.
 
-  \<^descr>[@{setting_def ISABELLE_PLATFORM32}\<open>\<^sup>*\<close>, @{setting_def
-  ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>, @{setting_def ISABELLE_PLATFORM}\<open>\<^sup>*\<close>] indicate the
-  standard Posix platform: \<^verbatim>\<open>x86\<close> for 32 bit and \<^verbatim>\<open>x86_64\<close> for 64 bit,
-  together with a symbolic name for the operating system (\<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>darwin\<close>,
-  \<^verbatim>\<open>cygwin\<close>). Some platforms support both 32 bit and 64 bit executables, but
-  this depends on various side-conditions.
+  \<^descr>[@{setting_def ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>, @{setting_def
+  ISABELLE_PLATFORM32}\<open>\<^sup>*\<close>] indicate the standard Posix platform: \<^verbatim>\<open>x86_64\<close>
+  for 64 bit and \<^verbatim>\<open>x86\<close> for 32 bit, together with a symbolic name for the
+  operating system (\<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>darwin\<close>, \<^verbatim>\<open>cygwin\<close>). All platforms support 64
+  bit executables, some platforms also support 32 bit executables.
 
-  In GNU bash scripts, it is possible to use the following expressions
-  (including the quotes) to specify a preference of 64 bit over 32 bit:
+  In GNU bash scripts, it is possible to use the following expressions (with
+  quotes) to specify a preference of 64 bit over 32 bit:
 
   @{verbatim [display] \<open>"${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}"\<close>}
 
-  In contrast, the subsequent expression prefers the 32 bit variant; this is
-  how @{setting ISABELLE_PLATFORM} is defined:
+  In contrast, the subsequent expression prefers the old 32 bit variant (which
+  is only relevant for unusual applications):
 
   @{verbatim [display] \<open>"${ISABELLE_PLATFORM32:-$ISABELLE_PLATFORM64}"\<close>}
 
-  \<^descr>[@{setting_def ISABELLE_WINDOWS_PLATFORM32}\<open>\<^sup>*\<close>, @{setting_def
-  ISABELLE_WINDOWS_PLATFORM64}\<open>\<^sup>*\<close>,] @{setting_def
-  ISABELLE_WINDOWS_PLATFORM}\<open>\<^sup>*\<close> indicate the native Windows platform. These
-  settings are analogous (but independent) of those for the standard Posix
-  subsystem: @{setting ISABELLE_PLATFORM32}, @{setting ISABELLE_PLATFORM64},
-  @{setting ISABELLE_PLATFORM}.
+  \<^descr>[@{setting_def ISABELLE_WINDOWS_PLATFORM64}\<open>\<^sup>*\<close>, @{setting_def
+  ISABELLE_WINDOWS_PLATFORM32}\<open>\<^sup>*\<close>] indicate the native Windows platform.
+  These settings are analogous (but independent) of those for the standard
+  Posix subsystem: @{setting ISABELLE_PLATFORM64}, @{setting
+  ISABELLE_PLATFORM32}.
 
   In GNU bash scripts, a preference for native Windows platform variants may
-  be specified like this:
+  be specified like this (first 64 bit, second 32 bit):
 
-  @{verbatim [display] \<open>"${ISABELLE_WINDOWS_PLATFORM:-$ISABELLE_PLATFORM}"\<close>}
-
-  @{verbatim [display] \<open>"${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"\<close>}
+  @{verbatim [display] \<open>"${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-
+  ${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"\<close>}
 
   \<^descr>[@{setting ISABELLE_TOOL}\<open>\<^sup>*\<close>] is automatically set to the full path name
   of the @{executable isabelle} executable.
--- a/src/HOL/Algebra/Divisibility.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Algebra/Divisibility.thy	Thu May 03 15:07:14 2018 +0200
@@ -2491,11 +2491,7 @@
   have "a' \<in> carrier G \<and> a' gcdof b c"
     apply (simp add: gcdof_greatestLower carr')
     apply (subst greatest_Lower_cong_l[of _ a])
-        apply (simp add: a'a)
-       apply (simp add: carr)
-      apply (simp add: carr)
-     apply (simp add: carr)
-    apply (simp add: gcdof_greatestLower[symmetric] agcd carr)
+        apply (simp_all add: a'a carr gcdof_greatestLower[symmetric] agcd)
     done
   then show ?thesis ..
 qed
--- a/src/HOL/Algebra/Order.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Algebra/Order.thy	Thu May 03 15:07:14 2018 +0200
@@ -30,34 +30,33 @@
 
 locale weak_partial_order = equivalence L for L (structure) +
   assumes le_refl [intro, simp]:
-      "x \<in> carrier L ==> x \<sqsubseteq> x"
+      "x \<in> carrier L \<Longrightarrow> x \<sqsubseteq> x"
     and weak_le_antisym [intro]:
-      "[| x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L |] ==> x .= y"
+      "\<lbrakk>x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x .= y"
     and le_trans [trans]:
-      "[| x \<sqsubseteq> y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L |] ==> x \<sqsubseteq> z"
+      "\<lbrakk>x \<sqsubseteq> y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
     and le_cong:
-      "\<lbrakk> x .= y; z .= w; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L; w \<in> carrier L \<rbrakk> \<Longrightarrow>
+      "\<lbrakk>x .= y; z .= w; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L; w \<in> carrier L\<rbrakk> \<Longrightarrow>
       x \<sqsubseteq> z \<longleftrightarrow> y \<sqsubseteq> w"
 
 definition
   lless :: "[_, 'a, 'a] => bool" (infixl "\<sqsubset>\<index>" 50)
   where "x \<sqsubset>\<^bsub>L\<^esub> y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> y \<and> x .\<noteq>\<^bsub>L\<^esub> y"
 
-
 subsubsection \<open>The order relation\<close>
 
 context weak_partial_order
 begin
 
 lemma le_cong_l [intro, trans]:
-  "\<lbrakk> x .= y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
+  "\<lbrakk>x .= y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
   by (auto intro: le_cong [THEN iffD2])
 
 lemma le_cong_r [intro, trans]:
-  "\<lbrakk> x \<sqsubseteq> y; y .= z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
+  "\<lbrakk>x \<sqsubseteq> y; y .= z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
   by (auto intro: le_cong [THEN iffD1])
 
-lemma weak_refl [intro, simp]: "\<lbrakk> x .= y; x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y"
+lemma weak_refl [intro, simp]: "\<lbrakk>x .= y; x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y"
   by (simp add: le_cong_l)
 
 end
@@ -142,93 +141,86 @@
   Lower :: "[_, 'a set] => 'a set"
   where "Lower L A = {l. (\<forall>x. x \<in> A \<inter> carrier L \<longrightarrow> l \<sqsubseteq>\<^bsub>L\<^esub> x)} \<inter> carrier L"
 
-lemma Upper_closed [intro!, simp]:
+lemma Lower_dual [simp]:
+  "Lower (inv_gorder L) A = Upper L A"
+  by (simp add:Upper_def Lower_def)
+
+lemma Upper_dual [simp]:
+  "Upper (inv_gorder L) A = Lower L A"
+  by (simp add:Upper_def Lower_def)
+
+lemma (in weak_partial_order) equivalence_dual: "equivalence (inv_gorder L)"
+  by (rule equivalence.intro) (auto simp: intro: sym trans)
+
+lemma  (in weak_partial_order) dual_weak_order: "weak_partial_order (inv_gorder L)"
+  by intro_locales (auto simp add: weak_partial_order_axioms_def le_cong intro: equivalence_dual le_trans)
+
+lemma (in weak_partial_order) dual_eq_iff [simp]: "A {.=}\<^bsub>inv_gorder L\<^esub> A' \<longleftrightarrow> A {.=} A'"
+  by (auto simp: set_eq_def elem_def)
+
+lemma dual_weak_order_iff:
+  "weak_partial_order (inv_gorder A) \<longleftrightarrow> weak_partial_order A"
+proof
+  assume "weak_partial_order (inv_gorder A)"
+  then interpret dpo: weak_partial_order "inv_gorder A"
+  rewrites "carrier (inv_gorder A) = carrier A"
+  and   "le (inv_gorder A)      = (\<lambda> x y. le A y x)"
+  and   "eq (inv_gorder A)      = eq A"
+    by (simp_all)
+  show "weak_partial_order A"
+    by (unfold_locales, auto intro: dpo.sym dpo.trans dpo.le_trans)
+next
+  assume "weak_partial_order A"
+  thus "weak_partial_order (inv_gorder A)"
+    by (metis weak_partial_order.dual_weak_order)
+qed
+
+lemma Upper_closed [iff]:
   "Upper L A \<subseteq> carrier L"
   by (unfold Upper_def) clarify
 
 lemma Upper_memD [dest]:
   fixes L (structure)
-  shows "[| u \<in> Upper L A; x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> u \<and> u \<in> carrier L"
+  shows "\<lbrakk>u \<in> Upper L A; x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u \<and> u \<in> carrier L"
   by (unfold Upper_def) blast
 
 lemma (in weak_partial_order) Upper_elemD [dest]:
-  "[| u .\<in> Upper L A; u \<in> carrier L; x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> u"
+  "\<lbrakk>u .\<in> Upper L A; u \<in> carrier L; x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u"
   unfolding Upper_def elem_def
   by (blast dest: sym)
 
 lemma Upper_memI:
   fixes L (structure)
-  shows "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> carrier L |] ==> x \<in> Upper L A"
+  shows "\<lbrakk>!! y. y \<in> A \<Longrightarrow> y \<sqsubseteq> x; x \<in> carrier L\<rbrakk> \<Longrightarrow> x \<in> Upper L A"
   by (unfold Upper_def) blast
 
 lemma (in weak_partial_order) Upper_elemI:
-  "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> carrier L |] ==> x .\<in> Upper L A"
+  "\<lbrakk>!! y. y \<in> A \<Longrightarrow> y \<sqsubseteq> x; x \<in> carrier L\<rbrakk> \<Longrightarrow> x .\<in> Upper L A"
   unfolding Upper_def by blast
 
 lemma Upper_antimono:
-  "A \<subseteq> B ==> Upper L B \<subseteq> Upper L A"
+  "A \<subseteq> B \<Longrightarrow> Upper L B \<subseteq> Upper L A"
   by (unfold Upper_def) blast
 
 lemma (in weak_partial_order) Upper_is_closed [simp]:
-  "A \<subseteq> carrier L ==> is_closed (Upper L A)"
+  "A \<subseteq> carrier L \<Longrightarrow> is_closed (Upper L A)"
   by (rule is_closedI) (blast intro: Upper_memI)+
 
 lemma (in weak_partial_order) Upper_mem_cong:
-  assumes a'carr: "a' \<in> carrier L" and Acarr: "A \<subseteq> carrier L"
-    and aa': "a .= a'"
-    and aelem: "a \<in> Upper L A"
+  assumes  "a' \<in> carrier L" "A \<subseteq> carrier L" "a .= a'" "a \<in> Upper L A"
   shows "a' \<in> Upper L A"
-proof (rule Upper_memI[OF _ a'carr])
-  fix y
-  assume yA: "y \<in> A"
-  hence "y \<sqsubseteq> a" by (intro Upper_memD[OF aelem, THEN conjunct1] Acarr)
-  also note aa'
-  finally
-      show "y \<sqsubseteq> a'"
-      by (simp add: a'carr subsetD[OF Acarr yA] subsetD[OF Upper_closed aelem])
-qed
+  by (metis assms Upper_closed Upper_is_closed closure_of_eq complete_classes)
+
+lemma (in weak_partial_order) Upper_semi_cong:
+  assumes "A \<subseteq> carrier L" "A {.=} A'"
+  shows "Upper L A \<subseteq> Upper L A'"
+  unfolding Upper_def
+   by clarsimp (meson assms equivalence.refl equivalence_axioms le_cong set_eqD2 subset_eq)
 
 lemma (in weak_partial_order) Upper_cong:
-  assumes Acarr: "A \<subseteq> carrier L" and A'carr: "A' \<subseteq> carrier L"
-    and AA': "A {.=} A'"
+  assumes "A \<subseteq> carrier L" "A' \<subseteq> carrier L" "A {.=} A'"
   shows "Upper L A = Upper L A'"
-unfolding Upper_def
-apply rule
- apply (rule, clarsimp) defer 1
- apply (rule, clarsimp) defer 1
-proof -
-  fix x a'
-  assume carr: "x \<in> carrier L" "a' \<in> carrier L"
-    and a'A': "a' \<in> A'"
-  assume aLxCond[rule_format]: "\<forall>a. a \<in> A \<and> a \<in> carrier L \<longrightarrow> a \<sqsubseteq> x"
-
-  from AA' and a'A' have "\<exists>a\<in>A. a' .= a" by (rule set_eqD2)
-  from this obtain a
-      where aA: "a \<in> A"
-      and a'a: "a' .= a"
-      by auto
-  note [simp] = subsetD[OF Acarr aA] carr
-
-  note a'a
-  also have "a \<sqsubseteq> x" by (simp add: aLxCond aA)
-  finally show "a' \<sqsubseteq> x" by simp
-next
-  fix x a
-  assume carr: "x \<in> carrier L" "a \<in> carrier L"
-    and aA: "a \<in> A"
-  assume a'LxCond[rule_format]: "\<forall>a'. a' \<in> A' \<and> a' \<in> carrier L \<longrightarrow> a' \<sqsubseteq> x"
-
-  from AA' and aA have "\<exists>a'\<in>A'. a .= a'" by (rule set_eqD1)
-  from this obtain a'
-      where a'A': "a' \<in> A'"
-      and aa': "a .= a'"
-      by auto
-  note [simp] = subsetD[OF A'carr a'A'] carr
-
-  note aa'
-  also have "a' \<sqsubseteq> x" by (simp add: a'LxCond a'A')
-  finally show "a \<sqsubseteq> x" by simp
-qed
+  using assms by (simp add: Upper_semi_cong set_eq_sym subset_antisym)
 
 lemma Lower_closed [intro!, simp]:
   "Lower L A \<subseteq> carrier L"
@@ -236,16 +228,16 @@
 
 lemma Lower_memD [dest]:
   fixes L (structure)
-  shows "[| l \<in> Lower L A; x \<in> A; A \<subseteq> carrier L |] ==> l \<sqsubseteq> x \<and> l \<in> carrier L"
+  shows "\<lbrakk>l \<in> Lower L A; x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> l \<sqsubseteq> x \<and> l \<in> carrier L"
   by (unfold Lower_def) blast
 
 lemma Lower_memI:
   fixes L (structure)
-  shows "[| !! y. y \<in> A ==> x \<sqsubseteq> y; x \<in> carrier L |] ==> x \<in> Lower L A"
+  shows "\<lbrakk>!! y. y \<in> A \<Longrightarrow> x \<sqsubseteq> y; x \<in> carrier L\<rbrakk> \<Longrightarrow> x \<in> Lower L A"
   by (unfold Lower_def) blast
 
 lemma Lower_antimono:
-  "A \<subseteq> B ==> Lower L B \<subseteq> Lower L A"
+  "A \<subseteq> B \<Longrightarrow> Lower L B \<subseteq> Lower L A"
   by (unfold Lower_def) blast
 
 lemma (in weak_partial_order) Lower_is_closed [simp]:
@@ -253,56 +245,15 @@
   by (rule is_closedI) (blast intro: Lower_memI dest: sym)+
 
 lemma (in weak_partial_order) Lower_mem_cong:
-  assumes a'carr: "a' \<in> carrier L" and Acarr: "A \<subseteq> carrier L"
-    and aa': "a .= a'"
-    and aelem: "a \<in> Lower L A"
+  assumes "a' \<in> carrier L"  "A \<subseteq> carrier L" "a .= a'" "a \<in> Lower L A"
   shows "a' \<in> Lower L A"
-using assms Lower_closed[of L A]
-by (intro Lower_memI) (blast intro: le_cong_l[OF aa'[symmetric]])
+  by (meson assms Lower_closed Lower_is_closed is_closed_eq subsetCE)
 
 lemma (in weak_partial_order) Lower_cong:
-  assumes Acarr: "A \<subseteq> carrier L" and A'carr: "A' \<subseteq> carrier L"
-    and AA': "A {.=} A'"
+  assumes "A \<subseteq> carrier L" "A' \<subseteq> carrier L" "A {.=} A'"
   shows "Lower L A = Lower L A'"
-unfolding Lower_def
-apply rule
- apply clarsimp defer 1
- apply clarsimp defer 1
-proof -
-  fix x a'
-  assume carr: "x \<in> carrier L" "a' \<in> carrier L"
-    and a'A': "a' \<in> A'"
-  assume "\<forall>a. a \<in> A \<and> a \<in> carrier L \<longrightarrow> x \<sqsubseteq> a"
-  hence aLxCond: "\<And>a. \<lbrakk>a \<in> A; a \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> a" by fast
-
-  from AA' and a'A' have "\<exists>a\<in>A. a' .= a" by (rule set_eqD2)
-  from this obtain a
-      where aA: "a \<in> A"
-      and a'a: "a' .= a"
-      by auto
-
-  from aA and subsetD[OF Acarr aA]
-      have "x \<sqsubseteq> a" by (rule aLxCond)
-  also note a'a[symmetric]
-  finally
-      show "x \<sqsubseteq> a'" by (simp add: carr subsetD[OF Acarr aA])
-next
-  fix x a
-  assume carr: "x \<in> carrier L" "a \<in> carrier L"
-    and aA: "a \<in> A"
-  assume "\<forall>a'. a' \<in> A' \<and> a' \<in> carrier L \<longrightarrow> x \<sqsubseteq> a'"
-  hence a'LxCond: "\<And>a'. \<lbrakk>a' \<in> A'; a' \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> a'" by fast+
-
-  from AA' and aA have "\<exists>a'\<in>A'. a .= a'" by (rule set_eqD1)
-  from this obtain a'
-      where a'A': "a' \<in> A'"
-      and aa': "a .= a'"
-      by auto
-  from a'A' and subsetD[OF A'carr a'A']
-      have "x \<sqsubseteq> a'" by (rule a'LxCond)
-  also note aa'[symmetric]
-  finally show "x \<sqsubseteq> a" by (simp add: carr subsetD[OF A'carr a'A'])
-qed
+  unfolding Upper_dual [symmetric]
+  by (rule weak_partial_order.Upper_cong [OF dual_weak_order]) (simp_all add: assms)
 
 text \<open>Jacobson: Theorem 8.1\<close>
 
@@ -325,29 +276,37 @@
   greatest :: "[_, 'a, 'a set] => bool"
   where "greatest L g A \<longleftrightarrow> A \<subseteq> carrier L \<and> g \<in> A \<and> (\<forall>x\<in>A. x \<sqsubseteq>\<^bsub>L\<^esub> g)"
 
-text (in weak_partial_order) \<open>Could weaken these to @{term "l \<in> carrier L \<and> l
-  .\<in> A"} and @{term "g \<in> carrier L \<and> g .\<in> A"}.\<close>
+text (in weak_partial_order) \<open>Could weaken these to @{term "l \<in> carrier L \<and> l .\<in> A"} and @{term "g \<in> carrier L \<and> g .\<in> A"}.\<close>
+
+lemma least_dual [simp]:
+  "least (inv_gorder L) x A = greatest L x A"
+  by (simp add:least_def greatest_def)
+
+lemma greatest_dual [simp]:
+  "greatest (inv_gorder L) x A = least L x A"
+  by (simp add:least_def greatest_def)
 
 lemma least_closed [intro, simp]:
-  "least L l A ==> l \<in> carrier L"
+  "least L l A \<Longrightarrow> l \<in> carrier L"
   by (unfold least_def) fast
 
 lemma least_mem:
-  "least L l A ==> l \<in> A"
+  "least L l A \<Longrightarrow> l \<in> A"
   by (unfold least_def) fast
 
 lemma (in weak_partial_order) weak_least_unique:
-  "[| least L x A; least L y A |] ==> x .= y"
+  "\<lbrakk>least L x A; least L y A\<rbrakk> \<Longrightarrow> x .= y"
   by (unfold least_def) blast
 
 lemma least_le:
   fixes L (structure)
-  shows "[| least L x A; a \<in> A |] ==> x \<sqsubseteq> a"
+  shows "\<lbrakk>least L x A; a \<in> A\<rbrakk> \<Longrightarrow> x \<sqsubseteq> a"
   by (unfold least_def) fast
 
 lemma (in weak_partial_order) least_cong:
-  "[| x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A |] ==> least L x A = least L x' A"
-  by (unfold least_def) (auto dest: sym)
+  "\<lbrakk>x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A\<rbrakk> \<Longrightarrow> least L x A = least L x' A"
+  unfolding least_def
+  by (meson is_closed_eq is_closed_eq_rev le_cong local.refl subset_iff)
 
 abbreviation is_lub :: "[_, 'a, 'a set] => bool"
 where "is_lub L x A \<equiv> least L x (Upper L A)"
@@ -363,16 +322,14 @@
   apply (rule least_cong) using assms by auto
 
 lemma (in weak_partial_order) least_Upper_cong_r:
-  assumes Acarrs: "A \<subseteq> carrier L" "A' \<subseteq> carrier L" (* unneccessary with current Upper? *)
-    and AA': "A {.=} A'"
+  assumes "A \<subseteq> carrier L" "A' \<subseteq> carrier L" "A {.=} A'"
   shows "least L x (Upper L A) = least L x (Upper L A')"
-apply (subgoal_tac "Upper L A = Upper L A'", simp)
-by (rule Upper_cong) fact+
+  using Upper_cong assms by auto
 
 lemma least_UpperI:
   fixes L (structure)
-  assumes above: "!! x. x \<in> A ==> x \<sqsubseteq> s"
-    and below: "!! y. y \<in> Upper L A ==> s \<sqsubseteq> y"
+  assumes above: "!! x. x \<in> A \<Longrightarrow> x \<sqsubseteq> s"
+    and below: "!! y. y \<in> Upper L A \<Longrightarrow> s \<sqsubseteq> y"
     and L: "A \<subseteq> carrier L"  "s \<in> carrier L"
   shows "least L s (Upper L A)"
 proof -
@@ -384,30 +341,31 @@
 
 lemma least_Upper_above:
   fixes L (structure)
-  shows "[| least L s (Upper L A); x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> s"
+  shows "\<lbrakk>least L s (Upper L A); x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> s"
   by (unfold least_def) blast
 
 lemma greatest_closed [intro, simp]:
-  "greatest L l A ==> l \<in> carrier L"
+  "greatest L l A \<Longrightarrow> l \<in> carrier L"
   by (unfold greatest_def) fast
 
 lemma greatest_mem:
-  "greatest L l A ==> l \<in> A"
+  "greatest L l A \<Longrightarrow> l \<in> A"
   by (unfold greatest_def) fast
 
 lemma (in weak_partial_order) weak_greatest_unique:
-  "[| greatest L x A; greatest L y A |] ==> x .= y"
+  "\<lbrakk>greatest L x A; greatest L y A\<rbrakk> \<Longrightarrow> x .= y"
   by (unfold greatest_def) blast
 
 lemma greatest_le:
   fixes L (structure)
-  shows "[| greatest L x A; a \<in> A |] ==> a \<sqsubseteq> x"
+  shows "\<lbrakk>greatest L x A; a \<in> A\<rbrakk> \<Longrightarrow> a \<sqsubseteq> x"
   by (unfold greatest_def) fast
 
 lemma (in weak_partial_order) greatest_cong:
-  "[| x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A |] ==>
+  "\<lbrakk>x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A\<rbrakk> \<Longrightarrow>
   greatest L x A = greatest L x' A"
-  by (unfold greatest_def) (auto dest: sym)
+  unfolding greatest_def
+  by (meson is_closed_eq_rev le_cong_r local.sym subset_eq)
 
 abbreviation is_glb :: "[_, 'a, 'a set] => bool"
 where "is_glb L x A \<equiv> greatest L x (Lower L A)"
@@ -418,21 +376,23 @@
 lemma (in weak_partial_order) greatest_Lower_cong_l:
   assumes "x .= x'"
     and "x \<in> carrier L" "x' \<in> carrier L"
-    and "A \<subseteq> carrier L" (* unneccessary with current Lower *)
   shows "greatest L x (Lower L A) = greatest L x' (Lower L A)"
-  apply (rule greatest_cong) using assms by auto
+proof -
+  have "\<forall>A. is_closed (Lower L (A \<inter> carrier L))"
+    by simp
+  then show ?thesis
+    by (simp add: Lower_def assms greatest_cong)
+qed
 
 lemma (in weak_partial_order) greatest_Lower_cong_r:
-  assumes Acarrs: "A \<subseteq> carrier L" "A' \<subseteq> carrier L"
-    and AA': "A {.=} A'"
+  assumes "A \<subseteq> carrier L" "A' \<subseteq> carrier L" "A {.=} A'"
   shows "greatest L x (Lower L A) = greatest L x (Lower L A')"
-apply (subgoal_tac "Lower L A = Lower L A'", simp)
-by (rule Lower_cong) fact+
+  using Lower_cong assms by auto
 
 lemma greatest_LowerI:
   fixes L (structure)
-  assumes below: "!! x. x \<in> A ==> i \<sqsubseteq> x"
-    and above: "!! y. y \<in> Lower L A ==> y \<sqsubseteq> i"
+  assumes below: "!! x. x \<in> A \<Longrightarrow> i \<sqsubseteq> x"
+    and above: "!! y. y \<in> Lower L A \<Longrightarrow> y \<sqsubseteq> i"
     and L: "A \<subseteq> carrier L"  "i \<in> carrier L"
   shows "greatest L i (Lower L A)"
 proof -
@@ -444,53 +404,9 @@
 
 lemma greatest_Lower_below:
   fixes L (structure)
-  shows "[| greatest L i (Lower L A); x \<in> A; A \<subseteq> carrier L |] ==> i \<sqsubseteq> x"
+  shows "\<lbrakk>greatest L i (Lower L A); x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> i \<sqsubseteq> x"
   by (unfold greatest_def) blast
 
-lemma Lower_dual [simp]:
-  "Lower (inv_gorder L) A = Upper L A"
-  by (simp add:Upper_def Lower_def)
-
-lemma Upper_dual [simp]:
-  "Upper (inv_gorder L) A = Lower L A"
-  by (simp add:Upper_def Lower_def)
-
-lemma least_dual [simp]:
-  "least (inv_gorder L) x A = greatest L x A"
-  by (simp add:least_def greatest_def)
-
-lemma greatest_dual [simp]:
-  "greatest (inv_gorder L) x A = least L x A"
-  by (simp add:least_def greatest_def)
-
-lemma (in weak_partial_order) dual_weak_order:
-  "weak_partial_order (inv_gorder L)"
-  apply (unfold_locales)
-  apply (simp_all)
-  apply (metis sym)
-  apply (metis trans)
-  apply (metis weak_le_antisym)
-  apply (metis le_trans)
-  apply (metis le_cong_l le_cong_r sym)
-done
-
-lemma dual_weak_order_iff:
-  "weak_partial_order (inv_gorder A) \<longleftrightarrow> weak_partial_order A"
-proof
-  assume "weak_partial_order (inv_gorder A)"
-  then interpret dpo: weak_partial_order "inv_gorder A"
-  rewrites "carrier (inv_gorder A) = carrier A"
-  and   "le (inv_gorder A)      = (\<lambda> x y. le A y x)"
-  and   "eq (inv_gorder A)      = eq A"
-    by (simp_all)
-  show "weak_partial_order A"
-    by (unfold_locales, auto intro: dpo.sym dpo.trans dpo.le_trans)
-next
-  assume "weak_partial_order A"
-  thus "weak_partial_order (inv_gorder A)"
-    by (metis weak_partial_order.dual_weak_order)
-qed
-
 
 subsubsection \<open>Intervals\<close>
 
@@ -513,7 +429,7 @@
     by (auto simp add: at_least_at_most_def)
 
   lemma at_least_at_most_member [intro]: 
-    "\<lbrakk> x \<in> carrier L; a \<sqsubseteq> x; x \<sqsubseteq> b \<rbrakk> \<Longrightarrow> x \<in> \<lbrace>a..b\<rbrace>"
+    "\<lbrakk>x \<in> carrier L; a \<sqsubseteq> x; x \<sqsubseteq> b\<rbrakk> \<Longrightarrow> x \<in> \<lbrace>a..b\<rbrace>"
     by (simp add: at_least_at_most_def)
 
 end
@@ -531,7 +447,7 @@
   fixes f :: "'a \<Rightarrow> 'b"
   assumes "weak_partial_order L1"
           "weak_partial_order L2"
-          "(\<And>x y. \<lbrakk> x \<in> carrier L1; y \<in> carrier L1; x \<sqsubseteq>\<^bsub>L1\<^esub> y \<rbrakk> 
+          "(\<And>x y. \<lbrakk>x \<in> carrier L1; y \<in> carrier L1; x \<sqsubseteq>\<^bsub>L1\<^esub> y\<rbrakk> 
                    \<Longrightarrow> f x \<sqsubseteq>\<^bsub>L2\<^esub> f y)"
   shows "isotone L1 L2 f"
   using assms by (auto simp add:isotone_def)
@@ -566,7 +482,7 @@
   "idempotent L f \<equiv> \<forall>x\<in>carrier L. f (f x) .=\<^bsub>L\<^esub> f x"
 
 lemma (in weak_partial_order) idempotent:
-  "\<lbrakk> Idem f; x \<in> carrier L \<rbrakk> \<Longrightarrow> f (f x) .= f x"
+  "\<lbrakk>Idem f; x \<in> carrier L\<rbrakk> \<Longrightarrow> f (f x) .= f x"
   by (auto simp add: idempotent_def)
 
 
@@ -596,7 +512,7 @@
 declare weak_le_antisym [rule del]
 
 lemma le_antisym [intro]:
-  "[| x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L |] ==> x = y"
+  "\<lbrakk>x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x = y"
   using weak_le_antisym unfolding eq_is_equal .
 
 lemma lless_eq:
@@ -627,8 +543,8 @@
   and   "eq (inv_gorder A)      = eq A"
     by (simp_all)
   show "partial_order A"
-    apply (unfold_locales, simp_all)
-    apply (metis po.sym, metis po.trans)
+    apply (unfold_locales, simp_all add: po.sym)
+    apply (metis po.trans)
     apply (metis po.weak_le_antisym, metis po.le_trans)
     apply (metis (full_types) po.eq_is_equal, metis po.eq_is_equal)
   done
@@ -641,11 +557,11 @@
 text \<open>Least and greatest, as predicate\<close>
 
 lemma (in partial_order) least_unique:
-  "[| least L x A; least L y A |] ==> x = y"
+  "\<lbrakk>least L x A; least L y A\<rbrakk> \<Longrightarrow> x = y"
   using weak_least_unique unfolding eq_is_equal .
 
 lemma (in partial_order) greatest_unique:
-  "[| greatest L x A; greatest L y A |] ==> x = y"
+  "\<lbrakk>greatest L x A; greatest L y A\<rbrakk> \<Longrightarrow> x = y"
   using weak_greatest_unique unfolding eq_is_equal .
 
 
@@ -709,12 +625,12 @@
 subsection \<open>Total Orders\<close>
 
 locale weak_total_order = weak_partial_order +
-  assumes total: "\<lbrakk> x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
+  assumes total: "\<lbrakk>x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
 
 text \<open>Introduction rule: the usual definition of total order\<close>
 
 lemma (in weak_partial_order) weak_total_orderI:
-  assumes total: "!!x y. \<lbrakk> x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
+  assumes total: "!!x y. \<lbrakk>x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
   shows "weak_total_order L"
   by unfold_locales (rule total)
 
@@ -722,7 +638,7 @@
 subsection \<open>Total orders where \<open>eq\<close> is the Equality\<close>
 
 locale total_order = partial_order +
-  assumes total_order_total: "\<lbrakk> x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
+  assumes total_order_total: "\<lbrakk>x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
 
 sublocale total_order < weak?: weak_total_order
   by unfold_locales (rule total_order_total)
@@ -730,7 +646,7 @@
 text \<open>Introduction rule: the usual definition of total order\<close>
 
 lemma (in partial_order) total_orderI:
-  assumes total: "!!x y. \<lbrakk> x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
+  assumes total: "!!x y. \<lbrakk>x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
   shows "total_order L"
   by unfold_locales (rule total)
 
--- a/src/HOL/Analysis/Bochner_Integration.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Bochner_Integration.thy	Thu May 03 15:07:14 2018 +0200
@@ -475,8 +475,8 @@
     by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real])
        (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases
              intro!: sum.cong ennreal_cong_mult
-             simp: sum_ennreal[symmetric] ac_simps ennreal_mult
-             simp del: sum_ennreal)
+             simp: ac_simps ennreal_mult
+             reorient: sum_ennreal)
   also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
     using f
     by (intro nn_integral_eq_simple_integral[symmetric])
@@ -504,7 +504,7 @@
     using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t
     by (auto intro!: simple_bochner_integral_eq_nn_integral)
   also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)"
-    by (auto intro!: nn_integral_mono simp: ennreal_plus[symmetric] simp del: ennreal_plus)
+    by (auto intro!: nn_integral_mono reorient: ennreal_plus)
        (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans
               norm_minus_commute norm_triangle_ineq4 order_refl)
   also have "\<dots> = ?S + ?T"
@@ -594,7 +594,7 @@
     proof (intro always_eventually allI)
       fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \<partial>M)"
         by (auto intro!: nn_integral_mono norm_diff_triangle_ineq
-                 simp del: ennreal_plus simp add: ennreal_plus[symmetric])
+                 reorient: ennreal_plus)
       also have "\<dots> = ?g i"
         by (intro nn_integral_add) auto
       finally show "?f i \<le> ?g i" .
@@ -747,7 +747,7 @@
   finally have s_fin: "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" .
 
   have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \<partial>M)"
-    by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric])
+    by (auto intro!: nn_integral_mono reorient: ennreal_plus)
        (metis add.commute norm_triangle_sub)
   also have "\<dots> = (\<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) + (\<integral>\<^sup>+x. norm (s i x) \<partial>M)"
     by (rule nn_integral_add) auto
@@ -783,7 +783,7 @@
         by (intro simple_bochner_integral_eq_nn_integral)
            (auto intro: s simple_bochner_integrable_compose2)
       also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \<partial>M)"
-        by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric])
+        by (auto intro!: nn_integral_mono reorient: ennreal_plus)
            (metis add.commute norm_minus_commute norm_triangle_sub)
       also have "\<dots> = ?t n"
         by (rule nn_integral_add) auto
@@ -828,7 +828,7 @@
       using tendsto_add[OF \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>] by simp
   qed
   then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> 0"
-    by (simp add: ennreal_0[symmetric] del: ennreal_0)
+    by (simp reorient: ennreal_0)
   ultimately have "norm (x - y) = 0"
     by (rule LIMSEQ_unique)
   then show "x = y" by simp
@@ -1174,7 +1174,7 @@
         by (intro simple_bochner_integral_bounded s f)
       also have "\<dots> < ennreal (e / 2) + e / 2"
         by (intro add_strict_mono M n m)
-      also have "\<dots> = e" using \<open>0<e\<close> by (simp del: ennreal_plus add: ennreal_plus[symmetric])
+      also have "\<dots> = e" using \<open>0<e\<close> by (simp reorient: ennreal_plus)
       finally show "dist (?s n) (?s m) < e"
         using \<open>0<e\<close> by (simp add: dist_norm ennreal_less_iff)
     qed
@@ -1219,7 +1219,7 @@
       fix x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
       from tendsto_diff[OF tendsto_const[of "u' x"] this]
       show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
-        by (simp add: tendsto_norm_zero_iff ennreal_0[symmetric] del: ennreal_0)
+        by (simp add: tendsto_norm_zero_iff reorient: ennreal_0)
     qed
   qed (insert bnd w_nonneg, auto)
   then show ?thesis by simp
@@ -2117,7 +2117,7 @@
       by auto
   qed
   then have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F"
-    by (simp add: ennreal_0[symmetric] del: ennreal_0)
+    by (simp reorient: ennreal_0)
   then have "((\<lambda>n. ((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" using tendsto_norm_zero_iff by blast
   then show ?thesis using Lim_null by auto
 qed
@@ -2215,7 +2215,7 @@
     ultimately have "(\<lambda>n. ennreal (norm(u (r n) x))) \<longlonglongrightarrow> 0"
       using tendsto_Limsup[of sequentially "\<lambda>n. ennreal (norm(u (r n) x))"] by auto
     then have "(\<lambda>n. norm(u (r n) x)) \<longlonglongrightarrow> 0"
-      by (simp add: ennreal_0[symmetric] del: ennreal_0)
+      by (simp reorient: ennreal_0)
     then have "(\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0"
       by (simp add: tendsto_norm_zero_iff)
   }
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy	Thu May 03 15:07:14 2018 +0200
@@ -2,9 +2,6 @@
     Author:     Robert Himmelmann, TU Muenchen (Translation from HOL light) and LCP
 *)
 
-(* ========================================================================= *)
-(* Results connected with topological dimension.                             *)
-(*                                                                           *)
 (* At the moment this is just Brouwer's fixpoint theorem. The proof is from  *)
 (* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518   *)
 (* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf".          *)
@@ -14,7 +11,6 @@
 (* the big advantage of Kuhn's proof over the usual Sperner's lemma one.     *)
 (*                                                                           *)
 (*              (c) Copyright, John Harrison 1998-2008                       *)
-(* ========================================================================= *)
 
 section \<open>Results connected with topological dimension\<close>
 
@@ -40,11 +36,7 @@
 lemma swap_image:
   "Fun.swap i j f ` A = (if i \<in> A then (if j \<in> A then f ` A else f ` ((A - {i}) \<union> {j}))
                                   else (if j \<in> A then f ` ((A - {j}) \<union> {i}) else f ` A))"
-  apply (auto simp: Fun.swap_def image_iff)
-  apply metis
-  apply (metis member_remove remove_def)
-  apply (metis member_remove remove_def)
-  done
+  by (auto simp: swap_def image_def) metis
 
 lemmas swap_apply1 = swap_apply(1)
 lemmas swap_apply2 = swap_apply(2)
@@ -191,9 +183,9 @@
     moreover obtain a where "rl a = Suc n" "a \<in> s"
       by (metis atMost_iff image_iff le_Suc_eq rl)
     ultimately have n: "{..n} = rl ` (s - {a})"
-      by (auto simp add: inj_on_image_set_diff Diff_subset rl)
+      by (auto simp: inj_on_image_set_diff Diff_subset rl)
     have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a}"
-      using inj_rl \<open>a \<in> s\<close> by (auto simp add: n inj_on_image_eq_iff[OF inj_rl] Diff_subset)
+      using inj_rl \<open>a \<in> s\<close> by (auto simp: n inj_on_image_eq_iff[OF inj_rl] Diff_subset)
     then show "card ?S = 1"
       unfolding card_S by simp }
 
@@ -202,7 +194,7 @@
     proof cases
       assume *: "{..n} \<subseteq> rl ` s"
       with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}"
-        by (auto simp add: atMost_Suc subset_insert_iff split: if_split_asm)
+        by (auto simp: atMost_Suc subset_insert_iff split: if_split_asm)
       then have "\<not> inj_on rl s"
         by (intro pigeonhole) simp
       then obtain a b where ab: "a \<in> s" "b \<in> s" "rl a = rl b" "a \<noteq> b"
@@ -210,7 +202,7 @@
       then have eq: "rl ` (s - {a}) = rl ` s"
         by auto
       with ab have inj: "inj_on rl (s - {a})"
-        by (intro eq_card_imp_inj_on) (auto simp add: rl_s card_Diff_singleton_if)
+        by (intro eq_card_imp_inj_on) (auto simp: rl_s card_Diff_singleton_if)
 
       { fix x assume "x \<in> s" "x \<notin> {a, b}"
         then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})"
@@ -275,7 +267,7 @@
     with upd have "upd ` {..< x} \<noteq> upd ` {..< y}"
       by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def)
     then have "enum x \<noteq> enum y"
-      by (auto simp add: enum_def fun_eq_iff) }
+      by (auto simp: enum_def fun_eq_iff) }
   then show ?thesis
     by (auto simp: inj_on_def)
 qed
@@ -325,7 +317,7 @@
   by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric])
 
 lemma enum_strict_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i < enum j \<longleftrightarrow> i < j"
-  using enum_mono[of i j] enum_inj[of i j] by (auto simp add: le_less)
+  using enum_mono[of i j] enum_inj[of i j] by (auto simp: le_less)
 
 lemma chain: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a \<le> b \<or> b \<le> a"
   by (auto simp: s_eq enum_mono)
@@ -346,7 +338,7 @@
   by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric])
 
 lemma out_eq_p: "a \<in> s \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p"
-  unfolding s_eq by (auto simp add: enum_eq_p)
+  unfolding s_eq by (auto simp: enum_eq_p)
 
 lemma s_le_p: "a \<in> s \<Longrightarrow> a j \<le> p"
   using out_eq_p[of a j] s_space by (cases "j < n") auto
@@ -578,7 +570,7 @@
         by (auto simp: image_iff Ball_def) arith
       then have upd_Suc: "\<And>i. i \<le> n \<Longrightarrow> (upd\<circ>Suc) ` {..< i} = upd ` {..< Suc i} - {n}"
         using \<open>upd 0 = n\<close> upd_inj
-        by (auto simp add: image_comp[symmetric] inj_on_image_set_diff[OF inj_upd])
+        by (auto simp: image_comp[symmetric] inj_on_image_set_diff[OF inj_upd])
       have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}"
         using \<open>upd 0 = n\<close> by auto
 
@@ -685,7 +677,7 @@
 qed
 
 lemma card_2_exists: "card s = 2 \<longleftrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y))"
-  by (auto simp add: card_Suc_eq eval_nat_numeral)
+  by (auto simp: card_Suc_eq eval_nat_numeral)
 
 lemma ksimplex_replace_2:
   assumes s: "ksimplex p n s" and "a \<in> s" and "n \<noteq> 0"
@@ -723,11 +715,11 @@
       obtain i' where "i' \<le> n" "enum i' \<noteq> enum 0" "enum i' (upd 0) \<noteq> p"
         unfolding s_eq by (auto intro: upd_space simp: enum_inj)
       then have "enum 1 \<le> enum i'" "enum i' (upd 0) < p"
-        using enum_le_p[of i' "upd 0"] by (auto simp add: enum_inj enum_mono upd_space)
+        using enum_le_p[of i' "upd 0"] by (auto simp: enum_inj enum_mono upd_space)
       then have "enum 1 (upd 0) < p"
-        by (auto simp add: le_fun_def intro: le_less_trans)
+        by (auto simp: le_fun_def intro: le_less_trans)
       then show "enum (Suc 0) \<in> {..<n} \<rightarrow> {..<p}"
-        using base \<open>n \<noteq> 0\<close> by (auto simp add: enum_0 enum_Suc PiE_iff extensional_def upd_space)
+        using base \<open>n \<noteq> 0\<close> by (auto simp: enum_0 enum_Suc PiE_iff extensional_def upd_space)
 
       { fix i assume "n \<le> i" then show "enum (Suc 0) i = p"
         using \<open>n \<noteq> 0\<close> by (auto simp: enum_eq_p) }
@@ -745,7 +737,7 @@
 
     { fix j assume j: "j < n"
       from j \<open>n \<noteq> 0\<close> have "f' j = enum (Suc j)"
-        by (auto simp add: f'_def enum_def upd_inj in_upd_image image_comp[symmetric] fun_eq_iff) }
+        by (auto simp: f'_def enum_def upd_inj in_upd_image image_comp[symmetric] fun_eq_iff) }
     note f'_eq_enum = this
     then have "enum ` Suc ` {..< n} = f' ` {..< n}"
       by (force simp: enum_inj)
@@ -859,10 +851,10 @@
       by (simp_all add: rot_def)
 
     { fix j assume j: "Suc j \<le> n" then have "b.enum (Suc j) = enum j"
-        by (induct j) (auto simp add: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) }
+        by (induct j) (auto simp: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) }
     note b_enum_eq_enum = this
     then have "enum ` {..< n} = b.enum ` Suc ` {..< n}"
-      by (auto simp add: image_comp intro!: image_cong)
+      by (auto simp: image_comp intro!: image_cong)
     also have "Suc ` {..< n} = {.. n} - {0}"
       by (auto simp: image_iff Ball_def) arith
     also have "{..< n} = {.. n} - {n}"
@@ -871,7 +863,7 @@
       unfolding s_eq \<open>a = enum i\<close> \<open>i = n\<close>
       using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"]
             inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"]
-      by (simp add: comp_def )
+      by (simp add: comp_def)
 
     have "b.enum 0 \<le> b.enum n"
       by (simp add: b.enum_mono)
@@ -956,7 +948,7 @@
       moreover note i
       ultimately have "enum j = b.enum j \<longleftrightarrow> j \<noteq> i"
         unfolding enum_def[abs_def] b.enum_def[abs_def]
-        by (auto simp add: fun_eq_iff swap_image i'_def
+        by (auto simp: fun_eq_iff swap_image i'_def
                            in_upd_image inj_on_image_set_diff[OF inj_upd]) }
     note enum_eq_benum = this
     then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})"
@@ -1001,7 +993,7 @@
         then obtain j where "t.enum (Suc l) = enum j" "j \<le> n" "enum j \<noteq> enum i"
           unfolding s_eq \<open>a = enum i\<close> by auto
         with i have "t.enum (Suc l) \<le> t.enum l \<or> t.enum k \<le> t.enum (Suc l)"
-          by (auto simp add: i'_def enum_mono enum_inj l k)
+          by (auto simp: i'_def enum_mono enum_inj l k)
         with \<open>Suc l < k\<close> \<open>k \<le> n\<close> show False
           by (simp add: t.enum_mono)
       qed
@@ -1041,7 +1033,7 @@
         assume u: "u l = upd (Suc i')"
         define B where "B = b.enum ` {..n}"
         have "b.enum i' = enum i'"
-          using enum_eq_benum[of i'] i by (auto simp add: i'_def gr0_conv_Suc)
+          using enum_eq_benum[of i'] i by (auto simp: i'_def gr0_conv_Suc)
         have "c = t.enum (Suc l)" unfolding c_eq ..
         also have "t.enum (Suc l) = b.enum (Suc i')"
           using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close>
@@ -1432,7 +1424,7 @@
 proof (rule ccontr)
   define n where "n = DIM('a)"
   have n: "1 \<le> n" "0 < n" "n \<noteq> 0"
-    unfolding n_def by (auto simp add: Suc_le_eq DIM_positive)
+    unfolding n_def by (auto simp: Suc_le_eq DIM_positive)
   assume "\<not> ?thesis"
   then have *: "\<not> (\<exists>x\<in>unit_cube. f x - x = 0)"
     by auto
@@ -1447,73 +1439,45 @@
     using assms(2)[unfolded image_subset_iff Ball_def]
     unfolding mem_unit_cube
     by auto
-  obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where
+  obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where label [rule_format]:
     "\<forall>x. \<forall>i\<in>Basis. label x i \<le> 1"
-    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0"
-    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1"
-    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i"
-    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i"
-    using kuhn_labelling_lemma[OF *] by blast
+    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0"
+    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1"
+    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i"
+    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i"
+    using kuhn_labelling_lemma[OF *] by auto
   note label = this [rule_format]
   have lem1: "\<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow>
     \<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)"
   proof safe
     fix x y :: 'a
-    assume x: "x \<in> unit_cube"
-    assume y: "y \<in> unit_cube"
+    assume x: "x \<in> unit_cube" and y: "y \<in> unit_cube"
     fix i
     assume i: "label x i \<noteq> label y i" "i \<in> Basis"
     have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow>
       \<bar>fx - x\<bar> \<le> \<bar>fy - fx\<bar> + \<bar>y - x\<bar>" by auto
     have "\<bar>(f x - x) \<bullet> i\<bar> \<le> \<bar>(f y - f x)\<bullet>i\<bar> + \<bar>(y - x)\<bullet>i\<bar>"
-      unfolding inner_simps
-      apply (rule *)
-      apply (cases "label x i = 0")
-      apply (rule disjI1)
-      apply rule
-      prefer 3
-      apply (rule disjI2)
-      apply rule
-    proof -
-      assume lx: "label x i = 0"
-      then have ly: "label y i = 1"
-        using i label(1)[of i y]
-        by auto
-      show "x \<bullet> i \<le> f x \<bullet> i"
-        apply (rule label(4)[rule_format])
-        using x y lx i(2)
-        apply auto
-        done
-      show "f y \<bullet> i \<le> y \<bullet> i"
-        apply (rule label(5)[rule_format])
-        using x y ly i(2)
-        apply auto
-        done
+    proof (cases "label x i = 0")
+      case True
+      then have fxy: "\<not> f y \<bullet> i \<le> y \<bullet> i \<Longrightarrow> f x \<bullet> i \<le> x \<bullet> i"
+        by (metis True i label(1) label(5) le_antisym less_one not_le_imp_less y)
+      show ?thesis
+      unfolding inner_simps         
+      by (rule *) (auto simp: True i label x y fxy)
     next
-      assume "label x i \<noteq> 0"
-      then have l: "label x i = 1" "label y i = 0"
-        using i label(1)[of i x] label(1)[of i y]
-        by auto
-      show "f x \<bullet> i \<le> x \<bullet> i"
-        apply (rule label(5)[rule_format])
-        using x y l i(2)
-        apply auto
-        done
-      show "y \<bullet> i \<le> f y \<bullet> i"
-        apply (rule label(4)[rule_format])
-        using x y l i(2)
-        apply auto
-        done
+      case False
+      then show ?thesis
+        using label [OF \<open>i \<in> Basis\<close>] i(1) x y
+        apply (auto simp: inner_diff_left le_Suc_eq)
+        by (metis "*")
     qed
     also have "\<dots> \<le> norm (f y - f x) + norm (y - x)"
-      apply (rule add_mono)
-      apply (rule Basis_le_norm[OF i(2)])+
-      done
+      by (simp add: add_mono i(2) norm_bound_Basis_le)
     finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)"
       unfolding inner_simps .
   qed
   have "\<exists>e>0. \<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>z\<in>unit_cube. \<forall>i\<in>Basis.
-    norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow>
+    norm (x - z) < e \<longrightarrow> norm (y - z) < e \<longrightarrow> label x i \<noteq> label y i \<longrightarrow>
       \<bar>(f(z) - z)\<bullet>i\<bar> < d / (real n)"
   proof -
     have d': "d / real n / 8 > 0"
@@ -1530,9 +1494,7 @@
       unfolding dist_norm
       by blast
     show ?thesis
-      apply (rule_tac x="min (e/2) (d/real n/8)" in exI)
-      apply safe
-    proof -
+    proof (intro exI conjI ballI impI)
       show "0 < min (e / 2) (d / real n / 8)"
         using d' e by auto
       fix x y z i
@@ -1551,10 +1513,9 @@
         unfolding inner_simps
       proof (rule *)
         show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)"
-          apply (rule lem1[rule_format])
-          using as i
-          apply auto
-          done
+          using as(1) as(2) as(6) i lem1 by blast
+        show "norm (f x - f z) < d / real n / 8"
+          using d' e as by auto
         show "\<bar>f x \<bullet> i - f z \<bullet> i\<bar> \<le> norm (f x - f z)" "\<bar>x \<bullet> i - z \<bullet> i\<bar> \<le> norm (x - z)"
           unfolding inner_diff_left[symmetric]
           by (rule Basis_le_norm[OF i])+
@@ -1563,30 +1524,14 @@
           unfolding norm_minus_commute
           by auto
         also have "\<dots> < e / 2 + e / 2"
-          apply (rule add_strict_mono)
-          using as(4,5)
-          apply auto
-          done
+          using as(4) as(5) by auto
         finally show "norm (f y - f x) < d / real n / 8"
-          apply -
-          apply (rule e(2))
-          using as
-          apply auto
-          done
+          using as(1) as(2) e(2) by auto
         have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8"
-          apply (rule add_strict_mono)
-          using as
-          apply auto
-          done
-        then show "norm (y - x) < 2 * (d / real n / 8)"
-          using tria
+          using as(4) as(5) by auto
+        with tria show "norm (y - x) < 2 * (d / real n / 8)"
           by auto
-        show "norm (f x - f z) < d / real n / 8"
-          apply (rule e(2))
-          using as e(1)
-          apply auto
-          done
-      qed (insert as, auto)
+      qed (use as in auto)
     qed
   qed
   then
@@ -1635,14 +1580,14 @@
   { fix x :: "nat \<Rightarrow> nat" and i assume "\<forall>i<n. x i \<le> p" "i < n" "x i = p \<or> x i = 0"
     then have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> (unit_cube::'a set)"
       using b'_Basis
-      by (auto simp add: mem_unit_cube inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) }
+      by (auto simp: mem_unit_cube inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) }
   note cube = this
   have q2: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow>
       (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)"
-    unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp add: b'')
+    unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp: b'')
   have q3: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow>
       (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
-    using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp add: b'')
+    using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp: b'')
   obtain q where q:
       "\<forall>i<n. q i < p"
       "\<forall>i<n.
@@ -1660,24 +1605,20 @@
     then have "z \<in> unit_cube"
       unfolding z_def mem_unit_cube
       using b'_Basis
-      by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
+      by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1)
     then have d_fz_z: "d \<le> norm (f z - z)"
       by (rule d)
     assume "\<not> ?thesis"
     then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n"
       using \<open>n > 0\<close>
-      by (auto simp add: not_le inner_diff)
+      by (auto simp: not_le inner_diff)
     have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)"
       unfolding inner_diff_left[symmetric]
       by (rule norm_le_l1)
     also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)"
-      apply (rule sum_strict_mono)
-      using as
-      apply auto
-      done
+      by (meson as finite_Basis nonempty_Basis sum_strict_mono)
     also have "\<dots> = d"
-      using DIM_positive[where 'a='a]
-      by (auto simp: n_def)
+      using DIM_positive[where 'a='a] by (auto simp: n_def)
     finally show False
       using d_fz_z by auto
   qed
@@ -1698,50 +1639,37 @@
     apply (rule order_trans)
     apply (rule rs(1)[OF b'_im,THEN conjunct2])
     using q(1)[rule_format,OF b'_im]
-    apply (auto simp add: Suc_le_eq)
+    apply (auto simp: Suc_le_eq)
     done
   then have "r' \<in> unit_cube"
     unfolding r'_def mem_unit_cube
     using b'_Basis
-    by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
+    by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1)
   define s' :: 'a where "s' = (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)"
   have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p"
-    apply (rule order_trans)
-    apply (rule rs(2)[OF b'_im, THEN conjunct2])
-    using q(1)[rule_format,OF b'_im]
-    apply (auto simp add: Suc_le_eq)
-    done
+    using b'_im q(1) rs(2) by fastforce
   then have "s' \<in> unit_cube"
     unfolding s'_def mem_unit_cube
-    using b'_Basis
-    by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
+    using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1)
   have "z \<in> unit_cube"
     unfolding z_def mem_unit_cube
     using b'_Basis q(1)[rule_format,OF b'_im] \<open>p > 0\<close>
-    by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le)
-  have *: "\<And>x. 1 + real x = real (Suc x)"
-    by auto
+    by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le)
   {
     have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
-      apply (rule sum_mono)
-      using rs(1)[OF b'_im]
-      apply (auto simp add:* field_simps simp del: of_nat_Suc)
-      done
+      by (rule sum_mono) (use rs(1)[OF b'_im] in force)
     also have "\<dots> < e * real p"
       using p \<open>e > 0\<close> \<open>p > 0\<close>
-      by (auto simp add: field_simps n_def)
+      by (auto simp: field_simps n_def)
     finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" .
   }
   moreover
   {
     have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
-      apply (rule sum_mono)
-      using rs(2)[OF b'_im]
-      apply (auto simp add:* field_simps simp del: of_nat_Suc)
-      done
+      by (rule sum_mono) (use rs(2)[OF b'_im] in force)
     also have "\<dots> < e * real p"
       using p \<open>e > 0\<close> \<open>p > 0\<close>
-      by (auto simp add: field_simps n_def)
+      by (auto simp: field_simps n_def)
     finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" .
   }
   ultimately
@@ -1749,7 +1677,7 @@
     unfolding r'_def s'_def z_def
     using \<open>p > 0\<close>
     apply (rule_tac[!] le_less_trans[OF norm_le_l1])
-    apply (auto simp add: field_simps sum_divide_distrib[symmetric] inner_diff_left)
+    apply (auto simp: field_simps sum_divide_distrib[symmetric] inner_diff_left)
     done
   then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
     using rs(3) i
@@ -1762,121 +1690,100 @@
 
 subsection \<open>Retractions\<close>
 
-definition "retraction s t r \<longleftrightarrow> t \<subseteq> s \<and> continuous_on s r \<and> r ` s \<subseteq> t \<and> (\<forall>x\<in>t. r x = x)"
+definition "retraction S T r \<longleftrightarrow> T \<subseteq> S \<and> continuous_on S r \<and> r ` S \<subseteq> T \<and> (\<forall>x\<in>T. r x = x)"
 
 definition retract_of (infixl "retract'_of" 50)
-  where "(t retract_of s) \<longleftrightarrow> (\<exists>r. retraction s t r)"
-
-lemma retraction_idempotent: "retraction s t r \<Longrightarrow> x \<in> s \<Longrightarrow>  r (r x) = r x"
+  where "(T retract_of S) \<longleftrightarrow> (\<exists>r. retraction S T r)"
+
+lemma retraction_idempotent: "retraction S T r \<Longrightarrow> x \<in> S \<Longrightarrow>  r (r x) = r x"
   unfolding retraction_def by auto
 
 subsection \<open>Preservation of fixpoints under (more general notion of) retraction\<close>
 
 lemma invertible_fixpoint_property:
-  fixes s :: "'a::euclidean_space set"
-    and t :: "'b::euclidean_space set"
-  assumes "continuous_on t i"
-    and "i ` t \<subseteq> s"
-    and "continuous_on s r"
-    and "r ` s \<subseteq> t"
-    and "\<forall>y\<in>t. r (i y) = y"
-    and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)"
-    and "continuous_on t g"
-    and "g ` t \<subseteq> t"
-  obtains y where "y \<in> t" and "g y = y"
+  fixes S :: "'a::euclidean_space set"
+    and T :: "'b::euclidean_space set"
+  assumes contt: "continuous_on T i"
+    and "i ` T \<subseteq> S"
+    and contr: "continuous_on S r"
+    and "r ` S \<subseteq> T"
+    and ri: "\<And>y. y \<in> T \<Longrightarrow> r (i y) = y"
+    and FP: "\<And>f. \<lbrakk>continuous_on S f; f ` S \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x"
+    and contg: "continuous_on T g"
+    and "g ` T \<subseteq> T"
+  obtains y where "y \<in> T" and "g y = y"
 proof -
-  have "\<exists>x\<in>s. (i \<circ> g \<circ> r) x = x"
-    apply (rule assms(6)[rule_format])
-    apply rule
-    apply (rule continuous_on_compose assms)+
-    apply ((rule continuous_on_subset)?, rule assms)+
-    using assms(2,4,8)
-    apply auto
-    apply blast
-    done
-  then obtain x where x: "x \<in> s" "(i \<circ> g \<circ> r) x = x" ..
-  then have *: "g (r x) \<in> t"
+  have "\<exists>x\<in>S. (i \<circ> g \<circ> r) x = x"
+  proof (rule FP)
+    show "continuous_on S (i \<circ> g \<circ> r)"
+      by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset)
+    show "(i \<circ> g \<circ> r) ` S \<subseteq> S"
+      using assms(2,4,8) by force
+  qed
+  then obtain x where x: "x \<in> S" "(i \<circ> g \<circ> r) x = x" ..
+  then have *: "g (r x) \<in> T"
     using assms(4,8) by auto
   have "r ((i \<circ> g \<circ> r) x) = r x"
     using x by auto
   then show ?thesis
-    apply (rule_tac that[of "r x"])
-    using x
-    unfolding o_def
-    unfolding assms(5)[rule_format,OF *]
-    using assms(4)
-    apply auto
-    done
+    using "*" ri that by auto
 qed
 
 lemma homeomorphic_fixpoint_property:
-  fixes s :: "'a::euclidean_space set"
-    and t :: "'b::euclidean_space set"
-  assumes "s homeomorphic t"
-  shows "(\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)) \<longleftrightarrow>
-    (\<forall>g. continuous_on t g \<and> g ` t \<subseteq> t \<longrightarrow> (\<exists>y\<in>t. g y = y))"
+  fixes S :: "'a::euclidean_space set"
+    and T :: "'b::euclidean_space set"
+  assumes "S homeomorphic T"
+  shows "(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> S \<longrightarrow> (\<exists>x\<in>S. f x = x)) \<longleftrightarrow>
+         (\<forall>g. continuous_on T g \<and> g ` T \<subseteq> T \<longrightarrow> (\<exists>y\<in>T. g y = y))"
+         (is "?lhs = ?rhs")
 proof -
-  obtain r i where
-      "\<forall>x\<in>s. i (r x) = x"
-      "r ` s = t"
-      "continuous_on s r"
-      "\<forall>y\<in>t. r (i y) = y"
-      "i ` t = s"
-      "continuous_on t i"
-    using assms
-    unfolding homeomorphic_def homeomorphism_def
-    by blast
-  then show ?thesis
-    apply -
-    apply rule
-    apply (rule_tac[!] allI impI)+
-    apply (rule_tac g=g in invertible_fixpoint_property[of t i s r])
-    prefer 10
-    apply (rule_tac g=f in invertible_fixpoint_property[of s r t i])
-    apply auto
-    done
+  obtain r i where r:
+      "\<forall>x\<in>S. i (r x) = x" "r ` S = T" "continuous_on S r"
+      "\<forall>y\<in>T. r (i y) = y" "i ` T = S" "continuous_on T i"
+    using assms unfolding homeomorphic_def homeomorphism_def  by blast
+  show ?thesis
+  proof
+    assume ?lhs
+    with r show ?rhs
+      by (metis invertible_fixpoint_property[of T i S r] order_refl)
+  next
+    assume ?rhs
+    with r show ?lhs
+      by (metis invertible_fixpoint_property[of S r T i] order_refl)
+  qed
 qed
 
 lemma retract_fixpoint_property:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-    and s :: "'a set"
-  assumes "t retract_of s"
-    and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)"
-    and "continuous_on t g"
-    and "g ` t \<subseteq> t"
-  obtains y where "y \<in> t" and "g y = y"
+    and S :: "'a set"
+  assumes "T retract_of S"
+    and FP: "\<And>f. \<lbrakk>continuous_on S f; f ` S \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x"
+    and contg: "continuous_on T g"
+    and "g ` T \<subseteq> T"
+  obtains y where "y \<in> T" and "g y = y"
 proof -
-  obtain h where "retraction s t h"
+  obtain h where "retraction S T h"
     using assms(1) unfolding retract_of_def ..
   then show ?thesis
     unfolding retraction_def
-    apply -
-    apply (rule invertible_fixpoint_property[OF continuous_on_id _ _ _ _ assms(2), of t h g])
-    prefer 7
-    apply (rule_tac y = y in that)
-    using assms
-    apply auto
-    done
+    using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP]
+    by (metis assms(4) contg image_ident that)
 qed
 
 
 subsection \<open>The Brouwer theorem for any set with nonempty interior\<close>
 
 lemma convex_unit_cube: "convex unit_cube"
-  apply (rule is_interval_convex)
-  apply (clarsimp simp add: is_interval_def mem_unit_cube)
-  apply (drule (1) bspec)+
-  apply auto
-  done
+  by (rule is_interval_convex) (fastforce simp add: is_interval_def mem_unit_cube)
 
 lemma brouwer_weak:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
-  assumes "compact s"
-    and "convex s"
-    and "interior s \<noteq> {}"
-    and "continuous_on s f"
-    and "f ` s \<subseteq> s"
-  obtains x where "x \<in> s" and "f x = x"
+  assumes "compact S"
+    and "convex S"
+    and "interior S \<noteq> {}"
+    and "continuous_on S f"
+    and "f ` S \<subseteq> S"
+  obtains x where "x \<in> S" and "f x = x"
 proof -
   let ?U = "unit_cube :: 'a set"
   have "\<Sum>Basis /\<^sub>R 2 \<in> interior ?U"
@@ -1890,7 +1797,7 @@
       unfolding unit_cube_def by force
   qed
   then have *: "interior ?U \<noteq> {}" by fast
-  have *: "?U homeomorphic s"
+  have *: "?U homeomorphic S"
     using homeomorphic_convex_compact[OF convex_unit_cube compact_unit_cube * assms(2,1,3)] .
   have "\<forall>f. continuous_on ?U f \<and> f ` ?U \<subseteq> ?U \<longrightarrow>
     (\<exists>x\<in>?U. f x = x)"
@@ -1898,7 +1805,7 @@
   then show ?thesis
     unfolding homeomorphic_fixpoint_property[OF *]
     using assms
-    by (auto simp: intro: that)
+    by (auto intro: that)
 qed
 
 
@@ -1920,49 +1827,37 @@
 
 lemma brouwer:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
-  assumes "compact s"
-    and "convex s"
-    and "s \<noteq> {}"
-    and "continuous_on s f"
-    and "f ` s \<subseteq> s"
-  obtains x where "x \<in> s" and "f x = x"
+  assumes S: "compact S" "convex S" "S \<noteq> {}"
+    and contf: "continuous_on S f"
+    and fim: "f ` S \<subseteq> S"
+  obtains x where "x \<in> S" and "f x = x"
 proof -
-  have "\<exists>e>0. s \<subseteq> cball 0 e"
-    using compact_imp_bounded[OF assms(1)]
-    unfolding bounded_pos
-    apply (erule_tac exE)
-    apply (rule_tac x=b in exI)
-    apply (auto simp add: dist_norm)
-    done
-  then obtain e where e: "e > 0" "s \<subseteq> cball 0 e"
+  have "\<exists>e>0. S \<subseteq> cball 0 e"
+    using compact_imp_bounded[OF \<open>compact S\<close>]  unfolding bounded_pos
+    by auto
+  then obtain e where e: "e > 0" "S \<subseteq> cball 0 e"
     by blast
-  have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point s) x = x"
-    apply (rule_tac brouwer_ball[OF e(1), of 0 "f \<circ> closest_point s"])
-    apply (rule continuous_on_compose )
-    apply (rule continuous_on_closest_point[OF assms(2) compact_imp_closed[OF assms(1)] assms(3)])
-    apply (rule continuous_on_subset[OF assms(4)])
-    apply (insert closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)])
-    using assms(5)[unfolded subset_eq]
-    using e(2)[unfolded subset_eq mem_cball]
-    apply (auto simp add: dist_norm)
-    done
-  then obtain x where x: "x \<in> cball 0 e" "(f \<circ> closest_point s) x = x" ..
-  have *: "closest_point s x = x"
-    apply (rule closest_point_self)
-    apply (rule assms(5)[unfolded subset_eq,THEN bspec[where x="x"], unfolded image_iff])
-    apply (rule_tac x="closest_point s x" in bexI)
-    using x
-    unfolding o_def
-    using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3), of x]
-    apply auto
-    done
+  have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point S) x = x"
+  proof (rule_tac brouwer_ball[OF e(1)])
+    show "continuous_on (cball 0 e) (f \<circ> closest_point S)"
+      apply (rule continuous_on_compose)
+      using S compact_eq_bounded_closed continuous_on_closest_point apply blast
+      by (meson S contf closest_point_in_set compact_imp_closed continuous_on_subset image_subsetI)
+    show "(f \<circ> closest_point S) ` cball 0 e \<subseteq> cball 0 e"
+      by clarsimp (metis S fim closest_point_exists(1) compact_eq_bounded_closed e(2) image_subset_iff mem_cball_0 subsetCE)
+  qed (use assms in auto)
+  then obtain x where x: "x \<in> cball 0 e" "(f \<circ> closest_point S) x = x" ..
+  have "x \<in> S"
+    by (metis closest_point_in_set comp_apply compact_imp_closed fim image_eqI S(1) S(3) subset_iff x(2))
+  then have *: "closest_point S x = x"
+    by (rule closest_point_self)
   show thesis
-    apply (rule_tac x="closest_point s x" in that)
-    unfolding x(2)[unfolded o_def]
-    apply (rule closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)])
-    using *
-    apply auto
-    done
+  proof
+    show "closest_point S x \<in> S"
+      by (simp add: "*" \<open>x \<in> S\<close>)
+    show "f (closest_point S x) = closest_point S x"
+      using "*" x(2) by auto
+  qed
 qed
 
 text \<open>So we get the no-retraction theorem.\<close>
@@ -1975,17 +1870,15 @@
   assume *: "frontier (cball a e) retract_of (cball a e)"
   have **: "\<And>xa. a - (2 *\<^sub>R a - xa) = - (a - xa)"
     using scaleR_left_distrib[of 1 1 a] by auto
-  obtain x where x:
-      "x \<in> {x. norm (a - x) = e}"
-      "2 *\<^sub>R a - x = x"
-    apply (rule retract_fixpoint_property[OF *, of "\<lambda>x. scaleR 2 a - x"])
-    apply (blast intro: brouwer_ball[OF assms])
-    apply (intro continuous_intros)
-    unfolding frontier_cball subset_eq Ball_def image_iff dist_norm sphere_def
-    apply (auto simp add: ** norm_minus_commute)
-    done
+  obtain x where x: "x \<in> {x. norm (a - x) = e}" "2 *\<^sub>R a - x = x"
+  proof (rule retract_fixpoint_property[OF *, of "\<lambda>x. scaleR 2 a - x"])
+    show "continuous_on (frontier (cball a e)) ((-) (2 *\<^sub>R a))"
+      by (intro continuous_intros)
+    show "(-) (2 *\<^sub>R a) ` frontier (cball a e) \<subseteq> frontier (cball a e)"
+      by clarsimp (metis "**" dist_norm norm_minus_cancel)
+  qed (auto simp: dist_norm intro: brouwer_ball[OF assms])
   then have "scaleR 2 a = scaleR 1 x + scaleR 1 x"
-    by (auto simp add: algebra_simps)
+    by (auto simp: algebra_simps)
   then have "a = x"
     unfolding scaleR_left_distrib[symmetric]
     by auto
@@ -2006,11 +1899,7 @@
   case False
   then show ?thesis
     unfolding contractible_def nullhomotopic_from_sphere_extension
-    apply (simp add: not_less)
-    apply (rule_tac x=id in exI)
-    apply (auto simp: continuous_on_def)
-    apply (meson dist_not_less_zero le_less less_le_trans)
-    done
+    using continuous_on_const less_eq_real_def by auto
 qed
 
 lemma connected_sphere_eq:
@@ -2035,9 +1924,8 @@
         by (metis dist_self greater insertI1 less_add_same_cancel1 mem_sphere mult_2 not_le zero_le_dist)
       then have "finite (sphere a r)"
         by auto
-      with L \<open>r > 0\<close> show "False"
-        apply (auto simp: connected_finite_iff_sing)
-        using xy by auto
+      with L \<open>r > 0\<close> xy show "False"
+        using connected_finite_iff_sing by auto
     qed
     with greater show ?rhs
       by (metis DIM_ge_Suc0 One_nat_def Suc_1 le_antisym not_less_eq_eq)
@@ -2098,12 +1986,10 @@
     unfolding retraction_def
   proof (intro conjI ballI)
     show "frontier (cball a B) \<subseteq> cball a B"
-      by (force simp:)
+      by force
     show "continuous_on (cball a B) h"
       unfolding h_def
-      apply (intro continuous_intros)
-      using contg continuous_on_subset notga apply auto
-      done
+      by (intro continuous_intros) (use contg continuous_on_subset notga in auto)
     show "h ` cball a B \<subseteq> frontier (cball a B)"
       using \<open>0 < B\<close> by (auto simp: h_def notga dist_norm)
     show "\<And>x. x \<in> frontier (cball a B) \<Longrightarrow> h x = x"
@@ -2117,76 +2003,76 @@
 subsection\<open>More Properties of Retractions\<close>
 
 lemma retraction:
-   "retraction s t r \<longleftrightarrow>
-    t \<subseteq> s \<and> continuous_on s r \<and> r ` s = t \<and> (\<forall>x \<in> t. r x = x)"
+   "retraction S T r \<longleftrightarrow>
+    T \<subseteq> S \<and> continuous_on S r \<and> r ` S = T \<and> (\<forall>x \<in> T. r x = x)"
 by (force simp: retraction_def)
 
 lemma retract_of_imp_extensible:
-  assumes "s retract_of t" and "continuous_on s f" and "f ` s \<subseteq> u"
-  obtains g where "continuous_on t g" "g ` t \<subseteq> u" "\<And>x. x \<in> s \<Longrightarrow> g x = f x"
+  assumes "S retract_of T" and "continuous_on S f" and "f ` S \<subseteq> U"
+  obtains g where "continuous_on T g" "g ` T \<subseteq> U" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
 using assms
 apply (clarsimp simp add: retract_of_def retraction)
-apply (rule_tac g = "f o r" in that)
+apply (rule_tac g = "f \<circ> r" in that)
 apply (auto simp: continuous_on_compose2)
 done
 
 lemma idempotent_imp_retraction:
-  assumes "continuous_on s f" and "f ` s \<subseteq> s" and "\<And>x. x \<in> s \<Longrightarrow> f(f x) = f x"
-    shows "retraction s (f ` s) f"
+  assumes "continuous_on S f" and "f ` S \<subseteq> S" and "\<And>x. x \<in> S \<Longrightarrow> f(f x) = f x"
+    shows "retraction S (f ` S) f"
 by (simp add: assms retraction)
 
 lemma retraction_subset:
-  assumes "retraction s t r" and "t \<subseteq> s'" and "s' \<subseteq> s"
-    shows "retraction s' t r"
-apply (simp add: retraction_def)
-by (metis assms continuous_on_subset image_mono retraction)
+  assumes "retraction S T r" and "T \<subseteq> s'" and "s' \<subseteq> S"
+  shows "retraction s' T r"
+  unfolding retraction_def
+  by (metis assms continuous_on_subset image_mono retraction)
 
 lemma retract_of_subset:
-  assumes "t retract_of s" and "t \<subseteq> s'" and "s' \<subseteq> s"
-    shows "t retract_of s'"
+  assumes "T retract_of S" and "T \<subseteq> s'" and "s' \<subseteq> S"
+    shows "T retract_of s'"
 by (meson assms retract_of_def retraction_subset)
 
-lemma retraction_refl [simp]: "retraction s s (\<lambda>x. x)"
+lemma retraction_refl [simp]: "retraction S S (\<lambda>x. x)"
 by (simp add: continuous_on_id retraction)
 
-lemma retract_of_refl [iff]: "s retract_of s"
+lemma retract_of_refl [iff]: "S retract_of S"
   using continuous_on_id retract_of_def retraction_def by fastforce
 
 lemma retract_of_imp_subset:
-   "s retract_of t \<Longrightarrow> s \<subseteq> t"
+   "S retract_of T \<Longrightarrow> S \<subseteq> T"
 by (simp add: retract_of_def retraction_def)
 
 lemma retract_of_empty [simp]:
-     "({} retract_of s) \<longleftrightarrow> s = {}"  "(s retract_of {}) \<longleftrightarrow> s = {}"
+     "({} retract_of S) \<longleftrightarrow> S = {}"  "(S retract_of {}) \<longleftrightarrow> S = {}"
 by (auto simp: retract_of_def retraction_def)
 
-lemma retract_of_singleton [iff]: "({x} retract_of s) \<longleftrightarrow> x \<in> s"
+lemma retract_of_singleton [iff]: "({x} retract_of S) \<longleftrightarrow> x \<in> S"
   using continuous_on_const
   by (auto simp: retract_of_def retraction_def)
 
 lemma retraction_comp:
-   "\<lbrakk>retraction s t f; retraction t u g\<rbrakk>
-        \<Longrightarrow> retraction s u (g o f)"
+   "\<lbrakk>retraction S T f; retraction T U g\<rbrakk>
+        \<Longrightarrow> retraction S U (g \<circ> f)"
 apply (auto simp: retraction_def intro: continuous_on_compose2)
 by blast
 
 lemma retract_of_trans [trans]:
-  assumes "s retract_of t" and "t retract_of u"
-    shows "s retract_of u"
+  assumes "S retract_of T" and "T retract_of U"
+    shows "S retract_of U"
 using assms by (auto simp: retract_of_def intro: retraction_comp)
 
 lemma closedin_retract:
-  fixes s :: "'a :: real_normed_vector set"
-  assumes "s retract_of t"
-    shows "closedin (subtopology euclidean t) s"
+  fixes S :: "'a :: real_normed_vector set"
+  assumes "S retract_of T"
+    shows "closedin (subtopology euclidean T) S"
 proof -
-  obtain r where "s \<subseteq> t" "continuous_on t r" "r ` t \<subseteq> s" "\<And>x. x \<in> s \<Longrightarrow> r x = x"
+  obtain r where "S \<subseteq> T" "continuous_on T r" "r ` T \<subseteq> S" "\<And>x. x \<in> S \<Longrightarrow> r x = x"
     using assms by (auto simp: retract_of_def retraction_def)
-  then have s: "s = {x \<in> t. (norm(r x - x)) = 0}" by auto
+  then have S: "S = {x \<in> T. (norm(r x - x)) = 0}" by auto
   show ?thesis
-    apply (subst s)
+    apply (subst S)
     apply (rule continuous_closedin_preimage_constant)
-    by (simp add: \<open>continuous_on t r\<close> continuous_on_diff continuous_on_id continuous_on_norm)
+    by (simp add: \<open>continuous_on T r\<close> continuous_on_diff continuous_on_id continuous_on_norm)
 qed
 
 lemma closedin_self [simp]:
@@ -2195,52 +2081,52 @@
   by (simp add: closedin_retract)
 
 lemma retract_of_contractible:
-  assumes "contractible t" "s retract_of t"
-    shows "contractible s"
+  assumes "contractible T" "S retract_of T"
+    shows "contractible S"
 using assms
 apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with)
 apply (rule_tac x="r a" in exI)
-apply (rule_tac x="r o h" in exI)
+apply (rule_tac x="r \<circ> h" in exI)
 apply (intro conjI continuous_intros continuous_on_compose)
 apply (erule continuous_on_subset | force)+
 done
 
 lemma retract_of_compact:
-     "\<lbrakk>compact t; s retract_of t\<rbrakk> \<Longrightarrow> compact s"
+     "\<lbrakk>compact T; S retract_of T\<rbrakk> \<Longrightarrow> compact S"
   by (metis compact_continuous_image retract_of_def retraction)
 
 lemma retract_of_closed:
-    fixes s :: "'a :: real_normed_vector set"
-    shows "\<lbrakk>closed t; s retract_of t\<rbrakk> \<Longrightarrow> closed s"
+    fixes S :: "'a :: real_normed_vector set"
+    shows "\<lbrakk>closed T; S retract_of T\<rbrakk> \<Longrightarrow> closed S"
   by (metis closedin_retract closedin_closed_eq)
 
 lemma retract_of_connected:
-    "\<lbrakk>connected t; s retract_of t\<rbrakk> \<Longrightarrow> connected s"
+    "\<lbrakk>connected T; S retract_of T\<rbrakk> \<Longrightarrow> connected S"
   by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction)
 
 lemma retract_of_path_connected:
-    "\<lbrakk>path_connected t; s retract_of t\<rbrakk> \<Longrightarrow> path_connected s"
+    "\<lbrakk>path_connected T; S retract_of T\<rbrakk> \<Longrightarrow> path_connected S"
   by (metis path_connected_continuous_image retract_of_def retraction)
 
 lemma retract_of_simply_connected:
-    "\<lbrakk>simply_connected t; s retract_of t\<rbrakk> \<Longrightarrow> simply_connected s"
+    "\<lbrakk>simply_connected T; S retract_of T\<rbrakk> \<Longrightarrow> simply_connected S"
 apply (simp add: retract_of_def retraction_def, clarify)
 apply (rule simply_connected_retraction_gen)
 apply (force simp: continuous_on_id elim!: continuous_on_subset)+
 done
 
 lemma retract_of_homotopically_trivial:
-  assumes ts: "t retract_of s"
-      and hom: "\<And>f g. \<lbrakk>continuous_on u f; f ` u \<subseteq> s;
-                       continuous_on u g; g ` u \<subseteq> s\<rbrakk>
-                       \<Longrightarrow> homotopic_with (\<lambda>x. True) u s f g"
-      and "continuous_on u f" "f ` u \<subseteq> t"
-      and "continuous_on u g" "g ` u \<subseteq> t"
-    shows "homotopic_with (\<lambda>x. True) u t f g"
+  assumes ts: "T retract_of S"
+      and hom: "\<And>f g. \<lbrakk>continuous_on U f; f ` U \<subseteq> S;
+                       continuous_on U g; g ` U \<subseteq> S\<rbrakk>
+                       \<Longrightarrow> homotopic_with (\<lambda>x. True) U S f g"
+      and "continuous_on U f" "f ` U \<subseteq> T"
+      and "continuous_on U g" "g ` U \<subseteq> T"
+    shows "homotopic_with (\<lambda>x. True) U T f g"
 proof -
-  obtain r where "r ` s \<subseteq> s" "continuous_on s r" "\<forall>x\<in>s. r (r x) = r x" "t = r ` s"
+  obtain r where "r ` S \<subseteq> S" "continuous_on S r" "\<forall>x\<in>S. r (r x) = r x" "T = r ` S"
     using ts by (auto simp: retract_of_def retraction)
-  then obtain k where "Retracts s r t k"
+  then obtain k where "Retracts S r T k"
     unfolding Retracts_def
     by (metis continuous_on_subset dual_order.trans image_iff image_mono)
   then show ?thesis
@@ -2251,15 +2137,15 @@
 qed
 
 lemma retract_of_homotopically_trivial_null:
-  assumes ts: "t retract_of s"
-      and hom: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s\<rbrakk>
-                     \<Longrightarrow> \<exists>c. homotopic_with (\<lambda>x. True) u s f (\<lambda>x. c)"
-      and "continuous_on u f" "f ` u \<subseteq> t"
-  obtains c where "homotopic_with (\<lambda>x. True) u t f (\<lambda>x. c)"
+  assumes ts: "T retract_of S"
+      and hom: "\<And>f. \<lbrakk>continuous_on U f; f ` U \<subseteq> S\<rbrakk>
+                     \<Longrightarrow> \<exists>c. homotopic_with (\<lambda>x. True) U S f (\<lambda>x. c)"
+      and "continuous_on U f" "f ` U \<subseteq> T"
+  obtains c where "homotopic_with (\<lambda>x. True) U T f (\<lambda>x. c)"
 proof -
-  obtain r where "r ` s \<subseteq> s" "continuous_on s r" "\<forall>x\<in>s. r (r x) = r x" "t = r ` s"
+  obtain r where "r ` S \<subseteq> S" "continuous_on S r" "\<forall>x\<in>S. r (r x) = r x" "T = r ` S"
     using ts by (auto simp: retract_of_def retraction)
-  then obtain k where "Retracts s r t k"
+  then obtain k where "Retracts S r T k"
     unfolding Retracts_def
     by (metis continuous_on_subset dual_order.trans image_iff image_mono)
   then show ?thesis
@@ -2269,35 +2155,34 @@
 qed
 
 lemma retraction_imp_quotient_map:
-   "retraction s t r
-    \<Longrightarrow> u \<subseteq> t
-            \<Longrightarrow> (openin (subtopology euclidean s) (s \<inter> r -` u) \<longleftrightarrow>
-                 openin (subtopology euclidean t) u)"
+   "retraction S T r
+    \<Longrightarrow> U \<subseteq> T
+            \<Longrightarrow> (openin (subtopology euclidean S) (S \<inter> r -` U) \<longleftrightarrow>
+                 openin (subtopology euclidean T) U)"
 apply (clarsimp simp add: retraction)
 apply (rule continuous_right_inverse_imp_quotient_map [where g=r])
 apply (auto simp: elim: continuous_on_subset)
 done
 
 lemma retract_of_locally_compact:
-    fixes s :: "'a :: {heine_borel,real_normed_vector} set"
-    shows  "\<lbrakk> locally compact s; t retract_of s\<rbrakk> \<Longrightarrow> locally compact t"
+    fixes S :: "'a :: {heine_borel,real_normed_vector} set"
+    shows  "\<lbrakk> locally compact S; T retract_of S\<rbrakk> \<Longrightarrow> locally compact T"
   by (metis locally_compact_closedin closedin_retract)
 
 lemma retract_of_Times:
-   "\<lbrakk>s retract_of s'; t retract_of t'\<rbrakk> \<Longrightarrow> (s \<times> t) retract_of (s' \<times> t')"
+   "\<lbrakk>S retract_of s'; T retract_of t'\<rbrakk> \<Longrightarrow> (S \<times> T) retract_of (s' \<times> t')"
 apply (simp add: retract_of_def retraction_def Sigma_mono, clarify)
 apply (rename_tac f g)
-apply (rule_tac x="\<lambda>z. ((f o fst) z, (g o snd) z)" in exI)
+apply (rule_tac x="\<lambda>z. ((f \<circ> fst) z, (g \<circ> snd) z)" in exI)
 apply (rule conjI continuous_intros | erule continuous_on_subset | force)+
 done
 
 lemma homotopic_into_retract:
-   "\<lbrakk>f ` s \<subseteq> t; g ` s \<subseteq> t; t retract_of u;
-        homotopic_with (\<lambda>x. True) s u f g\<rbrakk>
-        \<Longrightarrow> homotopic_with (\<lambda>x. True) s t f g"
+   "\<lbrakk>f ` S \<subseteq> T; g ` S \<subseteq> T; T retract_of U; homotopic_with (\<lambda>x. True) S U f g\<rbrakk>
+        \<Longrightarrow> homotopic_with (\<lambda>x. True) S T f g"
 apply (subst (asm) homotopic_with_def)
 apply (simp add: homotopic_with retract_of_def retraction_def, clarify)
-apply (rule_tac x="r o h" in exI)
+apply (rule_tac x="r \<circ> h" in exI)
 apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+
 done
 
@@ -2317,15 +2202,19 @@
 
 lemma deformation_retract_imp_homotopy_eqv:
   fixes S :: "'a::euclidean_space set"
-  assumes "homotopic_with (\<lambda>x. True) S S id r" "retraction S T r"
-    shows "S homotopy_eqv T"
-  apply (simp add: homotopy_eqv_def)
-  apply (rule_tac x=r in exI)
-  using assms apply (simp add: retraction_def)
-  apply (rule_tac x=id in exI)
-  apply (auto simp: continuous_on_id)
-   apply (metis homotopic_with_symD)
-  by (metis continuous_on_id' homotopic_with_equal homotopic_with_symD id_apply image_id subset_refl)
+  assumes "homotopic_with (\<lambda>x. True) S S id r" and r: "retraction S T r"
+  shows "S homotopy_eqv T"
+proof -
+  have "homotopic_with (\<lambda>x. True) S S (id \<circ> r) id"
+    by (simp add: assms(1) homotopic_with_symD)
+  moreover have "homotopic_with (\<lambda>x. True) T T (r \<circ> id) id"
+    using r unfolding retraction_def
+    by (metis (no_types, lifting) comp_id continuous_on_id' homotopic_with_equal homotopic_with_symD id_def image_id order_refl)
+  ultimately
+  show ?thesis
+    unfolding homotopy_eqv_def
+    by (metis continuous_on_id' id_def image_id r retraction_def)
+qed
 
 lemma deformation_retract:
   fixes S :: "'a::euclidean_space set"
@@ -2356,10 +2245,8 @@
   have "{a} retract_of S"
     by (simp add: \<open>a \<in> S\<close>)
   moreover have "homotopic_with (\<lambda>x. True) S S id (\<lambda>x. a)"
-    using assms
-    apply (clarsimp simp add: contractible_def)
-    apply (rule homotopic_with_trans, assumption)
-    by (metis assms(1) contractible_imp_path_connected homotopic_constant_maps homotopic_with_sym homotopic_with_trans insert_absorb insert_not_empty path_component_mem(1) path_connected_component)
+      using assms
+      by (auto simp: contractible_def continuous_on_const continuous_on_id homotopic_into_contractible image_subset_iff)
   moreover have "(\<lambda>x. a) ` S \<subseteq> {a}"
     by (simp add: image_subsetI)
   ultimately show ?thesis
@@ -2382,15 +2269,12 @@
     using assms by auto (metis imageI subset_iff)
   have contp': "continuous_on S p"
     by (rule continuous_on_subset [OF contp \<open>S \<subseteq> T\<close>])
-  have "continuous_on T (q \<circ> p)"
-    apply (rule continuous_on_compose [OF contp])
-    apply (simp add: *)
-    apply (rule continuous_on_inv [OF contp' \<open>compact S\<close>])
-    using assms by auto
+  have "continuous_on (p ` T) q"
+    by (simp add: "*" assms(1) assms(2) assms(5) continuous_on_inv contp' rev_subsetD)
+  then have "continuous_on T (q \<circ> p)"
+    by (rule continuous_on_compose [OF contp])
   then show ?thesis
-    apply (rule continuous_on_eq [of _ "q o p"])
-    apply (simp add: o_def)
-    done
+    by (rule continuous_on_eq [of _ "q \<circ> p"]) (simp add: o_def)
 qed
 
 lemma continuous_on_compact_surface_projection:
@@ -2441,21 +2325,19 @@
   have aaffS: "a \<in> affine hull S"
     by (meson arelS subsetD hull_inc rel_interior_subset)
   have "((\<lambda>z. z - a) ` (affine hull S - {a})) = ((\<lambda>z. z - a) ` (affine hull S)) - {0}"
-    by (auto simp: )
+    by auto
   moreover have "continuous_on (((\<lambda>z. z - a) ` (affine hull S)) - {0}) (\<lambda>x. dd x *\<^sub>R x)"
   proof (rule continuous_on_compact_surface_projection)
     show "compact (rel_frontier ((\<lambda>z. z - a) ` S))"
       by (simp add: \<open>bounded S\<close> bounded_translation_minus compact_rel_frontier_bounded)
     have releq: "rel_frontier ((\<lambda>z. z - a) ` S) = (\<lambda>z. z - a) ` rel_frontier S"
       using rel_frontier_translation [of "-a"] add.commute by simp
-    also have "... \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}"
+    also have "\<dots> \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}"
       using rel_frontier_affine_hull arelS rel_frontier_def by fastforce
     finally show "rel_frontier ((\<lambda>z. z - a) ` S) \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}" .
     show "cone ((\<lambda>z. z - a) ` (affine hull S))"
-      apply (rule subspace_imp_cone)
-      using aaffS
-      apply (simp add: subspace_affine image_comp o_def affine_translation_aux [of a])
-      done
+      by (rule subspace_imp_cone)
+         (use aaffS in \<open>simp add: subspace_affine image_comp o_def affine_translation_aux [of a]\<close>)
     show "(0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)) \<longleftrightarrow> (dd x = k)"
          if x: "x \<in> (\<lambda>z. z - a) ` (affine hull S) - {0}" for k x
     proof
@@ -2471,7 +2353,7 @@
         then have segsub: "open_segment a (a + k *\<^sub>R x) \<subseteq> rel_interior S"
           by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS])
         have "x \<noteq> 0" and xaffS: "a + x \<in> affine hull S"
-          using x by (auto simp: )
+          using x by auto
         then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \<in> rel_frontier S"
           using dd1 by auto
         moreover have "a + dd x *\<^sub>R x \<in> open_segment a (a + k *\<^sub>R x)"
@@ -2483,7 +2365,7 @@
           apply (metis (no_types) \<open>k \<noteq> 0\<close> divide_inverse_commute inverse_eq_divide mult.left_commute right_inverse)
           done
         ultimately show ?thesis
-          using segsub by (auto simp add: rel_frontier_def)
+          using segsub by (auto simp: rel_frontier_def)
       qed
       moreover have False if "k < dd x"
         using x k that rel_frontier_def
@@ -2497,7 +2379,7 @@
   have "continuous_on (affine hull S - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))"
     by (intro * continuous_intros continuous_on_compose)
   with affS have contdd: "continuous_on (T - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))"
-    by (blast intro: continuous_on_subset elim: )
+    by (blast intro: continuous_on_subset)
   show ?thesis
   proof
     show "homotopic_with (\<lambda>x. True) (T - {a}) (T - {a}) id (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
@@ -2510,11 +2392,10 @@
            if "x \<in> T - {a}" for x
       proof (clarsimp simp: in_segment, intro conjI)
         fix u::real assume u: "0 \<le> u" "u \<le> 1"
-        show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<in> T"
-          apply (rule convexD [OF \<open>convex T\<close>])
-          using that u apply (auto simp add: )
-          apply (metis add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 relS subsetD)
-          done
+        have "a + dd (x - a) *\<^sub>R (x - a) \<in> T"
+          by (metis DiffD1 DiffD2 add.commute add.right_neutral affS dd1 diff_add_cancel relS singletonI subsetCE that)
+        then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<in> T"
+          using convexD [OF \<open>convex T\<close>] that u by simp
         have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x - a)) = a \<longleftrightarrow>
                   (1 - u + u * d) *\<^sub>R (x - a) = 0" for d
           by (auto simp: algebra_simps)
@@ -2541,7 +2422,7 @@
       show "a + dd (x - a) *\<^sub>R (x - a) = x" if x: "x \<in> rel_frontier S" for x
       proof -
         have "x \<noteq> a"
-          using that arelS by (auto simp add: rel_frontier_def)
+          using that arelS by (auto simp: rel_frontier_def)
         have False if "dd (x - a) < 1"
         proof -
           have "x \<in> closure S"
@@ -2551,7 +2432,7 @@
           have  xaffS: "x \<in> affine hull S"
             using affS relS x by auto
           then have "0 < dd (x - a)" and inS: "a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S"
-            using dd1 by (auto simp add: \<open>x \<noteq> a\<close>)
+            using dd1 by (auto simp: \<open>x \<noteq> a\<close>)
           moreover have "a + dd (x - a) *\<^sub>R (x - a) \<in> open_segment a x"
             using  \<open>x \<noteq> a\<close> \<open>0 < dd (x - a)\<close>
             apply (simp add: in_segment)
@@ -2559,7 +2440,7 @@
             apply (simp add: algebra_simps that)
             done
           ultimately show ?thesis
-            using segsub by (auto simp add: rel_frontier_def)
+            using segsub by (auto simp: rel_frontier_def)
         qed
         moreover have False if "1 < dd (x - a)"
           using x that dd2 [of "x - a" 1] \<open>x \<noteq> a\<close> closure_affine_hull
@@ -2578,7 +2459,7 @@
   assumes "bounded S" "convex S" "a \<in> rel_interior S"
     shows "rel_frontier S retract_of (affine hull S - {a})"
 apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S "affine hull S" a])
-apply (auto simp add: affine_imp_convex rel_frontier_affine_hull retract_of_def assms)
+apply (auto simp: affine_imp_convex rel_frontier_affine_hull retract_of_def assms)
 done
 
 corollary rel_boundary_retract_of_punctured_affine_hull:
@@ -2643,7 +2524,7 @@
     using assms by (auto simp: path_component_def)
   then show ?thesis
     apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def)
-    apply (rule_tac x = "\<lambda>z. inverse(norm(snd z - (g o fst)z)) *\<^sub>R (snd z - (g o fst)z)" in exI)
+    apply (rule_tac x = "\<lambda>z. inverse(norm(snd z - (g \<circ> fst)z)) *\<^sub>R (snd z - (g \<circ> fst)z)" in exI)
     apply (intro conjI continuous_intros)
     apply (rule continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+
     done
@@ -2767,7 +2648,7 @@
     using hom by (force simp: homeomorphic_def)
   then have "continuous_on (f ` T) g"
     by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def)
-  then have contgf: "continuous_on T (g o f)"
+  then have contgf: "continuous_on T (g \<circ> f)"
     by (metis continuous_on_compose contf)
   have gfTC: "(g \<circ> f) ` T \<subseteq> C"
   proof -
@@ -2779,7 +2660,7 @@
                       "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x"
     by (metis Dugundji [OF C cloUT contgf gfTC])
   show ?thesis
-  proof (rule_tac g = "h o r o f'" in that)
+  proof (rule_tac g = "h \<circ> r \<circ> f'" in that)
     show "continuous_on U (h \<circ> r \<circ> f')"
       apply (intro continuous_on_compose f')
        using continuous_on_subset contr f' apply blast
@@ -2811,7 +2692,7 @@
   have [simp]: "S' \<subseteq> U" using clo closedin_limpt by blast
   show ?thesis
   proof (simp add: retraction_def retract_of_def, intro exI conjI)
-    show "continuous_on U (g o h')"
+    show "continuous_on U (g \<circ> h')"
       apply (intro continuous_on_compose h')
       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       done
@@ -2853,7 +2734,7 @@
     using clo closedin_imp_subset by auto
   show "T retract_of U"
   proof (simp add: retraction_def retract_of_def, intro exI conjI)
-    show "continuous_on U (g o h')"
+    show "continuous_on U (g \<circ> h')"
       apply (intro continuous_on_compose h')
       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       done
@@ -2919,7 +2800,7 @@
     using hom by (force simp: homeomorphic_def)
   have "continuous_on (f ` T) g"
     by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def homgh)
-  then have contgf: "continuous_on T (g o f)"
+  then have contgf: "continuous_on T (g \<circ> f)"
     by (intro continuous_on_compose contf)
   have gfTC: "(g \<circ> f) ` T \<subseteq> C"
   proof -
@@ -2933,7 +2814,7 @@
               and eq: "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x"
     by (metis Dugundji [OF C cloUT contgf gfTC])
   show ?thesis
-  proof (rule_tac V = "U \<inter> f' -` D" and g = "h o r o f'" in that)
+  proof (rule_tac V = "U \<inter> f' -` D" and g = "h \<circ> r \<circ> f'" in that)
     show "T \<subseteq> U \<inter> f' -` D"
       using cloUT closedin_imp_subset \<open>S' \<subseteq> D\<close> \<open>f ` T \<subseteq> S\<close> eq homeomorphism_image1 homgh
       by fastforce
@@ -2976,7 +2857,7 @@
     by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo])
   have "S' retract_of V"
   proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>S' \<subseteq> V\<close>)
-    show "continuous_on V (g o h')"
+    show "continuous_on V (g \<circ> h')"
       apply (intro continuous_on_compose h')
       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       done
@@ -3029,7 +2910,7 @@
     using clo closedin_imp_subset by auto
   have "T retract_of V"
   proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>T \<subseteq> V\<close>)
-    show "continuous_on V (g o h')"
+    show "continuous_on V (g \<circ> h')"
       apply (intro continuous_on_compose h')
       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       done
@@ -3086,7 +2967,7 @@
     using Diff_subset_conv \<open>U - Z \<subseteq> W\<close> by blast
   ultimately show ?thesis
     apply (rule_tac V=V and W = "U-W" in that)
-    using openin_imp_subset apply (force simp:)+
+    using openin_imp_subset apply force+
     done
 qed
 
@@ -3146,7 +3027,7 @@
   proof (simp add: retraction_def retract_of_def, intro exI conjI)
     show "S' \<subseteq> W" "S' \<subseteq> h -` X"
       using him WS' closedin_imp_subset by blast+
-    show "continuous_on (W \<inter> h -` X) (f o r o h)"
+    show "continuous_on (W \<inter> h -` X) (f \<circ> r \<circ> h)"
     proof (intro continuous_on_compose)
       show "continuous_on (W \<inter> h -` X) h"
         by (meson conth continuous_on_subset inf_le1)
@@ -3356,7 +3237,7 @@
 apply (clarsimp elim!: all_forward)
 apply (erule impCE, metis subset_trans)
 apply (clarsimp elim!: ex_forward)
-apply (rule_tac x="r o g" in exI)
+apply (rule_tac x="r \<circ> g" in exI)
 by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans)
 
 lemma AR_retract_of_AR:
@@ -3642,7 +3523,7 @@
   obtain r0
     where "S \<inter> T \<subseteq> W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \<subseteq> S \<inter> T"
       and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x"
-    using ret  by (force simp add: retract_of_def retraction_def)
+    using ret  by (force simp: retract_of_def retraction_def)
   have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x
     using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin)
   define r where "r \<equiv> \<lambda>x. if x \<in> W0 then r0 x else x"
@@ -3667,8 +3548,7 @@
                 and opeSW1: "openin (subtopology euclidean S') W1"
                 and "g ` W1 \<subseteq> S" and geqr: "\<And>x. x \<in> W0 \<union> S \<Longrightarrow> g x = r x"
     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> _ \<open>r ` (W0 \<union> S) \<subseteq> S\<close> cloS'WS])
-     apply (rule continuous_on_subset [OF contr])
-    apply (blast intro:  elim: )+
+     apply (rule continuous_on_subset [OF contr], blast+)
     done
   have cloT'WT: "closedin (subtopology euclidean T') (W0 \<union> T)"
     by (meson closedin_subset_trans UT cloUT' \<open>T \<subseteq> T'\<close> \<open>W \<subseteq> T'\<close> cloUW cloWW0 
@@ -3677,13 +3557,12 @@
                 and opeSW2: "openin (subtopology euclidean T') W2"
                 and "h ` W2 \<subseteq> T" and heqr: "\<And>x. x \<in> W0 \<union> T \<Longrightarrow> h x = r x"
     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> _ \<open>r ` (W0 \<union> T) \<subseteq> T\<close> cloT'WT])
-     apply (rule continuous_on_subset [OF contr])
-    apply (blast intro:  elim: )+
+     apply (rule continuous_on_subset [OF contr], blast+)
     done
   have "S' \<inter> T' = W"
     by (force simp: S'_def T'_def W_def)
   obtain O1 O2 where "open O1" "W1 = S' \<inter> O1" "open O2" "W2 = T' \<inter> O2"
-    using opeSW1 opeSW2 by (force simp add: openin_open)
+    using opeSW1 opeSW2 by (force simp: openin_open)
   show ?thesis
   proof
     have eq: "W1 - (W - U0) \<union> (W2 - (W - U0)) =
@@ -3692,25 +3571,23 @@
       by (auto simp: \<open>S' \<union> T' = U\<close> [symmetric] \<open>S' \<inter> T' = W\<close> [symmetric] \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close>)
     show "openin (subtopology euclidean U) (W1 - (W - U0) \<union> (W2 - (W - U0)))"
       apply (subst eq)
-      apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \<open>open O1\<close> \<open>open O2\<close>)
-      apply simp_all
+      apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \<open>open O1\<close> \<open>open O2\<close>, simp_all)
       done
     have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W1 - (W - U0))"
       using cloUS' apply (simp add: closedin_closed)
       apply (erule ex_forward)
       using U0 \<open>W0 \<union> S \<subseteq> W1\<close>
-      apply (auto simp add: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric])
+      apply (auto simp: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric])
       done
     have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W2 - (W - U0))"
       using cloUT' apply (simp add: closedin_closed)
       apply (erule ex_forward)
       using U0 \<open>W0 \<union> T \<subseteq> W2\<close>
-      apply (auto simp add: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric])
+      apply (auto simp: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric])
       done
     have *: "\<forall>x\<in>S \<union> T. (if x \<in> S' then g x else h x) = x"
       using ST \<open>S' \<inter> T' = W\<close> cloT'WT closedin_subset geqr heqr 
-      apply (auto simp: r_def)
-       apply fastforce
+      apply (auto simp: r_def, fastforce)
       using \<open>S \<subseteq> S'\<close> \<open>T \<subseteq> T'\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W1 = S' \<inter> O1\<close>  by auto
     have "\<exists>r. continuous_on (W1 - (W - U0) \<union> (W2 - (W - U0))) r \<and>
               r ` (W1 - (W - U0) \<union> (W2 - (W - U0))) \<subseteq> S \<union> T \<and> 
@@ -3725,7 +3602,7 @@
       done
     then show "S \<union> T retract_of W1 - (W - U0) \<union> (W2 - (W - U0))"
       using  \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> ST opeUU0 U0
-      by (auto simp add: retract_of_def retraction_def)
+      by (auto simp: retract_of_def retraction_def)
   qed
 qed
 
@@ -4059,15 +3936,15 @@
     by (auto simp: closest_point_self)
   have "rel_frontier S retract_of affine hull S - {a}"
     by (simp add: assms a rel_frontier_retract_of_punctured_affine_hull)
-  also have "... retract_of {x. closest_point (affine hull S) x \<noteq> a}"
+  also have "\<dots> retract_of {x. closest_point (affine hull S) x \<noteq> a}"
     apply (simp add: retract_of_def retraction_def ahS)
     apply (rule_tac x="closest_point (affine hull S)" in exI)
-    apply (auto simp add: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point)
+    apply (auto simp: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point)
     done
   finally have "rel_frontier S retract_of {x. closest_point (affine hull S) x \<noteq> a}" .
   moreover have "openin (subtopology euclidean UNIV) (UNIV \<inter> closest_point (affine hull S) -` (- {a}))"
     apply (rule continuous_openin_preimage_gen)
-    apply (auto simp add: False affine_imp_convex continuous_on_closest_point)
+    apply (auto simp: False affine_imp_convex continuous_on_closest_point)
     done
   ultimately show ?thesis
     unfolding ENR_def
@@ -4116,7 +3993,7 @@
       apply (rule continuous_on_cases_local [OF clS clT])
       using r by (auto simp: continuous_on_id)
   qed (use r in auto)
-  also have "... retract_of U"
+  also have "\<dots> retract_of U"
     by (rule Un)
   finally show ?thesis .
 qed
@@ -4499,7 +4376,7 @@
              and him: "h ` ({0..1} \<times> S) \<subseteq> U"
              and [simp]: "\<And>x. h(0, x) = f x" "\<And>x. h(1::real, x) = g x"
        using assms by (auto simp: homotopic_with_def)
-  define h' where "h' \<equiv>  \<lambda>z. if snd z \<in> S then h z else (f o snd) z"
+  define h' where "h' \<equiv>  \<lambda>z. if snd z \<in> S then h z else (f \<circ> snd) z"
   define B where "B \<equiv> {0::real} \<times> T \<union> {0..1} \<times> S"
   have clo0T: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0::real} \<times> T)"
     by (simp add: closedin_subtopology_refl closedin_Times)
@@ -4542,7 +4419,7 @@
                       "retraction V B r" for V r
       using that
       apply (clarsimp simp add: retraction_def)
-      apply (rule Vk [of V "h' o r"], assumption+)
+      apply (rule Vk [of V "h' \<circ> r"], assumption+)
         apply (metis continuous_on_compose conth' continuous_on_subset) 
       using \<open>h' ` B \<subseteq> U\<close> apply force+
       done
@@ -4629,7 +4506,7 @@
 proof
   assume ?lhs
   then obtain c where c: "homotopic_with (\<lambda>x. True) S T (\<lambda>x. c) f"
-    by (blast intro: homotopic_with_symD elim: )
+    by (blast intro: homotopic_with_symD)
   have "closedin (subtopology euclidean UNIV) S"
     using \<open>closed S\<close> closed_closedin by fastforce
   then obtain g where "continuous_on UNIV g" "range g \<subseteq> T"
@@ -4645,10 +4522,10 @@
   then obtain c where "homotopic_with (\<lambda>h. True) UNIV T g (\<lambda>x. c)"
     using nullhomotopic_from_contractible [of UNIV g T] contractible_UNIV by blast
   then show ?lhs
-    apply (rule_tac x="c" in exI)
+    apply (rule_tac x=c in exI)
     apply (rule homotopic_with_eq [of _ _ _ g "\<lambda>x. c"])
     apply (rule homotopic_with_subset_left)
-    apply (auto simp add: \<open>\<And>x. x \<in> S \<Longrightarrow> g x = f x\<close>)
+    apply (auto simp: \<open>\<And>x. x \<in> S \<Longrightarrow> g x = f x\<close>)
     done
 qed
 
@@ -4672,7 +4549,7 @@
            (is "?lhs = ?rhs")
 proof (cases "r = 0")
   case True with fim show ?thesis
-    apply (auto simp: )
+    apply auto
     using fim continuous_on_const apply fastforce
     by (metis contf contractible_sing nullhomotopic_into_contractible)
 next
@@ -4717,11 +4594,11 @@
     obtain g where "range g \<subseteq> sphere 0 1" "continuous_on UNIV g"
                         "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)"
       using notr
-      by (auto simp add: nullhomotopic_into_sphere_extension
+      by (auto simp: nullhomotopic_into_sphere_extension
                  [OF \<open>closed S\<close> continuous_on_Borsuk_map [OF \<open>a \<notin> S\<close>] False s01])
     with \<open>a \<notin> S\<close> show  "~ ?lhs"
       apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog)
-      apply (drule_tac x="g" in spec)
+      apply (drule_tac x=g in spec)
       using continuous_on_subset by fastforce 
   next
     assume "~ ?lhs"
@@ -5070,7 +4947,7 @@
         then obtain k where "y \<in> V k" and j: "\<forall>j<k. y \<notin> V j"
           by (metis image_iff V wop)
         with him t show "\<zeta>(t, y) \<in> T"
-          by (subst eq) (force simp:)+
+          by (subst eq) force+
       qed
       fix X y
       assume "X \<in> \<V>" "y \<in> X"
@@ -5291,4 +5168,16 @@
   using connected_complement_homeomorphic_convex_compact [OF assms]
   using \<open>compact T\<close> compact_eq_bounded_closed connected_open_path_connected hom homeomorphic_compactness by blast
 
+lemma path_connected_complement_homeomorphic_interval:
+  fixes S :: "'a::euclidean_space set"
+  assumes "S homeomorphic cbox a b" "2 \<le> DIM('a)"
+  shows "path_connected(-S)"
+  using assms compact_cbox convex_box(1) path_connected_complement_homeomorphic_convex_compact by blast
+
+lemma connected_complement_homeomorphic_interval:
+  fixes S :: "'a::euclidean_space set"
+  assumes "S homeomorphic cbox a b" "2 \<le> DIM('a)"
+  shows "connected(-S)"
+  using assms path_connected_complement_homeomorphic_interval path_connected_imp_connected by blast
+
 end
--- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy	Thu May 03 15:07:14 2018 +0200
@@ -1,3 +1,7 @@
+(* Title:      HOL/Analysis/Cartesian_Euclidean_Space.thy
+   Some material by Jose Divasón, Tim Makarios and L C Paulson
+*)
+
 section \<open>Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space\<close>
 
 theory Cartesian_Euclidean_Space
@@ -176,12 +180,10 @@
 qed
 
 lemma matrix_mult_transpose_dot_column:
-  fixes A :: "real^'n^'n"
   shows "transpose A ** A = (\<chi> i j. inner (column i A) (column j A))"
   by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def column_def inner_vec_def)
 
 lemma matrix_mult_transpose_dot_row:
-  fixes A :: "real^'n^'n"
   shows "A ** transpose A = (\<chi> i j. inner (row i A) (row j A))"
   by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def row_def inner_vec_def)
 
@@ -215,11 +217,77 @@
   by (simp add: linear_conv_bounded_linear linear_matrix_vector_mul_eq)
 
 lemma
-  fixes A :: "real^'n^'m"
+  fixes A :: "'a::{euclidean_space,real_algebra_1}^'n^'m"
   shows matrix_vector_mult_linear_continuous_at [continuous_intros]: "isCont (( *v) A) z"
     and matrix_vector_mult_linear_continuous_on [continuous_intros]: "continuous_on S (( *v) A)"
   by (simp_all add: linear_continuous_at linear_continuous_on)
 
+lemma scalar_invertible:
+  fixes A :: "('a::real_algebra_1)^'m^'n"
+  assumes "k \<noteq> 0" and "invertible A"
+  shows "invertible (k *\<^sub>R A)"
+proof -
+  obtain A' where "A ** A' = mat 1" and "A' ** A = mat 1"
+    using assms unfolding invertible_def by auto
+  with `k \<noteq> 0`
+  have "(k *\<^sub>R A) ** ((1/k) *\<^sub>R A') = mat 1" "((1/k) *\<^sub>R A') ** (k *\<^sub>R A) = mat 1"
+    by (simp_all add: assms matrix_scalar_ac)
+  thus "invertible (k *\<^sub>R A)"
+    unfolding invertible_def by auto
+qed
+
+lemma scalar_invertible_iff:
+  fixes A :: "('a::real_algebra_1)^'m^'n"
+  assumes "k \<noteq> 0" and "invertible A"
+  shows "invertible (k *\<^sub>R A) \<longleftrightarrow> k \<noteq> 0 \<and> invertible A"
+  by (simp add: assms scalar_invertible)
+
+lemma vector_transpose_matrix [simp]: "x v* transpose A = A *v x"
+  unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
+  by simp
+
+lemma transpose_matrix_vector [simp]: "transpose A *v x = x v* A"
+  unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
+  by simp
+
+lemma vector_scalar_commute:
+  fixes A :: "'a::{field}^'m^'n"
+  shows "A *v (c *s x) = c *s (A *v x)"
+  by (simp add: vector_scalar_mult_def matrix_vector_mult_def mult_ac sum_distrib_left)
+
+lemma scalar_vector_matrix_assoc:
+  fixes k :: "'a::{field}" and x :: "'a::{field}^'n" and A :: "'a^'m^'n"
+  shows "(k *s x) v* A = k *s (x v* A)"
+  by (metis transpose_matrix_vector vector_scalar_commute)
+ 
+lemma vector_matrix_mult_0 [simp]: "0 v* A = 0"
+  unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
+
+lemma vector_matrix_mult_0_right [simp]: "x v* 0 = 0"
+  unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
+
+lemma vector_matrix_mul_rid [simp]:
+  fixes v :: "('a::semiring_1)^'n"
+  shows "v v* mat 1 = v"
+  by (metis matrix_vector_mul_lid transpose_mat vector_transpose_matrix)
+
+lemma scaleR_vector_matrix_assoc:
+  fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
+  shows "(k *\<^sub>R x) v* A = k *\<^sub>R (x v* A)"
+  by (metis matrix_vector_mult_scaleR transpose_matrix_vector)
+
+lemma vector_scaleR_matrix_ac:
+  fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
+  shows "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)"
+proof -
+  have "x v* (k *\<^sub>R A) = (k *\<^sub>R x) v* A"
+    unfolding vector_matrix_mult_def
+    by (simp add: algebra_simps)
+  with scaleR_vector_matrix_assoc
+  show "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)"
+    by auto
+qed
+
 
 subsection\<open>Some bounds on components etc. relative to operator norm\<close>
 
@@ -402,22 +470,18 @@
   have fU: "finite ?U" by simp
   have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y)"
     unfolding matrix_right_invertible_surjective matrix_mult_sum surj_def
-    apply (subst eq_commute)
-    apply rule
-    done
+    by (simp add: eq_commute)
   have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> vec.span (columns A))" by blast
   { assume h: ?lhs
     { fix x:: "'a ^'n"
       from h[unfolded lhseq, rule_format, of x] obtain y :: "'a ^'m"
         where y: "sum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
       have "x \<in> vec.span (columns A)"
-        unfolding y[symmetric]
-        apply (rule vec.span_sum)
-        apply (rule vec.span_scale)
-        apply (rule vec.span_base)
-        unfolding columns_def
-        apply blast
-        done
+        unfolding y[symmetric] scalar_mult_eq_scaleR
+      proof (rule span_sum [OF span_mul])
+        show "column i A \<in> span (columns A)" for i
+          using columns_def span_inc by auto
+      qed
     }
     then have ?rhs unfolding rhseq by blast }
   moreover
@@ -428,17 +492,18 @@
         unfolding h by blast
       then have "?P y"
       proof (induction rule: vec.span_induct_alt)
-        show "\<exists>x::'a ^ 'm. sum (\<lambda>i. (x$i) *s column i A) ?U = 0"
-          by (rule exI[where x=0], simp)
+        case base
+        then show ?case
+          by (metis (full_types) matrix_mult_sum matrix_vector_mult_0_right)
       next
-        fix c y1 y2
-        assume y1: "y1 \<in> columns A" and y2: "?P y2"
+        case (step c y1 y2)
+        then obtain i where i: "i \<in> ?U" "y1 = column i A"
         from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
           unfolding columns_def by blast
-        from y2 obtain x:: "'a ^'m" where
-          x: "sum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
+        obtain x:: "real ^'m" where x: "sum (\<lambda>i. (x$i) *s column i A) ?U = y2"
+          using step by blast
         let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::'a^'m"
-        show "?P (c*s y1 + y2)"
+        show ?case
         proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left if_distribR cong del: if_weak_cong)
           fix j
           have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
@@ -446,9 +511,7 @@
             using i(1) by (simp add: field_simps)
           have "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
               else (x$xa) * ((column xa A$j))) ?U = sum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
-            apply (rule sum.cong[OF refl])
-            using th apply blast
-            done
+            by (rule sum.cong[OF refl]) (use th in blast)
           also have "\<dots> = sum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
             by (simp add: sum.distrib)
           also have "\<dots> = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
@@ -1038,8 +1101,8 @@
     obtain B where "independent B" "span(rows A) \<subseteq> span B"
               and B: "B \<subseteq> span(rows A)""card B = dim (span(rows A))"
       using basis_exists [of "span(rows A)"] by blast
-    then have eq: "span B = span(rows A)"
-      using span_subspace subspace_span by blast
+    with span_subspace have eq: "span B = span(rows A)"
+      by auto
     then have inj: "inj_on (( *v) A) (span B)"
       by (simp add: inj_on_def matrix_vector_mul_injective_on_rowspace)
     then have ind: "independent (( *v) A ` B)"
@@ -1208,15 +1271,13 @@
 
 definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)"
 
-lemma vector_1: "(vector[x]) $1 = x"
+lemma vector_1 [simp]: "(vector[x]) $1 = x"
   unfolding vector_def by simp
 
-lemma vector_2:
- "(vector[x,y]) $1 = x"
- "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"
+lemma vector_2 [simp]: "(vector[x,y]) $1 = x" "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"
   unfolding vector_def by simp_all
 
-lemma vector_3:
+lemma vector_3 [simp]:
  "(vector [x,y,z] ::('a::zero)^3)$1 = x"
  "(vector [x,y,z] ::('a::zero)^3)$2 = y"
  "(vector [x,y,z] ::('a::zero)^3)$3 = z"
@@ -1247,7 +1308,7 @@
   done
 
 lemma bounded_linear_component_cart[intro]: "bounded_linear (\<lambda>x::real^'n. x $ k)"
-  apply (rule bounded_linearI[where K=1])
+  apply (rule bounded_linear_intro[where K=1])
   using component_le_norm_cart[of _ k] unfolding real_norm_def by auto
 
 lemma interval_split_cart:
@@ -1263,4 +1324,4 @@
   bounded_linear.uniform_limit[OF bounded_linear_vec_nth]
   bounded_linear.uniform_limit[OF bounded_linear_component_cart]
 
-end
\ No newline at end of file
+end
--- a/src/HOL/Analysis/Cartesian_Space.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Cartesian_Space.thy	Thu May 03 15:07:14 2018 +0200
@@ -219,11 +219,7 @@
   { fix A A' :: "'a ^'n^'n"
     assume AA': "A ** A' = mat 1"
     have sA: "surj (( *v) A)"
-      unfolding surj_def
-      apply clarify
-      apply (rule_tac x="(A' *v y)" in exI)
-      apply (simp add: matrix_vector_mul_assoc AA')
-      done
+      using AA' matrix_right_invertible_surjective by auto
     from vec.linear_surjective_isomorphism[OF matrix_vector_mul_linear_gen sA]
     obtain f' :: "'a ^'n \<Rightarrow> 'a ^'n"
       where f': "Vector_Spaces.linear ( *s) ( *s) f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast
@@ -244,11 +240,64 @@
   shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). B ** A = mat 1)"
   by (metis invertible_def matrix_left_right_inverse)
 
-  lemma invertible_right_inverse:
+lemma invertible_right_inverse:
   fixes A :: "'a::{field}^'n^'n"
   shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). A** B = mat 1)"
   by (metis invertible_def matrix_left_right_inverse)
 
+lemma invertible_mult:
+  assumes inv_A: "invertible A"
+  and inv_B: "invertible B"
+  shows "invertible (A**B)"
+proof -
+  obtain A' where AA': "A ** A' = mat 1" and A'A: "A' ** A = mat 1" 
+    using inv_A unfolding invertible_def by blast
+  obtain B' where BB': "B ** B' = mat 1" and B'B: "B' ** B = mat 1" 
+    using inv_B unfolding invertible_def by blast
+  show ?thesis
+  proof (unfold invertible_def, rule exI[of _ "B'**A'"], rule conjI)
+    have "A ** B ** (B' ** A') = A ** (B ** (B' ** A'))" 
+      using matrix_mul_assoc[of A B "(B' ** A')", symmetric] .
+    also have "... = A ** (B ** B' ** A')" unfolding matrix_mul_assoc[of B "B'" "A'"] ..
+    also have "... = A ** (mat 1 ** A')" unfolding BB' ..
+    also have "... = A ** A'" unfolding matrix_mul_lid ..
+    also have "... = mat 1" unfolding AA' ..
+    finally show "A ** B ** (B' ** A') = mat (1::'a)" .    
+    have "B' ** A' ** (A ** B) = B' ** (A' ** (A ** B))" using matrix_mul_assoc[of B' A' "(A ** B)", symmetric] .
+    also have "... =  B' ** (A' ** A ** B)" unfolding matrix_mul_assoc[of A' A B] ..
+    also have "... =  B' ** (mat 1 ** B)" unfolding A'A ..
+    also have "... = B' ** B"  unfolding matrix_mul_lid ..
+    also have "... = mat 1" unfolding B'B ..
+    finally show "B' ** A' ** (A ** B) = mat 1" .
+  qed
+qed
+
+lemma transpose_invertible:
+  fixes A :: "real^'n^'n"
+  assumes "invertible A"
+  shows "invertible (transpose A)"
+  by (meson assms invertible_def matrix_left_right_inverse right_invertible_transpose)
+
+lemma vector_matrix_mul_assoc:
+  fixes v :: "('a::comm_semiring_1)^'n"
+  shows "(v v* M) v* N = v v* (M ** N)"
+proof -
+  from matrix_vector_mul_assoc
+  have "transpose N *v (transpose M *v v) = (transpose N ** transpose M) *v v" by fast
+  thus "(v v* M) v* N = v v* (M ** N)"
+    by (simp add: matrix_transpose_mul [symmetric])
+qed
+
+lemma matrix_scaleR_vector_ac:
+  fixes A :: "real^('m::finite)^'n"
+  shows "A *v (k *\<^sub>R v) = k *\<^sub>R A *v v"
+  by (metis matrix_vector_mult_scaleR transpose_scalar vector_scaleR_matrix_ac vector_transpose_matrix)
+
+lemma scaleR_matrix_vector_assoc:
+  fixes A :: "real^('m::finite)^'n"
+  shows "k *\<^sub>R (A *v v) = k *\<^sub>R A *v v"
+  by (metis matrix_scaleR_vector_ac matrix_vector_mult_scaleR)
+
 (*Finally, some interesting theorems and interpretations that don't appear in any file of the
   library.*)
 
--- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy	Thu May 03 15:07:14 2018 +0200
@@ -4130,7 +4130,7 @@
     have "winding_number \<gamma> y \<in> \<int>"  "winding_number \<gamma> z \<in>  \<int>"
       using that integer_winding_number [OF \<gamma> loop] sg \<open>y \<in> S\<close> by auto
     with ne show ?thesis
-      by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff)
+      by (auto simp: Ints_def reorient: of_int_diff)
   qed
   have cont: "continuous_on S (\<lambda>w. winding_number \<gamma> w)"
     using continuous_on_winding_number [OF \<gamma>] sg
@@ -6663,7 +6663,7 @@
       by (rule derivative_eq_intros | simp)+
     have y_le: "\<lbrakk>cmod (z - y) * 2 < r - cmod z\<rbrakk> \<Longrightarrow> cmod y \<le> cmod (of_real r + of_real (cmod z)) / 2" for z y
       using \<open>r > 0\<close>
-      apply (auto simp: algebra_simps norm_mult norm_divide norm_power of_real_add [symmetric] simp del: of_real_add)
+      apply (auto simp: algebra_simps norm_mult norm_divide norm_power reorient: of_real_add)
       using norm_triangle_ineq2 [of y z]
       apply (simp only: diff_le_eq norm_minus_commute mult_2)
       done
@@ -6671,7 +6671,7 @@
       using assms \<open>r > 0\<close> by simp
     moreover have "\<And>z. cmod z < r \<Longrightarrow> cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)"
       using \<open>r > 0\<close>
-      by (simp add: of_real_add [symmetric] del: of_real_add)
+      by (simp reorient: of_real_add)
     ultimately have sum: "\<And>z. cmod z < r \<Longrightarrow> summable (\<lambda>n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)"
       by (rule power_series_conv_imp_absconv_weak)
     have "\<exists>g g'. \<forall>z \<in> ball 0 r. (\<lambda>n.  (a n) * z ^ n) sums g z \<and>
@@ -6719,7 +6719,7 @@
       then have "0 \<le> r"
         by (meson less_eq_real_def norm_ge_zero order_trans)
       show ?thesis
-        using w by (simp add: dist_norm \<open>0\<le>r\<close> of_real_add [symmetric] del: of_real_add)
+        using w by (simp add: dist_norm \<open>0\<le>r\<close> reorient: of_real_add)
     qed
     have sum: "summable (\<lambda>n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))"
       using assms [OF inb] by (force simp add: summable_def dist_norm)
--- a/src/HOL/Analysis/Change_Of_Vars.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Change_Of_Vars.thy	Thu May 03 15:07:14 2018 +0200
@@ -1,3 +1,9 @@
+(*  Title:      HOL/Analysis/Change_Of_Vars.thy
+    Authors:    LC Paulson, based on material from HOL Light
+*)
+
+section\<open>Change of Variables Theorems\<close>
+
 theory Change_Of_Vars
   imports Vitali_Covering_Theorem Determinants
 
@@ -1274,8 +1280,7 @@
       proof (rule add_mono)
         have "(\<Sum>k\<le>n. real k * e * ?\<mu> (T k)) = (\<Sum>k\<le>n. integral (T k) (\<lambda>x. k * e))"
           by (simp add: lmeasure_integral [OF meas_t]
-                        integral_mult_right [symmetric] integral_mult_left [symmetric]
-                   del: integral_mult_right integral_mult_left)
+                   reorient: integral_mult_right integral_mult_left)
         also have "\<dots> \<le> (\<Sum>k\<le>n. integral (T k) (\<lambda>x.  (abs (det (matrix (f' x))))))"
         proof (rule sum_mono)
           fix k
@@ -1624,7 +1629,7 @@
   proof -
     obtain d where "d \<noteq> 0" and d: "\<And>y. f y = 0 \<Longrightarrow> d \<bullet> y = 0"
       using orthogonal_to_subspace_exists [OF less] orthogonal_def
-      by (metis (mono_tags, lifting) mem_Collect_eq span_clauses(1))
+      by (metis (mono_tags, lifting) mem_Collect_eq span_superset)
     then obtain k where "k > 0"
       and k: "\<And>e. e > 0 \<Longrightarrow> \<exists>y. y \<in> S - {0} \<and> norm y < e \<and> k * norm y \<le> \<bar>d \<bullet> y\<bar>"
       using lb by blast
--- a/src/HOL/Analysis/Complex_Analysis_Basics.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy	Thu May 03 15:07:14 2018 +0200
@@ -837,15 +837,13 @@
     qed
   } note ** = this
   show ?thesis
-  unfolding has_field_derivative_def
+    unfolding has_field_derivative_def
   proof (rule has_derivative_sequence [OF cvs _ _ x])
-    show "\<forall>n. \<forall>x\<in>s. (f n has_derivative (( * ) (f' n x))) (at x within s)"
-      by (metis has_field_derivative_def df)
-  next show "(\<lambda>n. f n x) \<longlonglongrightarrow> l"
+  show "(\<lambda>n. f n x) \<longlonglongrightarrow> l"
     by (rule tf)
-  next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
+  next show "\<And>e. e > 0 \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
     by (blast intro: **)
-  qed
+  qed (metis has_field_derivative_def df)
 qed
 
 lemma has_complex_derivative_series:
@@ -884,7 +882,7 @@
       by (metis df has_field_derivative_def mult_commute_abs)
   next show " ((\<lambda>n. f n x) sums l)"
     by (rule sf)
-  next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"
+  next show "\<And>e. e>0 \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"
     by (blast intro: **)
   qed
 qed
@@ -896,7 +894,7 @@
   assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
   assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)"
   assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" and x: "x \<in> s"
-  shows   "summable (\<lambda>n. f n x)" and "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x)"
+  shows  "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x)"
 proof -
   from assms(4) obtain g' where A: "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
     unfolding uniformly_convergent_on_def by blast
@@ -905,7 +903,6 @@
     by (intro has_field_derivative_series[of s f f' g' x0] assms A has_field_derivative_at_within)
   then obtain g where g: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x"
     "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast
-  from g[OF x] show "summable (\<lambda>n. f n x)" by (auto simp: summable_def)
   from g(2)[OF x] have g': "(g has_derivative ( * ) (g' x)) (at x)"
     by (simp add: has_field_derivative_def s)
   have "((\<lambda>x. \<Sum>n. f n x) has_derivative ( * ) (g' x)) (at x)"
@@ -915,15 +912,6 @@
     by (auto simp: summable_def field_differentiable_def has_field_derivative_def)
 qed
 
-lemma field_differentiable_series':
-  fixes f :: "nat \<Rightarrow> 'a::{real_normed_field,banach} \<Rightarrow> 'a"
-  assumes "convex s" "open s"
-  assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
-  assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)"
-  assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)"
-  shows   "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x0)"
-  using field_differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+
-
 subsection\<open>Bound theorem\<close>
 
 lemma field_differentiable_bound:
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Thu May 03 15:07:14 2018 +0200
@@ -68,8 +68,7 @@
   then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)"
     by auto
   then show ?thesis
-    using dim_span[of "cball (0 :: 'n::euclidean_space) e"]
-    by auto
+    using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp: dim_UNIV)
 qed
 
 lemma sum_not_0: "sum f A \<noteq> 0 \<Longrightarrow> \<exists>a \<in> A. f a \<noteq> 0"
@@ -119,8 +118,7 @@
     then have "x-a \<in> S" using assms by auto
     then have "x \<in> {a + v |v. v \<in> S}"
       apply auto
-      apply (rule exI[of _ "x-a"])
-      apply simp
+      apply (rule exI[of _ "x-a"], simp)
       done
     then have "x \<in> ((\<lambda>x. a+x) ` S)" by auto
   }
@@ -1301,7 +1299,7 @@
 proof -
   have *: "x - y + (y - z) = x - z" by auto
   show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *]
-    by (auto simp add:norm_minus_commute)
+    by (auto simp:norm_minus_commute)
 qed
 
 
@@ -1317,7 +1315,7 @@
   unfolding affine_def by auto
 
 lemma affine_sing [iff]: "affine {x}"
-  unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric])
+  unfolding affine_alt by (auto simp: scaleR_left_distrib [symmetric])
 
 lemma affine_UNIV [iff]: "affine UNIV"
   unfolding affine_def by auto
@@ -1350,293 +1348,200 @@
 lemma affine:
   fixes V::"'a::real_vector set"
   shows "affine V \<longleftrightarrow>
-    (\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> sum u s = 1 \<longrightarrow> (sum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)"
-  unfolding affine_def
-  apply rule
-  apply(rule, rule, rule)
-  apply(erule conjE)+
-  defer
-  apply (rule, rule, rule, rule, rule)
+         (\<forall>S u. finite S \<and> S \<noteq> {} \<and> S \<subseteq> V \<and> sum u S = 1 \<longrightarrow> (\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V)"
 proof -
-  fix x y u v
-  assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)"
-    "\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> sum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
-  then show "u *\<^sub>R x + v *\<^sub>R y \<in> V"
-    apply (cases "x = y")
-    using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]]
-      and as(1-3)
-    apply (auto simp add: scaleR_left_distrib[symmetric])
-    done
-next
-  fix s u
-  assume as: "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
-    "finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = (1::real)"
-  define n where "n = card s"
-  have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto
-  then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
-  proof (auto simp only: disjE)
-    assume "card s = 2"
-    then have "card s = Suc (Suc 0)"
-      by auto
-    then obtain a b where "s = {a, b}"
-      unfolding card_Suc_eq by auto
+  have "u *\<^sub>R x + v *\<^sub>R y \<in> V" if "x \<in> V" "y \<in> V" "u + v = (1::real)"
+    and *: "\<And>S u. \<lbrakk>finite S; S \<noteq> {}; S \<subseteq> V; sum u S = 1\<rbrakk> \<Longrightarrow> (\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V" for x y u v
+  proof (cases "x = y")
+    case True
+    then show ?thesis
+      using that by (metis scaleR_add_left scaleR_one)
+  next
+    case False
     then show ?thesis
-      using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5)
-      by (auto simp add: sum_clauses(2))
-  next
-    assume "card s > 2"
-    then show ?thesis using as and n_def
-    proof (induct n arbitrary: u s)
-      case 0
-      then show ?case by auto
+      using that *[of "{x,y}" "\<lambda>w. if w = x then u else v"] by auto
+  qed
+  moreover have "(\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V"
+                if *: "\<And>x y u v. \<lbrakk>x\<in>V; y\<in>V; u + v = 1\<rbrakk> \<Longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
+                  and "finite S" "S \<noteq> {}" "S \<subseteq> V" "sum u S = 1" for S u
+  proof -
+    define n where "n = card S"
+    consider "card S = 0" | "card S = 1" | "card S = 2" | "card S > 2" by linarith
+    then show "(\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V"
+    proof cases
+      assume "card S = 1"
+      then obtain a where "S={a}"
+        by (auto simp: card_Suc_eq)
+      then show ?thesis
+        using that by simp
+    next
+      assume "card S = 2"
+      then obtain a b where "S = {a, b}"
+        by (metis Suc_1 card_1_singletonE card_Suc_eq)
+      then show ?thesis
+        using *[of a b] that
+        by (auto simp: sum_clauses(2))
     next
-      case (Suc n)
-      fix s :: "'a set" and u :: "'a \<Rightarrow> real"
-      assume IA:
-        "\<And>u s.  \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s;
-          s \<noteq> {}; s \<subseteq> V; sum u s = 1; n = card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
-        and as:
-          "Suc n = card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
-           "finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = 1"
-      have "\<exists>x\<in>s. u x \<noteq> 1"
-      proof (rule ccontr)
-        assume "\<not> ?thesis"
-        then have "sum u s = real_of_nat (card s)"
-          unfolding card_eq_sum by auto
-        then show False
-          using as(7) and \<open>card s > 2\<close>
-          by (metis One_nat_def less_Suc0 Zero_not_Suc of_nat_1 of_nat_eq_iff numeral_2_eq_2)
+      assume "card S > 2"
+      then show ?thesis using that n_def
+      proof (induct n arbitrary: u S)
+        case 0
+        then show ?case by auto
+      next
+        case (Suc n u S)
+        have "sum u S = card S" if "\<not> (\<exists>x\<in>S. u x \<noteq> 1)"
+          using that unfolding card_eq_sum by auto
+        with Suc.prems obtain x where "x \<in> S" and x: "u x \<noteq> 1" by force
+        have c: "card (S - {x}) = card S - 1"
+          by (simp add: Suc.prems(3) \<open>x \<in> S\<close>)
+        have "sum u (S - {x}) = 1 - u x"
+          by (simp add: Suc.prems sum_diff1_ring \<open>x \<in> S\<close>)
+        with x have eq1: "inverse (1 - u x) * sum u (S - {x}) = 1"
+          by auto
+        have inV: "(\<Sum>y\<in>S - {x}. (inverse (1 - u x) * u y) *\<^sub>R y) \<in> V"
+        proof (cases "card (S - {x}) > 2")
+          case True
+          then have S: "S - {x} \<noteq> {}" "card (S - {x}) = n"
+            using Suc.prems c by force+
+          show ?thesis
+          proof (rule Suc.hyps)
+            show "(\<Sum>a\<in>S - {x}. inverse (1 - u x) * u a) = 1"
+              by (auto simp: eq1 sum_distrib_left[symmetric])
+          qed (use S Suc.prems True in auto)
+        next
+          case False
+          then have "card (S - {x}) = Suc (Suc 0)"
+            using Suc.prems c by auto
+          then obtain a b where ab: "(S - {x}) = {a, b}" "a\<noteq>b"
+            unfolding card_Suc_eq by auto
+          then show ?thesis
+            using eq1 \<open>S \<subseteq> V\<close>
+            by (auto simp: sum_distrib_left distrib_left intro!: Suc.prems(2)[of a b])
+        qed
+        have "u x + (1 - u x) = 1 \<Longrightarrow>
+          u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>y\<in>S - {x}. u y *\<^sub>R y) /\<^sub>R (1 - u x)) \<in> V"
+          by (rule Suc.prems) (use \<open>x \<in> S\<close> Suc.prems inV in \<open>auto simp: scaleR_right.sum\<close>)
+        moreover have "(\<Sum>a\<in>S. u a *\<^sub>R a) = u x *\<^sub>R x + (\<Sum>a\<in>S - {x}. u a *\<^sub>R a)"
+          by (meson Suc.prems(3) sum.remove \<open>x \<in> S\<close>)
+        ultimately show "(\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V"
+          by (simp add: x)
       qed
-      then obtain x where x:"x \<in> s" "u x \<noteq> 1" by auto
-
-      have c: "card (s - {x}) = card s - 1"
-        apply (rule card_Diff_singleton)
-        using \<open>x\<in>s\<close> as(4)
-        apply auto
-        done
-      have *: "s = insert x (s - {x})" "finite (s - {x})"
-        using \<open>x\<in>s\<close> and as(4) by auto
-      have **: "sum u (s - {x}) = 1 - u x"
-        using sum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto
-      have ***: "inverse (1 - u x) * sum u (s - {x}) = 1"
-        unfolding ** using \<open>u x \<noteq> 1\<close> by auto
-      have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
-      proof (cases "card (s - {x}) > 2")
-        case True
-        then have "s - {x} \<noteq> {}" "card (s - {x}) = n"
-          unfolding c and as(1)[symmetric]
-        proof (rule_tac ccontr)
-          assume "\<not> s - {x} \<noteq> {}"
-          then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp
-          then show False using True by auto
-        qed auto
-        then show ?thesis
-          apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
-          unfolding sum_distrib_left[symmetric]
-          using as and *** and True
-          apply auto
-          done
-      next
-        case False
-        then have "card (s - {x}) = Suc (Suc 0)"
-          using as(2) and c by auto
-        then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b"
-          unfolding card_Suc_eq by auto
-        then show ?thesis
-          using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
-          using *** *(2) and \<open>s \<subseteq> V\<close>
-          unfolding sum_distrib_left
-          by (auto simp add: sum_clauses(2))
-      qed
-      then have "u x + (1 - u x) = 1 \<Longrightarrow>
-          u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \<in> V"
-        apply -
-        apply (rule as(3)[rule_format])
-        unfolding  Real_Vector_Spaces.scaleR_right.sum
-        using x(1) as(6)
-        apply auto
-        done
-      then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
-        unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric]
-        apply (subst *)
-        unfolding sum_clauses(2)[OF *(2)]
-        using \<open>u x \<noteq> 1\<close>
-        apply auto
-        done
-    qed
-  next
-    assume "card s = 1"
-    then obtain a where "s={a}"
-      by (auto simp add: card_Suc_eq)
-    then show ?thesis
-      using as(4,5) by simp
-  qed (insert \<open>s\<noteq>{}\<close> \<open>finite s\<close>, auto)
-qed
+    qed (use \<open>S\<noteq>{}\<close> \<open>finite S\<close> in auto)
+  qed
+  ultimately show ?thesis
+    unfolding affine_def by meson
+qed
+
 
 lemma affine_hull_explicit:
-  "affine hull p =
-    {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> sum (\<lambda>v. (u v) *\<^sub>R v) s = y}"
-  apply (rule hull_unique)
-  apply (subst subset_eq)
-  prefer 3
-  apply rule
-  unfolding mem_Collect_eq
-  apply (erule exE)+
-  apply (erule conjE)+
-  prefer 2
-  apply rule
-proof -
-  fix x
-  assume "x\<in>p"
-  then show "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-    apply (rule_tac x="{x}" in exI)
-    apply (rule_tac x="\<lambda>x. 1" in exI)
-    apply auto
-    done
-next
-  fix t x s u
-  assume as: "p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}"
-    "s \<subseteq> p" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-  then show "x \<in> t"
-    using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]]
-    by auto
-next
-  show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}"
+  "affine hull p = {y. \<exists>S u. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p \<and> sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}"
+  (is "_ = ?rhs")
+proof (rule hull_unique)
+  show "p \<subseteq> ?rhs"
+  proof (intro subsetI CollectI exI conjI)
+    show "\<And>x. sum (\<lambda>z. 1) {x} = 1"
+      by auto
+  qed auto
+  show "?rhs \<subseteq> T" if "p \<subseteq> T" "affine T" for T
+    using that unfolding affine by blast
+  show "affine ?rhs"
     unfolding affine_def
-    apply (rule, rule, rule, rule, rule)
-    unfolding mem_Collect_eq
-  proof -
-    fix u v :: real
+  proof clarify
+    fix u v :: real and sx ux sy uy
     assume uv: "u + v = 1"
-    fix x
-    assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-    then obtain sx ux where
-      x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "sum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x"
-      by auto
-    fix y
-    assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
-    then obtain sy uy where
-      y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "sum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto
-    have xy: "finite (sx \<union> sy)"
-      using x(1) y(1) by auto
+      and x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "sum ux sx = (1::real)"
+      and y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "sum uy sy = (1::real)" 
     have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy"
       by auto
-    show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and>
-        sum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y"
-      apply (rule_tac x="sx \<union> sy" in exI)
-      apply (rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI)
-      unfolding scaleR_left_distrib sum.distrib if_smult scaleR_zero_left
-        ** sum.inter_restrict[OF xy, symmetric]
-      unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric]
-        and sum_distrib_left[symmetric]
-      unfolding x y
-      using x(1-3) y(1-3) uv
-      apply simp
-      done
+    show "\<exists>S w. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p \<and>
+        sum w S = 1 \<and> (\<Sum>v\<in>S. w v *\<^sub>R v) = u *\<^sub>R (\<Sum>v\<in>sx. ux v *\<^sub>R v) + v *\<^sub>R (\<Sum>v\<in>sy. uy v *\<^sub>R v)"
+    proof (intro exI conjI)
+      show "finite (sx \<union> sy)"
+        using x y by auto
+      show "sum (\<lambda>i. (if i\<in>sx then u * ux i else 0) + (if i\<in>sy then v * uy i else 0)) (sx \<union> sy) = 1"
+        using x y uv
+        by (simp add: sum_Un sum.distrib sum.inter_restrict[symmetric] sum_distrib_left [symmetric] **)
+      have "(\<Sum>i\<in>sx \<union> sy. ((if i \<in> sx then u * ux i else 0) + (if i \<in> sy then v * uy i else 0)) *\<^sub>R i)
+          = (\<Sum>i\<in>sx. (u * ux i) *\<^sub>R i) + (\<Sum>i\<in>sy. (v * uy i) *\<^sub>R i)"
+        using x y
+        unfolding scaleR_left_distrib scaleR_zero_left if_smult
+        by (simp add: sum_Un sum.distrib sum.inter_restrict[symmetric]  **)
+      also have "\<dots> = u *\<^sub>R (\<Sum>v\<in>sx. ux v *\<^sub>R v) + v *\<^sub>R (\<Sum>v\<in>sy. uy v *\<^sub>R v)"
+        unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] by blast
+      finally show "(\<Sum>i\<in>sx \<union> sy. ((if i \<in> sx then u * ux i else 0) + (if i \<in> sy then v * uy i else 0)) *\<^sub>R i) 
+                  = u *\<^sub>R (\<Sum>v\<in>sx. ux v *\<^sub>R v) + v *\<^sub>R (\<Sum>v\<in>sy. uy v *\<^sub>R v)" .
+    qed (use x y in auto)
   qed
 qed
 
 lemma affine_hull_finite:
-  assumes "finite s"
-  shows "affine hull s = {y. \<exists>u. sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
-  unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq
-  apply (rule, rule)
-  apply (erule exE)+
-  apply (erule conjE)+
-  defer
-  apply (erule exE)
-  apply (erule conjE)
+  assumes "finite S"
+  shows "affine hull S = {y. \<exists>u. sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}"
 proof -
-  fix x u
-  assume "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-  then show "\<exists>sa u. finite sa \<and>
-      \<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x"
-    apply (rule_tac x=s in exI, rule_tac x=u in exI)
-    using assms
-    apply auto
-    done
-next
-  fix x t u
-  assume "t \<subseteq> s"
-  then have *: "s \<inter> t = t"
-    by auto
-  assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
-  then show "\<exists>u. sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-    apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
-    unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms, symmetric] and *
-    apply auto
-    done
+  have *: "\<exists>h. sum h S = 1 \<and> (\<Sum>v\<in>S. h v *\<^sub>R v) = x" 
+    if "F \<subseteq> S" "finite F" "F \<noteq> {}" and sum: "sum u F = 1" and x: "(\<Sum>v\<in>F. u v *\<^sub>R v) = x" for x F u
+  proof -
+    have "S \<inter> F = F"
+      using that by auto
+    show ?thesis
+    proof (intro exI conjI)
+      show "(\<Sum>x\<in>S. if x \<in> F then u x else 0) = 1"
+        by (metis (mono_tags, lifting) \<open>S \<inter> F = F\<close> assms sum.inter_restrict sum)
+      show "(\<Sum>v\<in>S. (if v \<in> F then u v else 0) *\<^sub>R v) = x"
+        by (simp add: if_smult cong: if_cong) (metis (no_types) \<open>S \<inter> F = F\<close> assms sum.inter_restrict x)
+    qed
+  qed
+  show ?thesis
+    unfolding affine_hull_explicit using assms
+    by (fastforce dest: *)
 qed
 
 
 subsubsection%unimportant \<open>Stepping theorems and hence small special cases\<close>
 
 lemma affine_hull_empty[simp]: "affine hull {} = {}"
-  by (rule hull_unique) auto
-
-(*could delete: it simply rewrites sum expressions, but it's used twice*)
+  by simp
+
 lemma affine_hull_finite_step:
   fixes y :: "'a::real_vector"
-  shows
-    "(\<exists>u. sum u {} = w \<and> sum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
-    and
-    "finite s \<Longrightarrow>
-      (\<exists>u. sum u (insert a s) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow>
-      (\<exists>v u. sum u s = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs")
+  shows "finite S \<Longrightarrow>
+      (\<exists>u. sum u (insert a S) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a S) = y) \<longleftrightarrow>
+      (\<exists>v u. sum u S = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) S = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs")
 proof -
-  show ?th1 by simp
-  assume fin: "finite s"
+  assume fin: "finite S"
   show "?lhs = ?rhs"
   proof
     assume ?lhs
-    then obtain u where u: "sum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
+    then obtain u where u: "sum u (insert a S) = w \<and> (\<Sum>x\<in>insert a S. u x *\<^sub>R x) = y"
       by auto
     show ?rhs
-    proof (cases "a \<in> s")
+    proof (cases "a \<in> S")
       case True
-      then have *: "insert a s = s" by auto
-      show ?thesis
-        using u[unfolded *]
-        apply(rule_tac x=0 in exI)
-        apply auto
-        done
+      then show ?thesis
+        using u by (simp add: insert_absorb) (metis diff_zero real_vector.scale_zero_left)
     next
       case False
-      then show ?thesis
-        apply (rule_tac x="u a" in exI)
-        using u and fin
-        apply auto
-        done
+      show ?thesis
+        by (rule exI [where x="u a"]) (use u fin False in auto)
     qed
   next
     assume ?rhs
-    then obtain v u where vu: "sum u s = w - v"  "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
+    then obtain v u where vu: "sum u S = w - v"  "(\<Sum>x\<in>S. u x *\<^sub>R x) = y - v *\<^sub>R a"
       by auto
     have *: "\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)"
       by auto
     show ?lhs
-    proof (cases "a \<in> s")
+    proof (cases "a \<in> S")
       case True
-      then show ?thesis
-        apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
-        unfolding sum_clauses(2)[OF fin]
-        apply simp
-        unfolding scaleR_left_distrib and sum.distrib
-        unfolding vu and * and scaleR_zero_left
-        apply (auto simp add: sum.delta[OF fin])
-        done
+      show ?thesis
+        by (rule exI [where x="\<lambda>x. (if x=a then v else 0) + u x"])
+           (simp add: True scaleR_left_distrib sum.distrib sum_clauses fin vu * cong: if_cong)
     next
       case False
-      then have **:
-        "\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)"
-        "\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto
-      from False show ?thesis
-        apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
-        unfolding sum_clauses(2)[OF fin] and * using vu
-        using sum.cong [of s _ "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF _ **(2)]
-        using sum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)]
-        apply auto
-        done
+      then show ?thesis
+        apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI) 
+        apply (simp add: vu sum_clauses(2)[OF fin] *)
+        by (simp add: sum_delta_notmem(3) vu)
     qed
   qed
 qed
@@ -1652,7 +1557,7 @@
   have "?lhs = {y. \<exists>u. sum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
     using affine_hull_finite[of "{a,b}"] by auto
   also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
-    by (simp add: affine_hull_finite_step(2)[of "{b}" a])
+    by (simp add: affine_hull_finite_step[of "{b}" a])
   also have "\<dots> = ?rhs" unfolding * by auto
   finally show ?thesis by auto
 qed
@@ -1667,12 +1572,9 @@
   show ?thesis
     apply (simp add: affine_hull_finite affine_hull_finite_step)
     unfolding *
-    apply auto
-    apply (rule_tac x=v in exI)
-    apply (rule_tac x=va in exI)
-    apply auto
-    apply (rule_tac x=u in exI)
-    apply force
+    apply safe
+     apply (metis add.assoc)
+    apply (rule_tac x=u in exI, force)
     done
 qed
 
@@ -1710,56 +1612,57 @@
 subsubsection%unimportant \<open>Some relations between affine hull and subspaces\<close>
 
 lemma affine_hull_insert_subset_span:
-  "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
-  unfolding subset_eq Ball_def
-  unfolding affine_hull_explicit span_explicit mem_Collect_eq
-  apply (rule, rule)
-  apply (erule exE)+
-  apply (erule conjE)+
+  "affine hull (insert a S) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> S}}"
 proof -
-  fix x t u
-  assume as: "finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
-  have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}"
-    using as(3) by auto
-  then show "\<exists>v. x = a + v \<and> (\<exists>S u. v = (\<Sum>v\<in>S. u v *\<^sub>R v) \<and> finite S \<and> S \<subseteq> {x - a |x. x \<in> s} )"
-    apply (rule_tac x="x - a" in exI)
-    apply (rule conjI, simp)
-    apply (rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI)
-    apply (rule_tac x="\<lambda>x. u (x + a)" in exI)
-    by (simp_all add: as sum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib
-        sum_subtractf scaleR_left.sum[symmetric] sum_diff1 scaleR_left_diff_distrib)
+  have "\<exists>v T u. x = a + v \<and> (finite T \<and> T \<subseteq> {x - a |x. x \<in> S} \<and> (\<Sum>v\<in>T. u v *\<^sub>R v) = v)"
+    if "finite F" "F \<noteq> {}" "F \<subseteq> insert a S" "sum u F = 1" "(\<Sum>v\<in>F. u v *\<^sub>R v) = x"
+    for x F u
+  proof -
+    have *: "(\<lambda>x. x - a) ` (F - {a}) \<subseteq> {x - a |x. x \<in> S}"
+      using that by auto
+    show ?thesis
+    proof (intro exI conjI)
+      show "finite ((\<lambda>x. x - a) ` (F - {a}))"
+        by (simp add: that(1))
+      show "(\<Sum>v\<in>(\<lambda>x. x - a) ` (F - {a}). u(v+a) *\<^sub>R v) = x-a"
+        by (simp add: sum.reindex[unfolded inj_on_def] algebra_simps
+            sum_subtractf scaleR_left.sum[symmetric] sum_diff1 that)
+    qed (use \<open>F \<subseteq> insert a S\<close> in auto)
+  qed
+  then show ?thesis
+    unfolding affine_hull_explicit span_explicit by auto
 qed
 
 lemma affine_hull_insert_span:
-  assumes "a \<notin> s"
-  shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
-  apply (rule, rule affine_hull_insert_subset_span)
-  unfolding subset_eq Ball_def
-  unfolding affine_hull_explicit and mem_Collect_eq
-proof (rule, rule, erule exE, erule conjE)
-  fix y v
-  assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
-  then obtain t u where obt: "finite t" "t \<subseteq> {x - a |x. x \<in> s}" "a + (\<Sum>v\<in>t. u v *\<^sub>R v) = y"
-    unfolding span_explicit by auto
-  define f where "f = (\<lambda>x. x + a) ` t"
-  have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a"
-    unfolding f_def using obt by (auto simp add: sum.reindex[unfolded inj_on_def])
-  have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f"
-    using f(2) assms by auto
-  show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
-    apply (rule_tac x = "insert a f" in exI)
-    apply (rule_tac x = "\<lambda>x. if x=a then 1 - sum (\<lambda>x. u (x - a)) f else u (x - a)" in exI)
-    using assms and f
-    unfolding sum_clauses(2)[OF f(1)] and if_smult
-    unfolding sum.If_cases[OF f(1), of "\<lambda>x. x = a"]
-    apply (auto simp add: sum_subtractf scaleR_left.sum algebra_simps *)
-    done
+  assumes "a \<notin> S"
+  shows "affine hull (insert a S) = {a + v | v . v \<in> span {x - a | x.  x \<in> S}}"
+proof -
+  have *: "\<exists>G u. finite G \<and> G \<noteq> {} \<and> G \<subseteq> insert a S \<and> sum u G = 1 \<and> (\<Sum>v\<in>G. u v *\<^sub>R v) = y"
+    if "v \<in> span {x - a |x. x \<in> S}" "y = a + v" for y v
+  proof -
+    from that
+    obtain T u where u: "finite T" "T \<subseteq> {x - a |x. x \<in> S}" "a + (\<Sum>v\<in>T. u v *\<^sub>R v) = y"
+      unfolding span_explicit by auto
+    define F where "F = (\<lambda>x. x + a) ` T"
+    have F: "finite F" "F \<subseteq> S" "(\<Sum>v\<in>F. u (v - a) *\<^sub>R (v - a)) = y - a"
+      unfolding F_def using u by (auto simp: sum.reindex[unfolded inj_on_def])
+    have *: "F \<inter> {a} = {}" "F \<inter> - {a} = F"
+      using F assms by auto
+    show "\<exists>G u. finite G \<and> G \<noteq> {} \<and> G \<subseteq> insert a S \<and> sum u G = 1 \<and> (\<Sum>v\<in>G. u v *\<^sub>R v) = y"
+      apply (rule_tac x = "insert a F" in exI)
+      apply (rule_tac x = "\<lambda>x. if x=a then 1 - sum (\<lambda>x. u (x - a)) F else u (x - a)" in exI)
+      using assms F
+      apply (auto simp:  sum_clauses sum.If_cases if_smult sum_subtractf scaleR_left.sum algebra_simps *)
+      done
+  qed
+  show ?thesis
+    by (intro subset_antisym affine_hull_insert_subset_span) (auto simp: affine_hull_explicit dest!: *)
 qed
 
 lemma affine_hull_span:
-  assumes "a \<in> s"
-  shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}"
-  using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
+  assumes "a \<in> S"
+  shows "affine hull S = {a + v | v. v \<in> span {x - a | x. x \<in> S - {a}}}"
+  using affine_hull_insert_span[of a "S - {a}", unfolded insert_Diff[OF assms]] by auto
 
 
 subsubsection%unimportant \<open>Parallel affine sets\<close>
@@ -1769,17 +1672,12 @@
 
 lemma affine_parallel_expl_aux:
   fixes S T :: "'a::real_vector set"
-  assumes "\<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T"
+  assumes "\<And>x. x \<in> S \<longleftrightarrow> a + x \<in> T"
   shows "T = (\<lambda>x. a + x) ` S"
 proof -
-  {
-    fix x
-    assume "x \<in> T"
-    then have "( - a) + x \<in> S"
-      using assms by auto
-    then have "x \<in> ((\<lambda>x. a + x) ` S)"
-      using imageI[of "-a+x" S "(\<lambda>x. a+x)"] by auto
-  }
+  have "x \<in> ((\<lambda>x. a + x) ` S)" if "x \<in> T" for x
+    using that
+    by (simp add: image_iff) (metis add.commute diff_add_cancel assms)
   moreover have "T \<ge> (\<lambda>x. a + x) ` S"
     using assms by auto
   ultimately show ?thesis by auto
@@ -1791,9 +1689,7 @@
 
 lemma affine_parallel_reflex: "affine_parallel S S"
   unfolding affine_parallel_def
-  apply (rule exI[of _ "0"])
-  apply auto
-  done
+  using image_add_0 by blast
 
 lemma affine_parallel_commut:
   assumes "affine_parallel A B"
@@ -2109,7 +2005,7 @@
   shows "c *\<^sub>R x \<in> cone hull S"
   by (metis assms cone_cone_hull hull_inc mem_cone)
 
-lemma%important cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}"
+proposition%important cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}"
   (is "?lhs = ?rhs")
 proof%unimportant -
   {
@@ -2135,8 +2031,7 @@
     assume "x \<in> S"
     then have "1 *\<^sub>R x \<in> ?rhs"
       apply auto
-      apply (rule_tac x = 1 in exI)
-      apply auto
+      apply (rule_tac x = 1 in exI, auto)
       done
     then have "x \<in> ?rhs" by auto
   }
@@ -2169,7 +2064,7 @@
   then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> ( *\<^sub>R) c ` S = S)"
     using cone_iff[of S] assms by auto
   then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> ( *\<^sub>R) c ` closure S = closure S)"
-    using closure_subset by (auto simp add: closure_scaleR)
+    using closure_subset by (auto simp: closure_scaleR)
   then show ?thesis
     using False cone_iff[of "closure S"] by auto
 qed
@@ -2194,66 +2089,60 @@
    "~ affine_dependent s \<Longrightarrow> ~ affine_dependent(s - t)"
 by (meson Diff_subset affine_dependent_subset)
 
-lemma%important affine_dependent_explicit:
+proposition%important affine_dependent_explicit:
   "affine_dependent p \<longleftrightarrow>
-    (\<exists>s u. finite s \<and> s \<subseteq> p \<and> sum u s = 0 \<and>
-      (\<exists>v\<in>s. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) s = 0)"
-  unfolding%unimportant affine_dependent_def affine_hull_explicit mem_Collect_eq
-  apply rule
-  apply (erule bexE, erule exE, erule exE)
-  apply (erule conjE)+
-  defer
-  apply (erule exE, erule exE)
-  apply (erule conjE)+
-  apply (erule bexE)
+    (\<exists>S u. finite S \<and> S \<subseteq> p \<and> sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) S = 0)"
 proof -
-  fix x s u
-  assume as: "x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-  have "x \<notin> s" using as(1,4) by auto
-  show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> sum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
-    apply (rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI)
-    unfolding if_smult and sum_clauses(2)[OF as(2)] and sum_delta_notmem[OF \<open>x\<notin>s\<close>] and as
-    using as
-    apply auto
-    done
-next
-  fix s u v
-  assume as: "finite s" "s \<subseteq> p" "sum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0"
-  have "s \<noteq> {v}"
-    using as(3,6) by auto
-  then show "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-    apply (rule_tac x=v in bexI)
-    apply (rule_tac x="s - {v}" in exI)
-    apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI)
-    unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric]
-    unfolding sum_distrib_left[symmetric] and sum_diff1[OF as(1)]
-    using as
-    apply auto
-    done
+  have "\<exists>S u. finite S \<and> S \<subseteq> p \<and> sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> (\<Sum>w\<in>S. u w *\<^sub>R w) = 0"
+    if "(\<Sum>w\<in>S. u w *\<^sub>R w) = x" "x \<in> p" "finite S" "S \<noteq> {}" "S \<subseteq> p - {x}" "sum u S = 1" for x S u
+  proof (intro exI conjI)
+    have "x \<notin> S" 
+      using that by auto
+    then show "(\<Sum>v \<in> insert x S. if v = x then - 1 else u v) = 0"
+      using that by (simp add: sum_delta_notmem)
+    show "(\<Sum>w \<in> insert x S. (if w = x then - 1 else u w) *\<^sub>R w) = 0"
+      using that \<open>x \<notin> S\<close> by (simp add: if_smult sum_delta_notmem cong: if_cong)
+  qed (use that in auto)
+  moreover have "\<exists>x\<in>p. \<exists>S u. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p - {x} \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = x"
+    if "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0" "finite S" "S \<subseteq> p" "sum u S = 0" "v \<in> S" "u v \<noteq> 0" for S u v
+  proof (intro bexI exI conjI)
+    have "S \<noteq> {v}"
+      using that by auto
+    then show "S - {v} \<noteq> {}"
+      using that by auto
+    show "(\<Sum>x \<in> S - {v}. - (1 / u v) * u x) = 1"
+      unfolding sum_distrib_left[symmetric] sum_diff1[OF \<open>finite S\<close>] by (simp add: that)
+    show "(\<Sum>x\<in>S - {v}. (- (1 / u v) * u x) *\<^sub>R x) = v"
+      unfolding sum_distrib_left [symmetric] scaleR_scaleR[symmetric]
+                scaleR_right.sum [symmetric] sum_diff1[OF \<open>finite S\<close>] 
+      using that by auto
+    show "S - {v} \<subseteq> p - {v}"
+      using that by auto
+  qed (use that in auto)
+  ultimately show ?thesis
+    unfolding affine_dependent_def affine_hull_explicit by auto
 qed
 
 lemma affine_dependent_explicit_finite:
-  fixes s :: "'a::real_vector set"
-  assumes "finite s"
-  shows "affine_dependent s \<longleftrightarrow>
-    (\<exists>u. sum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) s = 0)"
+  fixes S :: "'a::real_vector set"
+  assumes "finite S"
+  shows "affine_dependent S \<longleftrightarrow>
+    (\<exists>u. sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) S = 0)"
   (is "?lhs = ?rhs")
 proof
   have *: "\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)"
     by auto
   assume ?lhs
   then obtain t u v where
-    "finite t" "t \<subseteq> s" "sum u t = 0" "v\<in>t" "u v \<noteq> 0"  "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
+    "finite t" "t \<subseteq> S" "sum u t = 0" "v\<in>t" "u v \<noteq> 0"  "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
     unfolding affine_dependent_explicit by auto
   then show ?rhs
     apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
-    apply auto unfolding * and sum.inter_restrict[OF assms, symmetric]
-    unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>]
-    apply auto
+    apply (auto simp: * sum.inter_restrict[OF assms, symmetric] Int_absorb1[OF \<open>t\<subseteq>S\<close>])
     done
 next
   assume ?rhs
-  then obtain u v where "sum u s = 0"  "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
+  then obtain u v where "sum u S = 0"  "v\<in>S" "u v \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
     by auto
   then show ?lhs unfolding affine_dependent_explicit
     using assms by auto
@@ -2267,15 +2156,15 @@
   by (rule Topological_Spaces.topological_space_class.connectedD)
 
 lemma convex_connected:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "convex s"
-  shows "connected s"
+  fixes S :: "'a::real_normed_vector set"
+  assumes "convex S"
+  shows "connected S"
 proof (rule connectedI)
   fix A B
-  assume "open A" "open B" "A \<inter> B \<inter> s = {}" "s \<subseteq> A \<union> B"
+  assume "open A" "open B" "A \<inter> B \<inter> S = {}" "S \<subseteq> A \<union> B"
   moreover
-  assume "A \<inter> s \<noteq> {}" "B \<inter> s \<noteq> {}"
-  then obtain a b where a: "a \<in> A" "a \<in> s" and b: "b \<in> B" "b \<in> s" by auto
+  assume "A \<inter> S \<noteq> {}" "B \<inter> S \<noteq> {}"
+  then obtain a b where a: "a \<in> A" "a \<in> S" and b: "b \<in> B" "b \<in> S" by auto
   define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u
   then have "continuous_on {0 .. 1} f"
     by (auto intro!: continuous_intros)
@@ -2286,8 +2175,8 @@
     using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def)
   moreover have "b \<in> B \<inter> f ` {0 .. 1}"
     using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def)
-  moreover have "f ` {0 .. 1} \<subseteq> s"
-    using \<open>convex s\<close> a b unfolding convex_def f_def by auto
+  moreover have "f ` {0 .. 1} \<subseteq> S"
+    using \<open>convex S\<close> a b unfolding convex_def f_def by auto
   ultimately show False by auto
 qed
 
@@ -2372,7 +2261,7 @@
 lemma convex_ball [iff]:
   fixes x :: "'a::real_normed_vector"
   shows "convex (ball x e)"
-proof (auto simp add: convex_def)
+proof (auto simp: convex_def)
   fix y z
   assume yz: "dist x y < e" "dist x z < e"
   fix u v :: real
@@ -2403,7 +2292,7 @@
     then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e"
       using convex_bound_le[OF yz uv] by auto
   }
-  then show ?thesis by (auto simp add: convex_def Ball_def)
+  then show ?thesis by (auto simp: convex_def Ball_def)
 qed
 
 lemma connected_ball [iff]:
@@ -2481,8 +2370,8 @@
 proof
   show "convex hull (s \<times> t) \<subseteq> (convex hull s) \<times> (convex hull t)"
     by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull)
-  have "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)"
-  proof (intro hull_induct)
+  have "(x, y) \<in> convex hull (s \<times> t)" if x: "x \<in> convex hull s" and y: "y \<in> convex hull t" for x y
+  proof (rule hull_induct [OF x], rule hull_induct [OF y])
     fix x y assume "x \<in> s" and "y \<in> t"
     then show "(x, y) \<in> convex hull (s \<times> t)"
       by (simp add: hull_inc)
@@ -2492,22 +2381,22 @@
       by (intro convex_linear_vimage convex_translation convex_convex_hull,
         simp add: linear_iff)
     also have "?S = {y. (x, y) \<in> convex hull (s \<times> t)}"
-      by (auto simp add: image_def Bex_def)
+      by (auto simp: image_def Bex_def)
     finally show "convex {y. (x, y) \<in> convex hull (s \<times> t)}" .
   next
-    show "convex {x. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)}"
-    proof (unfold Collect_ball_eq, rule convex_INT [rule_format])
+    show "convex {x. (x, y) \<in> convex hull s \<times> t}"
+    proof -
       fix y let ?S = "((\<lambda>x. (x, 0)) -` (\<lambda>p. (0, - y) + p) ` (convex hull s \<times> t))"
       have "convex ?S"
       by (intro convex_linear_vimage convex_translation convex_convex_hull,
         simp add: linear_iff)
       also have "?S = {x. (x, y) \<in> convex hull (s \<times> t)}"
-        by (auto simp add: image_def Bex_def)
+        by (auto simp: image_def Bex_def)
       finally show "convex {x. (x, y) \<in> convex hull (s \<times> t)}" .
     qed
   qed
   then show "(convex hull s) \<times> (convex hull t) \<subseteq> convex hull (s \<times> t)"
-    unfolding subset_eq split_paired_Ball_Sigma .
+    unfolding subset_eq split_paired_Ball_Sigma by blast
 qed
 
 
@@ -2520,118 +2409,114 @@
   by (rule hull_unique) auto
 
 lemma convex_hull_insert:
-  fixes s :: "'a::real_vector set"
-  assumes "s \<noteq> {}"
-  shows "convex hull (insert a s) =
-    {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and> b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}"
+  fixes S :: "'a::real_vector set"
+  assumes "S \<noteq> {}"
+  shows "convex hull (insert a S) =
+         {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and> b \<in> (convex hull S) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}"
   (is "_ = ?hull")
-  apply (rule, rule hull_minimal, rule)
-  unfolding insert_iff
-  prefer 3
-  apply rule
-proof -
+proof (intro equalityI hull_minimal subsetI)
   fix x
-  assume x: "x = a \<or> x \<in> s"
+  assume "x \<in> insert a S"
+  then have "\<exists>u\<ge>0. \<exists>v\<ge>0. u + v = 1 \<and> (\<exists>b. b \<in> convex hull S \<and> x = u *\<^sub>R a + v *\<^sub>R b)"
+  unfolding insert_iff
+  proof
+    assume "x = a"
+    then show ?thesis
+      by (rule_tac x=1 in exI) (use assms hull_subset in fastforce)
+  next
+    assume "x \<in> S"
+    with hull_subset[of S convex] show ?thesis
+      by force
+  qed
   then show "x \<in> ?hull"
-    apply rule
-    unfolding mem_Collect_eq
-    apply (rule_tac x=1 in exI)
-    defer
-    apply (rule_tac x=0 in exI)
-    using assms hull_subset[of s convex]
-    apply auto
-    done
+    by simp
 next
   fix x
   assume "x \<in> ?hull"
-  then obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b"
+  then obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull S" "x = u *\<^sub>R a + v *\<^sub>R b"
     by auto
-  have "a \<in> convex hull insert a s" "b \<in> convex hull insert a s"
-    using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4)
+  have "a \<in> convex hull insert a S" "b \<in> convex hull insert a S"
+    using hull_mono[of S "insert a S" convex] hull_mono[of "{a}" "insert a S" convex] and obt(4)
     by auto
-  then show "x \<in> convex hull insert a s"
+  then show "x \<in> convex hull insert a S"
     unfolding obt(5) using obt(1-3)
     by (rule convexD [OF convex_convex_hull])
 next
   show "convex ?hull"
   proof (rule convexI)
     fix x y u v
-    assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull"
-    from as(4) obtain u1 v1 b1 where
-      obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1"
+    assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" and x: "x \<in> ?hull" and y: "y \<in> ?hull"
+    from x obtain u1 v1 b1 where
+      obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull S" and xeq: "x = u1 *\<^sub>R a + v1 *\<^sub>R b1"
       by auto
-    from as(5) obtain u2 v2 b2 where
-      obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2"
+    from y obtain u2 v2 b2 where
+      obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull S" and yeq: "y = u2 *\<^sub>R a + v2 *\<^sub>R b2"
       by auto
     have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x"
-      by (auto simp add: algebra_simps)
-    have **: "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y =
+      by (auto simp: algebra_simps)
+    have "\<exists>b \<in> convex hull S. u *\<^sub>R x + v *\<^sub>R y =
       (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)"
     proof (cases "u * v1 + v * v2 = 0")
       case True
       have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x"
-        by (auto simp add: algebra_simps)
-      from True have ***: "u * v1 = 0" "v * v2 = 0"
-        using mult_nonneg_nonneg[OF \<open>u\<ge>0\<close> \<open>v1\<ge>0\<close>] mult_nonneg_nonneg[OF \<open>v\<ge>0\<close> \<open>v2\<ge>0\<close>]
+        by (auto simp: algebra_simps)
+      have eq0: "u * v1 = 0" "v * v2 = 0"
+        using True mult_nonneg_nonneg[OF \<open>u\<ge>0\<close> \<open>v1\<ge>0\<close>] mult_nonneg_nonneg[OF \<open>v\<ge>0\<close> \<open>v2\<ge>0\<close>]
         by arith+
       then have "u * u1 + v * u2 = 1"
         using as(3) obt1(3) obt2(3) by auto
       then show ?thesis
-        unfolding obt1(5) obt2(5) *
-        using assms hull_subset[of s convex]
-        by (auto simp add: *** scaleR_right_distrib)
+        using "*" eq0 as obt1(4) xeq yeq by auto
     next
       case False
       have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)"
-        using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
+        using as(3) obt1(3) obt2(3) by (auto simp: field_simps)
       also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)"
-        using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
+        using as(3) obt1(3) obt2(3) by (auto simp: field_simps)
       also have "\<dots> = u * v1 + v * v2"
         by simp
       finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
-      have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2"
+      let ?b = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2"
+      have zeroes: "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2"
         using as(1,2) obt1(1,2) obt2(1,2) by auto
-      then show ?thesis
-        unfolding obt1(5) obt2(5)
-        unfolding * and **
-        using False
-        apply (rule_tac
-          x = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI)
-        defer
-        apply (rule convexD [OF convex_convex_hull])
-        using obt1(4) obt2(4)
-        unfolding add_divide_distrib[symmetric] and zero_le_divide_iff
-        apply (auto simp add: scaleR_left_distrib scaleR_right_distrib)
-        done
+      show ?thesis
+      proof
+        show "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (?b - (u * u1) *\<^sub>R ?b - (v * u2) *\<^sub>R ?b)"
+          unfolding xeq yeq * **
+          using False by (auto simp: scaleR_left_distrib scaleR_right_distrib)
+        show "?b \<in> convex hull S"
+          using False zeroes obt1(4) obt2(4)
+          by (auto simp: convexD [OF convex_convex_hull] scaleR_left_distrib scaleR_right_distrib  add_divide_distrib[symmetric]  zero_le_divide_iff)
+      qed
     qed
+    then obtain b where b: "b \<in> convex hull S" 
+       "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" ..
+
     have u1: "u1 \<le> 1"
       unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto
     have u2: "u2 \<le> 1"
       unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto
     have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v"
-      apply (rule add_mono)
-      apply (rule_tac [!] mult_right_mono)
-      using as(1,2) obt1(1,2) obt2(1,2)
-      apply auto
-      done
+    proof (rule add_mono)
+      show "u1 * u \<le> max u1 u2 * u" "u2 * v \<le> max u1 u2 * v"
+        by (simp_all add: as mult_right_mono)
+    qed
     also have "\<dots> \<le> 1"
       unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto
-    finally show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull"
-      unfolding mem_Collect_eq
-      apply (rule_tac x="u * u1 + v * u2" in exI)
-      apply (rule conjI)
-      defer
-      apply (rule_tac x="1 - u * u1 - v * u2" in exI)
-      unfolding Bex_def
-      using as(1,2) obt1(1,2) obt2(1,2) **
-      apply (auto simp add: algebra_simps)
-      done
+    finally have le1: "u1 * u + u2 * v \<le> 1" .    
+    show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull"
+    proof (intro CollectI exI conjI)
+      show "0 \<le> u * u1 + v * u2"
+        by (simp add: as(1) as(2) obt1(1) obt2(1))
+      show "0 \<le> 1 - u * u1 - v * u2"
+        by (simp add: le1 diff_diff_add mult.commute)
+    qed (use b in \<open>auto simp: algebra_simps\<close>)
   qed
 qed
 
 lemma convex_hull_insert_alt:
    "convex hull (insert a S) =
-      (if S = {} then {a}
+     (if S = {} then {a}
       else {(1 - u) *\<^sub>R a + u *\<^sub>R x |x u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> convex hull S})"
   apply (auto simp: convex_hull_insert)
   using diff_eq_eq apply fastforce
@@ -2639,147 +2524,81 @@
 
 subsubsection%unimportant \<open>Explicit expression for convex hull\<close>
 
-lemma%important convex_hull_indexed:
-  fixes s :: "'a::real_vector set"
-  shows "convex hull s =
-    {y. \<exists>k u x.
-      (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
-      (sum u {1..k} = 1) \<and> (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}"
-  (is "?xyz = ?hull")
-  apply%unimportant (rule hull_unique)
-  apply rule
-  defer
-  apply (rule convexI)
-proof -
-  fix x
-  assume "x\<in>s"
-  then show "x \<in> ?hull"
-    unfolding mem_Collect_eq
-    apply (rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI)
-    apply auto
-    done
+proposition%important convex_hull_indexed:
+  fixes S :: "'a::real_vector set"
+  shows "convex hull S =
+    {y. \<exists>k u x. (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> S) \<and>
+                (sum u {1..k} = 1) \<and> (\<Sum>i = 1..k. u i *\<^sub>R x i) = y}"
+    (is "?xyz = ?hull")
+proof (rule hull_unique [OF _ convexI])
+  show "S \<subseteq> ?hull" 
+    by (clarsimp, rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI, auto)
 next
-  fix t
-  assume as: "s \<subseteq> t" "convex t"
-  show "?hull \<subseteq> t"
-    apply rule
-    unfolding mem_Collect_eq
-    apply (elim exE conjE)
-  proof -
-    fix x k u y
-    assume assm:
-      "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s"
-      "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
-    show "x\<in>t"
-      unfolding assm(3) [symmetric]
-      apply (rule as(2)[unfolded convex, rule_format])
-      using assm(1,2) as(1) apply auto
-      done
-  qed
+  fix T
+  assume "S \<subseteq> T" "convex T"
+  then show "?hull \<subseteq> T"
+    by (blast intro: convex_sum)
 next
   fix x y u v
   assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
   assume xy: "x \<in> ?hull" "y \<in> ?hull"
   from xy obtain k1 u1 x1 where
-    x: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "sum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x"
+    x [rule_format]: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> S" 
+                      "sum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x"
     by auto
   from xy obtain k2 u2 x2 where
-    y: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "sum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y"
+    y [rule_format]: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> S" 
+                     "sum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y"
     by auto
-  have *: "\<And>P (x1::'a) x2 s1 s2 i.
-    (if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)"
-    "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
-    prefer 3
-    apply (rule, rule)
-    unfolding image_iff
-    apply (rule_tac x = "x - k1" in bexI)
-    apply (auto simp add: not_le)
-    done
+  have *: "\<And>P (x::'a) y s t i. (if P i then s else t) *\<^sub>R (if P i then x else y) = (if P i then s *\<^sub>R x else t *\<^sub>R y)"
+          "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
+    by auto
   have inj: "inj_on (\<lambda>i. i + k1) {1..k2}"
     unfolding inj_on_def by auto
+  let ?uu = "\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)"
+  let ?xx = "\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)"
   show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull"
-    apply rule
-    apply (rule_tac x="k1 + k2" in exI)
-    apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
-    apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI)
-    apply (rule, rule)
-    defer
-    apply rule
-    unfolding * and sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and
-      sum.reindex[OF inj] and o_def Collect_mem_eq
-    unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric]
-  proof -
-    fix i
-    assume i: "i \<in> {1..k1+k2}"
-    show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and>
-      (if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s"
-    proof (cases "i\<in>{1..k1}")
-      case True
-      then show ?thesis
-        using uv(1) x(1)[THEN bspec[where x=i]] by auto
-    next
-      case False
-      define j where "j = i - k1"
-      from i False have "j \<in> {1..k2}"
-        unfolding j_def by auto
-      then show ?thesis
-        using False uv(2) y(1)[THEN bspec[where x=j]]
-        by (auto simp: j_def[symmetric])
-    qed
-  qed (auto simp add: not_le x(2,3) y(2,3) uv(3))
+  proof (intro CollectI exI conjI ballI)
+    show "0 \<le> ?uu i" "?xx i \<in> S" if "i \<in> {1..k1+k2}" for i
+      using that by (auto simp add: le_diff_conv uv(1) x(1) uv(2) y(1))
+    show "(\<Sum>i = 1..k1 + k2. ?uu i) = 1"  "(\<Sum>i = 1..k1 + k2. ?uu i *\<^sub>R ?xx i) = u *\<^sub>R x + v *\<^sub>R y"
+      unfolding * sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]]
+        sum.reindex[OF inj] Collect_mem_eq o_def
+      unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric]
+      by (simp_all add: sum_distrib_left[symmetric]  x(2,3) y(2,3) uv(3))
+  qed 
 qed
 
 lemma convex_hull_finite:
-  fixes s :: "'a::real_vector set"
-  assumes "finite s"
-  shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
-    sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}"
-  (is "?HULL = ?set")
-proof (rule hull_unique, auto simp add: convex_def[of ?set])
+  fixes S :: "'a::real_vector set"
+  assumes "finite S"
+  shows "convex hull S = {y. \<exists>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) S = y}"
+  (is "?HULL = _")
+proof (rule hull_unique [OF _ convexI]; clarify)
   fix x
-  assume "x \<in> s"
-  then show "\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x"
-    apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI)
-    apply auto
-    unfolding sum.delta'[OF assms] and sum_delta''[OF assms]
-    apply auto
-    done
+  assume "x \<in> S"
+  then show "\<exists>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>x\<in>S. u x *\<^sub>R x) = x"
+    by (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) (auto simp: sum.delta'[OF assms] sum_delta''[OF assms])
 next
   fix u v :: real
   assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
-  fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "sum ux s = (1::real)"
-  fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "sum uy s = (1::real)"
-  {
-    fix x
-    assume "x\<in>s"
-    then have "0 \<le> u * ux x + v * uy x"
-      using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2)
-      by auto
-  }
+  fix ux assume ux [rule_format]: "\<forall>x\<in>S. 0 \<le> ux x" "sum ux S = (1::real)"
+  fix uy assume uy [rule_format]: "\<forall>x\<in>S. 0 \<le> uy x" "sum uy S = (1::real)"
+  have "0 \<le> u * ux x + v * uy x" if "x\<in>S" for x
+    by (simp add: that uv ux(1) uy(1))
   moreover
-  have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
-    unfolding sum.distrib and sum_distrib_left[symmetric] and ux(2) uy(2)
+  have "(\<Sum>x\<in>S. u * ux x + v * uy x) = 1"
+    unfolding sum.distrib and sum_distrib_left[symmetric] ux(2) uy(2)
     using uv(3) by auto
   moreover
-  have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
-    unfolding scaleR_left_distrib and sum.distrib and scaleR_scaleR[symmetric]
-      and scaleR_right.sum [symmetric]
+  have "(\<Sum>x\<in>S. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>S. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>S. uy x *\<^sub>R x)"
+    unfolding scaleR_left_distrib sum.distrib scaleR_scaleR[symmetric] scaleR_right.sum [symmetric]
     by auto
   ultimately
-  show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> sum uc s = 1 \<and>
-      (\<Sum>x\<in>s. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
-    apply (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI)
-    apply auto
-    done
-next
-  fix t
-  assume t: "s \<subseteq> t" "convex t"
-  fix u
-  assume u: "\<forall>x\<in>s. 0 \<le> u x" "sum u s = (1::real)"
-  then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t"
-    using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]]
-    using assms and t(1) by auto
-qed
+  show "\<exists>uc. (\<forall>x\<in>S. 0 \<le> uc x) \<and> sum uc S = 1 \<and>
+             (\<Sum>x\<in>S. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>S. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>S. uy x *\<^sub>R x)"
+    by (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI, auto)
+qed (use assms in \<open>auto simp: convex_explicit\<close>)
 
 
 subsubsection%unimportant \<open>Another formulation from Lars Schewe\<close>
@@ -2787,7 +2606,7 @@
 lemma convex_hull_explicit:
   fixes p :: "'a::real_vector set"
   shows "convex hull p =
-    {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
+    {y. \<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}"
   (is "?lhs = ?rhs")
 proof -
   {
@@ -2817,10 +2636,9 @@
       using sum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
       unfolding scaleR_left.sum using obt(3) by auto
     ultimately
-    have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+    have "\<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = x"
       apply (rule_tac x="y ` {1..k}" in exI)
-      apply (rule_tac x="\<lambda>v. sum u {i\<in>{1..k}. y i = v}" in exI)
-      apply auto
+      apply (rule_tac x="\<lambda>v. sum u {i\<in>{1..k}. y i = v}" in exI, auto)
       done
     then have "x\<in>?rhs" by auto
   }
@@ -2828,55 +2646,50 @@
   {
     fix y
     assume "y\<in>?rhs"
-    then obtain s u where
-      obt: "finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+    then obtain S u where
+      obt: "finite S" "S \<subseteq> p" "\<forall>x\<in>S. 0 \<le> u x" "sum u S = 1" "(\<Sum>v\<in>S. u v *\<^sub>R v) = y"
       by auto
 
-    obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s"
+    obtain f where f: "inj_on f {1..card S}" "f ` {1..card S} = S"
       using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto
 
     {
       fix i :: nat
-      assume "i\<in>{1..card s}"
-      then have "f i \<in> s"
-        apply (subst f(2)[symmetric])
-        apply auto
-        done
+      assume "i\<in>{1..card S}"
+      then have "f i \<in> S"
+        using f(2) by blast
       then have "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto
     }
-    moreover have *: "finite {1..card s}" by auto
+    moreover have *: "finite {1..card S}" by auto
     {
       fix y
-      assume "y\<in>s"
-      then obtain i where "i\<in>{1..card s}" "f i = y"
-        using f using image_iff[of y f "{1..card s}"]
+      assume "y\<in>S"
+      then obtain i where "i\<in>{1..card S}" "f i = y"
+        using f using image_iff[of y f "{1..card S}"]
         by auto
-      then have "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}"
+      then have "{x. Suc 0 \<le> x \<and> x \<le> card S \<and> f x = y} = {i}"
         apply auto
         using f(1)[unfolded inj_on_def]
-        apply(erule_tac x=x in ballE)
-        apply auto
-        done
-      then have "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
-      then have "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
-          "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
-        by (auto simp add: sum_constant_scaleR)
+        by (metis One_nat_def atLeastAtMost_iff)
+      then have "card {x. Suc 0 \<le> x \<and> x \<le> card S \<and> f x = y} = 1" by auto
+      then have "(\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x)) = u y"
+          "(\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
+        by (auto simp: sum_constant_scaleR)
     }
-    then have "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y"
+    then have "(\<Sum>x = 1..card S. u (f x)) = 1" "(\<Sum>i = 1..card S. u (f i) *\<^sub>R f i) = y"
       unfolding sum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
         and sum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
       unfolding f
-      using sum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
-      using sum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
+      using sum.cong [of S S "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
+      using sum.cong [of S S "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x))" u]
       unfolding obt(4,5)
       by auto
     ultimately
     have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> sum u {1..k} = 1 \<and>
         (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
-      apply (rule_tac x="card s" in exI)
+      apply (rule_tac x="card S" in exI)
       apply (rule_tac x="u \<circ> f" in exI)
-      apply (rule_tac x=f in exI)
-      apply fastforce
+      apply (rule_tac x=f in exI, fastforce)
       done
     then have "y \<in> ?lhs"
       unfolding convex_hull_indexed by auto
@@ -2889,70 +2702,57 @@
 subsubsection%unimportant \<open>A stepping theorem for that expansion\<close>
 
 lemma convex_hull_finite_step:
-  fixes s :: "'a::real_vector set"
-  assumes "finite s"
+  fixes S :: "'a::real_vector set"
+  assumes "finite S"
   shows
-    "(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> sum u (insert a s) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y)
-      \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)"
+    "(\<exists>u. (\<forall>x\<in>insert a S. 0 \<le> u x) \<and> sum u (insert a S) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a S) = y)
+      \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) S = y - v *\<^sub>R a)"
   (is "?lhs = ?rhs")
-proof (rule, case_tac[!] "a\<in>s")
-  assume "a \<in> s"
-  then have *: "insert a s = s" by auto
+proof (rule, case_tac[!] "a\<in>S")
+  assume "a \<in> S"
+  then have *: "insert a S = S" by auto
   assume ?lhs
   then show ?rhs
-    unfolding *
-    apply (rule_tac x=0 in exI)
-    apply auto
-    done
+    unfolding *  by (rule_tac x=0 in exI, auto)
 next
   assume ?lhs
   then obtain u where
-      u: "\<forall>x\<in>insert a s. 0 \<le> u x" "sum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
+      u: "\<forall>x\<in>insert a S. 0 \<le> u x" "sum u (insert a S) = w" "(\<Sum>x\<in>insert a S. u x *\<^sub>R x) = y"
     by auto
-  assume "a \<notin> s"
+  assume "a \<notin> S"
   then show ?rhs
     apply (rule_tac x="u a" in exI)
     using u(1)[THEN bspec[where x=a]]
     apply simp
     apply (rule_tac x=u in exI)
-    using u[unfolded sum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close>
+    using u[unfolded sum_clauses(2)[OF assms]] and \<open>a\<notin>S\<close>
     apply auto
     done
 next
-  assume "a \<in> s"
-  then have *: "insert a s = s" by auto
-  have fin: "finite (insert a s)" using assms by auto
+  assume "a \<in> S"
+  then have *: "insert a S = S" by auto
+  have fin: "finite (insert a S)" using assms by auto
   assume ?rhs
-  then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
+  then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>S. 0 \<le> u x" "sum u S = w - v" "(\<Sum>x\<in>S. u x *\<^sub>R x) = y - v *\<^sub>R a"
     by auto
   show ?lhs
     apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI)
     unfolding scaleR_left_distrib and sum.distrib and sum_delta''[OF fin] and sum.delta'[OF fin]
     unfolding sum_clauses(2)[OF assms]
-    using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close>
+    using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>S\<close>
     apply auto
     done
 next
   assume ?rhs
-  then obtain v u where
-    uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
+  then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>S. 0 \<le> u x" "sum u S = w - v" "(\<Sum>x\<in>S. u x *\<^sub>R x) = y - v *\<^sub>R a"
     by auto
-  moreover
-  assume "a \<notin> s"
+  moreover assume "a \<notin> S"
   moreover
-  have "(\<Sum>x\<in>s. if a = x then v else u x) = sum u s"
-    and "(\<Sum>x\<in>s. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
-    apply (rule_tac sum.cong) apply rule
-    defer
-    apply (rule_tac sum.cong) apply rule
-    using \<open>a \<notin> s\<close>
-    apply auto
-    done
+  have "(\<Sum>x\<in>S. if a = x then v else u x) = sum u S"  "(\<Sum>x\<in>S. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>S. u x *\<^sub>R x)"
+    using \<open>a \<notin> S\<close>
+    by (auto simp: intro!: sum.cong)
   ultimately show ?lhs
-    apply (rule_tac x="\<lambda>x. if a = x then v else u x" in exI)
-    unfolding sum_clauses(2)[OF assms]
-    apply auto
-    done
+    by (rule_tac x="\<lambda>x. if a = x then v else u x" in exI) (auto simp: sum_clauses(2)[OF assms])
 qed
 
 
@@ -2969,12 +2769,9 @@
     unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc]
     apply auto
     apply (rule_tac x=v in exI)
-    apply (rule_tac x="1 - v" in exI)
-    apply simp
-    apply (rule_tac x=u in exI)
-    apply simp
-    apply (rule_tac x="\<lambda>x. v" in exI)
-    apply simp
+    apply (rule_tac x="1 - v" in exI, simp)
+    apply (rule_tac x=u in exI, simp)
+    apply (rule_tac x="\<lambda>x. v" in exI, simp)
     done
 qed
 
@@ -2989,7 +2786,7 @@
     unfolding *
     apply auto
     apply (rule_tac[!] x=u in exI)
-    apply (auto simp add: algebra_simps)
+    apply (auto simp: algebra_simps)
     done
 qed
 
@@ -2999,22 +2796,17 @@
   have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}"
     by auto
   have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
-    by (auto simp add: field_simps)
+    by (auto simp: field_simps)
   show ?thesis
     unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and *
     unfolding convex_hull_finite_step[OF fin(3)]
-    apply (rule Collect_cong)
-    apply simp
+    apply (rule Collect_cong, simp)
     apply auto
     apply (rule_tac x=va in exI)
-    apply (rule_tac x="u c" in exI)
-    apply simp
-    apply (rule_tac x="1 - v - w" in exI)
-    apply simp
-    apply (rule_tac x=v in exI)
-    apply simp
-    apply (rule_tac x="\<lambda>x. w" in exI)
-    apply simp
+    apply (rule_tac x="u c" in exI, simp)
+    apply (rule_tac x="1 - v - w" in exI, simp)
+    apply (rule_tac x=v in exI, simp)
+    apply (rule_tac x="\<lambda>x. w" in exI, simp)
     done
 qed
 
@@ -3025,7 +2817,7 @@
     by auto
   show ?thesis
     unfolding convex_hull_3
-    apply (auto simp add: *)
+    apply (auto simp: *)
     apply (rule_tac x=v in exI)
     apply (rule_tac x=w in exI)
     apply (simp add: algebra_simps)
@@ -3084,36 +2876,24 @@
     apply auto
     done
   have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
-    unfolding sum_clauses(2)[OF fin]
-    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
-    apply auto
-    unfolding *
-    apply auto
-    done
+    unfolding sum_clauses(2)[OF fin] * using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> by auto
   moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0"
-    apply (rule_tac x="v + a" in bexI)
-    using obt(3,4) and \<open>0\<notin>S\<close>
-    unfolding t_def
-    apply auto
-    done
+    using obt(3,4) \<open>0\<notin>S\<close>
+    by (rule_tac x="v + a" in bexI) (auto simp: t_def)
   moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)"
-    apply (rule sum.cong)
-    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
-    apply auto
-    done
+    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> by (auto intro!: sum.cong)
   have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
     unfolding scaleR_left.sum
     unfolding t_def and sum.reindex[OF inj] and o_def
     using obt(5)
-    by (auto simp add: sum.distrib scaleR_right_distrib)
+    by (auto simp: sum.distrib scaleR_right_distrib)
   then have "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0"
     unfolding sum_clauses(2)[OF fin]
     using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
-    by (auto simp add: *)
+    by (auto simp: *)
   ultimately show ?thesis
     unfolding affine_dependent_explicit
-    apply (rule_tac x="insert a t" in exI)
-    apply auto
+    apply (rule_tac x="insert a t" in exI, auto)
     done
 qed
 
@@ -3130,10 +2910,8 @@
       using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1]
       apply (erule_tac x="2*\<^sub>R x" in ballE)
       apply (erule_tac x="2*\<^sub>R y" in ballE)
-      apply (erule_tac x="1/2" in allE)
-      apply simp
-      apply (erule_tac x="1/2" in allE)
-      apply auto
+      apply (erule_tac x="1/2" in allE, simp)
+      apply (erule_tac x="1/2" in allE, auto)
       done
   }
   then show ?thesis
@@ -3150,49 +2928,36 @@
   have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
     by auto
   have "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
-    unfolding *
-    apply (rule card_image)
-    unfolding inj_on_def
-    apply auto
-    done
+    unfolding * by (simp add: card_image inj_on_def)
   also have "\<dots> > DIM('a)" using assms(2)
     unfolding card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] by auto
   finally show ?thesis
     apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
     apply (rule dependent_imp_affine_dependent)
-    apply (rule dependent_biggerset)
-    apply auto
+    apply (rule dependent_biggerset, auto)
     done
 qed
 
 lemma affine_dependent_biggerset_general:
-  assumes "finite (s :: 'a::euclidean_space set)"
-    and "card s \<ge> dim s + 2"
-  shows "affine_dependent s"
+  assumes "finite (S :: 'a::euclidean_space set)"
+    and "card S \<ge> dim S + 2"
+  shows "affine_dependent S"
 proof -
-  from assms(2) have "s \<noteq> {}" by auto
-  then obtain a where "a\<in>s" by auto
-  have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
+  from assms(2) have "S \<noteq> {}" by auto
+  then obtain a where "a\<in>S" by auto
+  have *: "{x - a |x. x \<in> S - {a}} = (\<lambda>x. x - a) ` (S - {a})"
     by auto
-  have **: "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
-    unfolding *
-    apply (rule card_image)
-    unfolding inj_on_def
-    apply auto
-    done
-  have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
-    apply (rule subset_le_dim)
-    unfolding subset_eq
-    using \<open>a\<in>s\<close>
-    apply (auto simp add:span_base span_diff)
-    done
-  also have "\<dots> < dim s + 1" by auto
-  also have "\<dots> \<le> card (s - {a})"
+  have **: "card {x - a |x. x \<in> S - {a}} = card (S - {a})"
+    by (metis (no_types, lifting) "*" card_image diff_add_cancel inj_on_def)
+  have "dim {x - a |x. x \<in> S - {a}} \<le> dim S"
+    using \<open>a\<in>S\<close> by (auto simp: span_superset span_diff intro: subset_le_dim)
+  also have "\<dots> < dim S + 1" by auto
+  also have "\<dots> \<le> card (S - {a})"
     using assms
-    using card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>]
+    using card_Diff_singleton[OF assms(1) \<open>a\<in>S\<close>]
     by auto
   finally show ?thesis
-    apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
+    apply (subst insert_Diff[OF \<open>a\<in>S\<close>, symmetric])
     apply (rule dependent_imp_affine_dependent)
     apply (rule dependent_biggerset_general)
     unfolding **
@@ -3384,10 +3149,10 @@
     using assms by auto
   then have h0: "independent  ((\<lambda>x. -a + x) ` (S-{a}))"
     using affine_dependent_iff_dependent2 assms by auto
-  then obtain B where B:
+  obtain B where B:
     "(\<lambda>x. -a+x) ` (S - {a}) \<subseteq> B \<and> B \<subseteq> (\<lambda>x. -a+x) ` V \<and> independent B \<and> (\<lambda>x. -a+x) ` V \<subseteq> span B"
-     using maximal_independent_subset_extend[of "(\<lambda>x. -a+x) ` (S-{a})" "(\<lambda>x. -a + x) ` V"] assms
-     by blast
+    using assms
+    by (blast intro: maximal_independent_subset_extend[OF _ h0, of "(\<lambda>x. -a + x) ` V"])
   define T where "T = (\<lambda>x. a+x) ` insert 0 B"
   then have "T = insert a ((\<lambda>x. a+x) ` B)"
     by auto
@@ -3457,8 +3222,7 @@
       some_eq_ex[of "\<lambda>d. \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1"]
     apply auto
     apply (rule exI[of _ "int (card B) - (1 :: int)"])
-    apply (rule exI[of _ "B"])
-    apply auto
+    apply (rule exI[of _ "B"], auto)
     done
 qed
 
@@ -3493,10 +3257,7 @@
     then have "card ((\<lambda>x. -a + x) ` (B - {a})) > 0"
       using fin by auto
     moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
-       apply (rule card_image)
-       using translate_inj_on
-       apply (auto simp del: uminus_add_conv_diff)
-       done
+      by (rule card_image) (use translate_inj_on in blast)
     ultimately have "card (B-{a}) > 0" by auto
     then have *: "finite (B - {a})"
       using card_gt_0_iff[of "(B - {a})"] by auto
@@ -3592,23 +3353,10 @@
     by auto
   let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
   have "\<exists>f. linear f \<and> f ` B = d \<and> f ` span B = ?t \<and> inj_on f (span B)"
-    apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"])
-    apply (rule subspace_span)
-    apply (rule subspace_substandard)
-    defer
-    apply (rule span_superset)
-    apply (rule assms)
-    defer
-    unfolding dim_span[of B]
-    apply(rule B)
-    unfolding span_substd_basis[OF d, symmetric]
-    apply (rule span_superset)
-    apply (rule independent_substdbasis[OF d])
-    apply rule
-    apply assumption
-    unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d]
-    apply auto
-    done
+  proof (intro basis_to_basis_subspace_isomorphism subspace_span subspace_substandard span_inc)
+    show "d \<subseteq> {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
+      using d inner_not_same_Basis by blast
+  qed (auto simp: span_substd_basis independent_substdbasis dim_substandard d t B assms)
   with t \<open>card B = dim B\<close> d show ?thesis by auto
 qed
 
@@ -3694,7 +3442,7 @@
   assume "a \<noteq> b"
   then have "aff_dim{a,b} = card{a,b} - 1"
     using affine_independent_2 [of a b] aff_dim_affine_independent by fastforce
-  also have "... = 1"
+  also have "\<dots> = 1"
     using \<open>a \<noteq> b\<close> by simp
   finally show "aff_dim {a, b} = 1" .
 qed
@@ -3923,9 +3671,9 @@
     by blast
   then have "card {b - a |b. b \<in> S - {a}} = card ((\<lambda>b. b-a) ` (S - {a}))"
     by simp
-  also have "... = card (S - {a})"
+  also have "\<dots> = card (S - {a})"
     by (metis (no_types, lifting) card_image diff_add_cancel inj_onI)
-  also have "... = card S - 1"
+  also have "\<dots> = card S - 1"
     by (simp add: aff_independent_finite assms)
   finally have 4: "card {b - a |b. b \<in> S - {a}} = card S - 1" .
   have "finite S"
@@ -4156,8 +3904,7 @@
   assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
   then obtain N where "?P N" by auto
   then have "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n"
-    apply (rule_tac ex_least_nat_le)
-    apply auto
+    apply (rule_tac ex_least_nat_le, auto)
     done
   then obtain n where "?P n" and smallest: "\<forall>k<n. \<not> ?P k"
     by blast
@@ -4178,8 +3925,7 @@
     proof (rule ccontr, simp add: not_less)
       assume as:"\<forall>x\<in>s. 0 \<le> w x"
       then have "sum w (s - {v}) \<ge> 0"
-        apply (rule_tac sum_nonneg)
-        apply auto
+        apply (rule_tac sum_nonneg, auto)
         done
       then have "sum w s > 0"
         unfolding sum.remove[OF obt(1) \<open>v\<in>s\<close>]
@@ -4229,7 +3975,7 @@
       apply (rule_tac x="(s - {a})" in exI)
       apply (rule_tac x="\<lambda>v. u v + t * w v" in exI)
       using obt(1-3) and t and a
-      apply (auto simp add: * scaleR_left_distrib)
+      apply (auto simp: * scaleR_left_distrib)
       done
     then show False
       using smallest[THEN spec[where x="n - 1"]] by auto
@@ -4245,9 +3991,8 @@
         (is "?lhs = ?rhs")
 proof
   show "?lhs \<subseteq> ?rhs"
-    apply (subst convex_hull_caratheodory_aff_dim)
-    apply clarify
-    apply (rule_tac x="s" in exI)
+    apply (subst convex_hull_caratheodory_aff_dim, clarify)
+    apply (rule_tac x=s in exI)
     apply (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull])
     done
 next
@@ -4272,7 +4017,7 @@
 next
   fix x
   assume "x \<in> ?rhs" then show "x \<in> ?lhs"
-    by (auto simp add: convex_hull_explicit)
+    by (auto simp: convex_hull_explicit)
 qed
 
 theorem%important caratheodory:
@@ -4331,14 +4076,13 @@
 qed
 
 lemma mem_rel_interior: "x \<in> rel_interior S \<longleftrightarrow> (\<exists>T. open T \<and> x \<in> T \<inter> S \<and> T \<inter> affine hull S \<subseteq> S)"
-  by (auto simp add: rel_interior)
+  by (auto simp: rel_interior)
 
 lemma mem_rel_interior_ball:
   "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S)"
   apply (simp add: rel_interior, safe)
-  apply (force simp add: open_contains_ball)
-  apply (rule_tac x = "ball x e" in exI)
-  apply simp
+  apply (force simp: open_contains_ball)
+  apply (rule_tac x = "ball x e" in exI, simp)
   done
 
 lemma rel_interior_ball:
@@ -4348,10 +4092,9 @@
 lemma mem_rel_interior_cball:
   "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S)"
   apply (simp add: rel_interior, safe)
-  apply (force simp add: open_contains_cball)
+  apply (force simp: open_contains_cball)
   apply (rule_tac x = "ball x e" in exI)
-  apply (simp add: subset_trans [OF ball_subset_cball])
-  apply auto
+  apply (simp add: subset_trans [OF ball_subset_cball], auto)
   done
 
 lemma rel_interior_cball:
@@ -4359,7 +4102,7 @@
   using mem_rel_interior_cball [of _ S] by auto
 
 lemma rel_interior_empty [simp]: "rel_interior {} = {}"
-   by (auto simp add: rel_interior_def)
+   by (auto simp: rel_interior_def)
 
 lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
   by (metis affine_hull_eq affine_sing)
@@ -4367,8 +4110,7 @@
 lemma rel_interior_sing [simp]:
     fixes a :: "'n::euclidean_space"  shows "rel_interior {a} = {a}"
   apply (auto simp: rel_interior_ball)
-  apply (rule_tac x=1 in exI)
-  apply force
+  apply (rule_tac x=1 in exI, force)
   done
 
 lemma subset_rel_interior:
@@ -4376,16 +4118,16 @@
   assumes "S \<subseteq> T"
     and "affine hull S = affine hull T"
   shows "rel_interior S \<subseteq> rel_interior T"
-  using assms by (auto simp add: rel_interior_def)
+  using assms by (auto simp: rel_interior_def)
 
 lemma rel_interior_subset: "rel_interior S \<subseteq> S"
-  by (auto simp add: rel_interior_def)
+  by (auto simp: rel_interior_def)
 
 lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S"
-  using rel_interior_subset by (auto simp add: closure_def)
+  using rel_interior_subset by (auto simp: closure_def)
 
 lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S"
-  by (auto simp add: rel_interior interior_def)
+  by (auto simp: rel_interior interior_def)
 
 lemma interior_rel_interior:
   fixes S :: "'n::euclidean_space set"
@@ -4464,7 +4206,7 @@
     fix y
     assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S"
     have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
-      using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
+      using \<open>e > 0\<close> by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib)
     have "x \<in> affine hull S"
       using assms hull_subset[of S] by auto
     moreover have "1 / e + - ((1 - e) / e) = 1"
@@ -4476,17 +4218,17 @@
       unfolding dist_norm norm_scaleR[symmetric]
       apply (rule arg_cong[where f=norm])
       using \<open>e > 0\<close>
-      apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
+      apply (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
       done
     also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)"
       by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
     also have "\<dots> < d"
       using as[unfolded dist_norm] and \<open>e > 0\<close>
-      by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
+      by (auto simp:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
     finally have "y \<in> S"
       apply (subst *)
       apply (rule assms(1)[unfolded convex_alt,rule_format])
-      apply (rule d[unfolded subset_eq,rule_format])
+      apply (rule d[THEN subsetD])
       unfolding mem_ball
       using assms(3-5) **
       apply auto
@@ -4518,7 +4260,7 @@
     then have "y \<in> interior {a..}"
       apply (simp add: mem_interior)
       apply (rule_tac x="(y-a)" in exI)
-      apply (auto simp add: dist_norm)
+      apply (auto simp: dist_norm)
       done
   }
   moreover
@@ -4528,7 +4270,7 @@
     then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}"
       using mem_interior_cball[of y "{a..}"] by auto
     moreover from e have "y - e \<in> cball y e"
-      by (auto simp add: cball_def dist_norm)
+      by (auto simp: cball_def dist_norm)
     ultimately have "a \<le> y - e" by blast
     then have "a < y" using e by auto
   }
@@ -4558,7 +4300,7 @@
     then have "y \<in> interior {..a}"
       apply (simp add: mem_interior)
       apply (rule_tac x="(a-y)" in exI)
-      apply (auto simp add: dist_norm)
+      apply (auto simp: dist_norm)
       done
   }
   moreover
@@ -4568,7 +4310,7 @@
     then obtain e where e: "e > 0" "cball y e \<subseteq> {..a}"
       using mem_interior_cball[of y "{..a}"] by auto
     moreover from e have "y + e \<in> cball y e"
-      by (auto simp add: cball_def dist_norm)
+      by (auto simp: cball_def dist_norm)
     ultimately have "a \<ge> y + e" by auto
     then have "a > y" using e by auto
   }
@@ -4578,9 +4320,9 @@
 lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<..<b :: real}"
 proof-
   have "{a..b} = {a..} \<inter> {..b}" by auto
-  also have "interior ... = {a<..} \<inter> {..<b}"
+  also have "interior \<dots> = {a<..} \<inter> {..<b}"
     by (simp add: interior_real_semiline interior_real_semiline')
-  also have "... = {a<..<b}" by auto
+  also have "\<dots> = {a<..<b}" by auto
   finally show ?thesis .
 qed
 
@@ -4599,7 +4341,7 @@
 lemma frontier_real_Iic [simp]:
   fixes a :: real
   shows "frontier {..a} = {a}"
-  unfolding frontier_def by (auto simp add: interior_real_semiline')
+  unfolding frontier_def by (auto simp: interior_real_semiline')
 
 lemma rel_interior_real_box [simp]:
   fixes a b :: real
@@ -4638,8 +4380,7 @@
 
 lemma openin_rel_interior: "openin (subtopology euclidean (affine hull S)) (rel_interior S)"
   apply (simp add: rel_interior_def)
-  apply (subst openin_subopen)
-  apply blast
+  apply (subst openin_subopen, blast)
   done
 
 lemma openin_set_rel_interior:
@@ -4718,8 +4459,7 @@
   proof (cases "x \<in> S")
     case True
     then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close>
-      apply (rule_tac bexI[where x=x])
-      apply (auto)
+      apply (rule_tac bexI[where x=x], auto)
       done
   next
     case False
@@ -4739,7 +4479,7 @@
     next
       case False
       then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
-        using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by (auto)
+        using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto
       then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)"
         using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
       then show ?thesis
@@ -4755,11 +4495,11 @@
   define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)"
   have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
     unfolding z_def using \<open>e > 0\<close>
-    by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
+    by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
   have zball: "z \<in> ball c d"
     using mem_ball z_def dist_norm[of c]
     using y and assms(4,5)
-    by (auto simp add:field_simps norm_minus_commute)
+    by (auto simp:field_simps norm_minus_commute)
   have "x \<in> affine hull S"
     using closure_affine_hull assms by auto
   moreover have "y \<in> affine hull S"
@@ -4770,7 +4510,7 @@
     using z_def affine_affine_hull[of S]
       mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"]
       assms
-    by (auto simp add: field_simps)
+    by (auto simp: field_simps)
   then have "z \<in> S" using d zball by auto
   obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d"
     using zball open_ball[of c d] openE[of "ball c d" z] by auto
@@ -4868,24 +4608,17 @@
 lemma affine_hull_linear_image:
   assumes "bounded_linear f"
   shows "f ` (affine hull s) = affine hull f ` s"
-  apply rule
-  unfolding subset_eq ball_simps
-  apply (rule_tac[!] hull_induct, rule hull_inc)
-  prefer 3
-  apply (erule imageE)
-  apply (rule_tac x=xa in image_eqI)
-  apply assumption
-  apply (rule hull_subset[unfolded subset_eq, rule_format])
-  apply assumption
 proof -
   interpret f: bounded_linear f by fact
-  show "affine {x. f x \<in> affine hull f ` s}"
+  have "affine {x. f x \<in> affine hull f ` s}"
     unfolding affine_def
-    by (auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format])
-  show "affine {x. x \<in> f ` (affine hull s)}"
+    by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format])
+  moreover have "affine {x. x \<in> f ` (affine hull s)}"
     using affine_affine_hull[unfolded affine_def, of s]
-    unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
-qed auto
+    unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric])
+  ultimately show ?thesis
+    by (auto simp: hull_inc elim!: hull_induct)
+qed 
 
 
 lemma rel_interior_injective_on_span_linear_image:
@@ -5012,114 +4745,77 @@
 subsection%unimportant \<open>Openness and compactness are preserved by convex hull operation\<close>
 
 lemma open_convex_hull[intro]:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "open s"
-  shows "open (convex hull s)"
-  unfolding open_contains_cball convex_hull_explicit
-  unfolding mem_Collect_eq ball_simps(8)
-proof (rule, rule)
-  fix a
-  assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a"
-  then obtain t u where obt: "finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a"
-    by auto
+  fixes S :: "'a::real_normed_vector set"
+  assumes "open S"
+  shows "open (convex hull S)"
+proof (clarsimp simp: open_contains_cball convex_hull_explicit)
+  fix T and u :: "'a\<Rightarrow>real"
+  assume obt: "finite T" "T\<subseteq>S" "\<forall>x\<in>T. 0 \<le> u x" "sum u T = 1" 
 
   from assms[unfolded open_contains_cball] obtain b
-    where b: "\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s"
-    using bchoice[of s "\<lambda>x e. e > 0 \<and> cball x e \<subseteq> s"] by auto
-  have "b ` t \<noteq> {}"
+    where b: "\<And>x. x\<in>S \<Longrightarrow> 0 < b x \<and> cball x (b x) \<subseteq> S" by metis
+  have "b ` T \<noteq> {}"
     using obt by auto
-  define i where "i = b ` t"
-
-  show "\<exists>e > 0.
-    cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}"
-    apply (rule_tac x = "Min i" in exI)
-    unfolding subset_eq
-    apply rule
-    defer
-    apply rule
-    unfolding mem_Collect_eq
-  proof -
+  define i where "i = b ` T"
+  let ?\<Phi> = "\<lambda>y. \<exists>F. finite F \<and> F \<subseteq> S \<and> (\<exists>u. (\<forall>x\<in>F. 0 \<le> u x) \<and> sum u F = 1 \<and> (\<Sum>v\<in>F. u v *\<^sub>R v) = y)"
+  let ?a = "\<Sum>v\<in>T. u v *\<^sub>R v"
+  show "\<exists>e > 0. cball ?a e \<subseteq> {y. ?\<Phi> y}"
+  proof (intro exI subsetI conjI)
     show "0 < Min i"
-      unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` t\<noteq>{}\<close>]
-      using b
-      apply simp
-      apply rule
-      apply (erule_tac x=x in ballE)
-      using \<open>t\<subseteq>s\<close>
-      apply auto
-      done
+      unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` T\<noteq>{}\<close>]
+      using b \<open>T\<subseteq>S\<close> by auto
   next
     fix y
-    assume "y \<in> cball a (Min i)"
-    then have y: "norm (a - y) \<le> Min i"
+    assume "y \<in> cball ?a (Min i)"
+    then have y: "norm (?a - y) \<le> Min i"
       unfolding dist_norm[symmetric] by auto
-    {
-      fix x
-      assume "x \<in> t"
+    { fix x
+      assume "x \<in> T"
       then have "Min i \<le> b x"
-        unfolding i_def
-        apply (rule_tac Min_le)
-        using obt(1)
-        apply auto
-        done
-      then have "x + (y - a) \<in> cball x (b x)"
+        by (simp add: i_def obt(1))
+      then have "x + (y - ?a) \<in> cball x (b x)"
         using y unfolding mem_cball dist_norm by auto
-      moreover from \<open>x\<in>t\<close> have "x \<in> s"
-        using obt(2) by auto
-      ultimately have "x + (y - a) \<in> s"
-        using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast
+      moreover have "x \<in> S"
+        using \<open>x\<in>T\<close> \<open>T\<subseteq>S\<close> by auto
+      ultimately have "x + (y - ?a) \<in> S"
+        using y b by blast
     }
     moreover
-    have *: "inj_on (\<lambda>v. v + (y - a)) t"
+    have *: "inj_on (\<lambda>v. v + (y - ?a)) T"
       unfolding inj_on_def by auto
-    have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1"
-      unfolding sum.reindex[OF *] o_def using obt(4) by auto
-    moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y"
-      unfolding sum.reindex[OF *] o_def using obt(4,5)
+    have "(\<Sum>v\<in>(\<lambda>v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y"
+      unfolding sum.reindex[OF *] o_def using obt(4)
       by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib)
-    ultimately
-    show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
-      apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI)
-      apply (rule_tac x="\<lambda>v. u (v - (y - a))" in exI)
-      using obt(1, 3)
-      apply auto
-      done
+    ultimately show "y \<in> {y. ?\<Phi> y}"
+    proof (intro CollectI exI conjI)
+      show "finite ((\<lambda>v. v + (y - ?a)) ` T)"
+        by (simp add: obt(1))
+      show "sum (\<lambda>v. u (v - (y - ?a))) ((\<lambda>v. v + (y - ?a)) ` T) = 1"
+        unfolding sum.reindex[OF *] o_def using obt(4) by auto
+    qed (use obt(1, 3) in auto)
   qed
 qed
 
 lemma compact_convex_combinations:
-  fixes s t :: "'a::real_normed_vector set"
-  assumes "compact s" "compact t"
-  shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> t}"
+  fixes S T :: "'a::real_normed_vector set"
+  assumes "compact S" "compact T"
+  shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> T}"
 proof -
-  let ?X = "{0..1} \<times> s \<times> t"
+  let ?X = "{0..1} \<times> S \<times> T"
   let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
-  have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> t} = ?h ` ?X"
-    apply (rule set_eqI)
-    unfolding image_iff mem_Collect_eq
-    apply rule
-    apply auto
-    apply (rule_tac x=u in rev_bexI)
-    apply simp
-    apply (erule rev_bexI)
-    apply (erule rev_bexI)
-    apply simp
-    apply auto
-    done
+  have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> T} = ?h ` ?X"
+    by force
   have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
     unfolding continuous_on by (rule ballI) (intro tendsto_intros)
-  then show ?thesis
-    unfolding *
-    apply (rule compact_continuous_image)
-    apply (intro compact_Times compact_Icc assms)
-    done
+  with assms show ?thesis
+    by (simp add: * compact_Times compact_continuous_image)
 qed
 
 lemma finite_imp_compact_convex_hull:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "finite s"
-  shows "compact (convex hull s)"
-proof (cases "s = {}")
+  fixes S :: "'a::real_normed_vector set"
+  assumes "finite S"
+  shows "compact (convex hull S)"
+proof (cases "S = {}")
   case True
   then show ?thesis by simp
 next
@@ -5142,8 +4838,7 @@
       unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>]
       apply safe
       apply (rule_tac x=a in exI, simp)
-      apply (rule_tac x="1 - a" in exI, simp)
-      apply fast
+      apply (rule_tac x="1 - a" in exI, simp, fast)
       apply (rule_tac x="(u, b)" in image_eqI, simp_all)
       done
     finally show "compact (convex hull (insert x A))" .
@@ -5151,20 +4846,20 @@
 qed
 
 lemma compact_convex_hull:
-  fixes s :: "'a::euclidean_space set"
-  assumes "compact s"
-  shows "compact (convex hull s)"
-proof (cases "s = {}")
+  fixes S :: "'a::euclidean_space set"
+  assumes "compact S"
+  shows "compact (convex hull S)"
+proof (cases "S = {}")
   case True
   then show ?thesis using compact_empty by simp
 next
   case False
-  then obtain w where "w \<in> s" by auto
+  then obtain w where "w \<in> S" by auto
   show ?thesis
-    unfolding caratheodory[of s]
+    unfolding caratheodory[of S]
   proof (induct ("DIM('a) + 1"))
     case 0
-    have *: "{x.\<exists>sa. finite sa \<and> sa \<subseteq> s \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}"
+    have *: "{x.\<exists>sa. finite sa \<and> sa \<subseteq> S \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}"
       using compact_empty by auto
     from 0 show ?case unfolding * by simp
   next
@@ -5172,27 +4867,27 @@
     show ?case
     proof (cases "n = 0")
       case True
-      have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
+      have "{x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T} = S"
         unfolding set_eq_iff and mem_Collect_eq
       proof (rule, rule)
         fix x
-        assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-        then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t"
+        assume "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T"
+        then obtain T where T: "finite T" "T \<subseteq> S" "card T \<le> Suc n" "x \<in> convex hull T"
           by auto
-        show "x \<in> s"
-        proof (cases "card t = 0")
+        show "x \<in> S"
+        proof (cases "card T = 0")
           case True
           then show ?thesis
-            using t(4) unfolding card_0_eq[OF t(1)] by simp
+            using T(4) unfolding card_0_eq[OF T(1)] by simp
         next
           case False
-          then have "card t = Suc 0" using t(3) \<open>n=0\<close> by auto
-          then obtain a where "t = {a}" unfolding card_Suc_eq by auto
-          then show ?thesis using t(2,4) by simp
+          then have "card T = Suc 0" using T(3) \<open>n=0\<close> by auto
+          then obtain a where "T = {a}" unfolding card_Suc_eq by auto
+          then show ?thesis using T(2,4) by simp
         qed
       next
-        fix x assume "x\<in>s"
-        then show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+        fix x assume "x\<in>S"
+        then show "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T"
           apply (rule_tac x="{x}" in exI)
           unfolding convex_hull_singleton
           apply auto
@@ -5201,57 +4896,56 @@
       then show ?thesis using assms by simp
     next
       case False
-      have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
+      have "{x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T} =
         {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
-          0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> {x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> x \<in> convex hull t}}"
+          0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> {x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> x \<in> convex hull T}}"
         unfolding set_eq_iff and mem_Collect_eq
       proof (rule, rule)
         fix x
         assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
-          0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)"
-        then obtain u v c t where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
-          "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n"  "v \<in> convex hull t"
+          0 \<le> c \<and> c \<le> 1 \<and> u \<in> S \<and> (\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> v \<in> convex hull T)"
+        then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
+          "0 \<le> c \<and> c \<le> 1" "u \<in> S" "finite T" "T \<subseteq> S" "card T \<le> n"  "v \<in> convex hull T"
           by auto
-        moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u t"
+        moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u T"
           apply (rule convexD_alt)
-          using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex]
-          using obt(7) and hull_mono[of t "insert u t"]
+          using obt(2) and convex_convex_hull and hull_subset[of "insert u T" convex]
+          using obt(7) and hull_mono[of T "insert u T"]
           apply auto
           done
-        ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-          apply (rule_tac x="insert u t" in exI)
-          apply (auto simp add: card_insert_if)
+        ultimately show "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T"
+          apply (rule_tac x="insert u T" in exI)
+          apply (auto simp: card_insert_if)
           done
       next
         fix x
-        assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-        then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t"
+        assume "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T"
+        then obtain T where T: "finite T" "T \<subseteq> S" "card T \<le> Suc n" "x \<in> convex hull T"
           by auto
         show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
-          0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)"
-        proof (cases "card t = Suc n")
+          0 \<le> c \<and> c \<le> 1 \<and> u \<in> S \<and> (\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> v \<in> convex hull T)"
+        proof (cases "card T = Suc n")
           case False
-          then have "card t \<le> n" using t(3) by auto
+          then have "card T \<le> n" using T(3) by auto
           then show ?thesis
             apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI)
-            using \<open>w\<in>s\<close> and t
-            apply (auto intro!: exI[where x=t])
+            using \<open>w\<in>S\<close> and T
+            apply (auto intro!: exI[where x=T])
             done
         next
           case True
-          then obtain a u where au: "t = insert a u" "a\<notin>u"
-            apply (drule_tac card_eq_SucD)
-            apply auto
+          then obtain a u where au: "T = insert a u" "a\<notin>u"
+            apply (drule_tac card_eq_SucD, auto)
             done
           show ?thesis
           proof (cases "u = {}")
             case True
-            then have "x = a" using t(4)[unfolded au] by auto
+            then have "x = a" using T(4)[unfolded au] by auto
             show ?thesis unfolding \<open>x = a\<close>
               apply (rule_tac x=a in exI)
               apply (rule_tac x=a in exI)
               apply (rule_tac x=1 in exI)
-              using t and \<open>n \<noteq> 0\<close>
+              using T and \<open>n \<noteq> 0\<close>
               unfolding au
               apply (auto intro!: exI[where x="{a}"])
               done
@@ -5259,14 +4953,14 @@
             case False
             obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1"
               "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
-              using t(4)[unfolded au convex_hull_insert[OF False]]
+              using T(4)[unfolded au convex_hull_insert[OF False]]
               by auto
             have *: "1 - vx = ux" using obt(3) by auto
             show ?thesis
               apply (rule_tac x=a in exI)
               apply (rule_tac x=b in exI)
               apply (rule_tac x=vx in exI)
-              using obt and t(1-3)
+              using obt and T(1-3)
               unfolding au and * using card_insert_disjoint[OF _ au(2)]
               apply (auto intro!: exI[where x=u])
               done
@@ -5318,25 +5012,25 @@
   using dist_increases_online[of d a 0] unfolding dist_norm by auto
 
 lemma simplex_furthest_lt:
-  fixes s :: "'a::real_inner set"
-  assumes "finite s"
-  shows "\<forall>x \<in> convex hull s.  x \<notin> s \<longrightarrow> (\<exists>y \<in> convex hull s. norm (x - a) < norm(y - a))"
+  fixes S :: "'a::real_inner set"
+  assumes "finite S"
+  shows "\<forall>x \<in> convex hull S.  x \<notin> S \<longrightarrow> (\<exists>y \<in> convex hull S. norm (x - a) < norm(y - a))"
   using assms
 proof induct
-  fix x s
-  assume as: "finite s" "x\<notin>s" "\<forall>x\<in>convex hull s. x \<notin> s \<longrightarrow> (\<exists>y\<in>convex hull s. norm (x - a) < norm (y - a))"
-  show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow>
-    (\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))"
-  proof (rule, rule, cases "s = {}")
+  fix x S
+  assume as: "finite S" "x\<notin>S" "\<forall>x\<in>convex hull S. x \<notin> S \<longrightarrow> (\<exists>y\<in>convex hull S. norm (x - a) < norm (y - a))"
+  show "\<forall>xa\<in>convex hull insert x S. xa \<notin> insert x S \<longrightarrow>
+    (\<exists>y\<in>convex hull insert x S. norm (xa - a) < norm (y - a))"
+  proof (intro impI ballI, cases "S = {}")
     case False
     fix y
-    assume y: "y \<in> convex hull insert x s" "y \<notin> insert x s"
-    obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "y = u *\<^sub>R x + v *\<^sub>R b"
+    assume y: "y \<in> convex hull insert x S" "y \<notin> insert x S"
+    obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b"
       using y(1)[unfolded convex_hull_insert[OF False]] by auto
-    show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)"
-    proof (cases "y \<in> convex hull s")
+    show "\<exists>z\<in>convex hull insert x S. norm (y - a) < norm (z - a)"
+    proof (cases "y \<in> convex hull S")
       case True
-      then obtain z where "z \<in> convex hull s" "norm (y - a) < norm (z - a)"
+      then obtain z where "z \<in> convex hull S" "norm (y - a) < norm (z - a)"
         using as(3)[THEN bspec[where x=y]] and y(2) by auto
       then show ?thesis
         apply (rule_tac x=z in bexI)
@@ -5363,7 +5057,7 @@
         proof
           assume "x = b"
           then have "y = b" unfolding obt(5)
-            using obt(3) by (auto simp add: scaleR_left_distrib[symmetric])
+            using obt(3) by (auto simp: scaleR_left_distrib[symmetric])
           then show False using obt(4) and False by simp
         qed
         then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
@@ -5375,15 +5069,12 @@
             unfolding dist_commute[of a]
             unfolding dist_norm obt(5)
             by (simp add: algebra_simps)
-          moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
-            unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
-            apply (rule_tac x="u + w" in exI)
-            apply rule
-            defer
-            apply (rule_tac x="v - w" in exI)
-            using \<open>u \<ge> 0\<close> and w and obt(3,4)
-            apply auto
-            done
+          moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x S"
+            unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>]
+          proof (intro CollectI conjI exI)
+            show "u + w \<ge> 0" "v - w \<ge> 0"
+              using obt(1) w by auto
+          qed (use obt in auto)
           ultimately show ?thesis by auto
         next
           assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
@@ -5391,39 +5082,36 @@
             unfolding dist_commute[of a]
             unfolding dist_norm obt(5)
             by (simp add: algebra_simps)
-          moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
-            unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
-            apply (rule_tac x="u - w" in exI)
-            apply rule
-            defer
-            apply (rule_tac x="v + w" in exI)
-            using \<open>u \<ge> 0\<close> and w and obt(3,4)
-            apply auto
-            done
+          moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x S"
+            unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>]
+          proof (intro CollectI conjI exI)
+            show "u - w \<ge> 0" "v + w \<ge> 0"
+              using obt(1) w by auto
+          qed (use obt in auto)
           ultimately show ?thesis by auto
         qed
       qed auto
     qed
   qed auto
-qed (auto simp add: assms)
+qed (auto simp: assms)
 
 lemma simplex_furthest_le:
-  fixes s :: "'a::real_inner set"
-  assumes "finite s"
-    and "s \<noteq> {}"
-  shows "\<exists>y\<in>s. \<forall>x\<in> convex hull s. norm (x - a) \<le> norm (y - a)"
+  fixes S :: "'a::real_inner set"
+  assumes "finite S"
+    and "S \<noteq> {}"
+  shows "\<exists>y\<in>S. \<forall>x\<in> convex hull S. norm (x - a) \<le> norm (y - a)"
 proof -
-  have "convex hull s \<noteq> {}"
-    using hull_subset[of s convex] and assms(2) by auto
-  then obtain x where x: "x \<in> convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)"
-    using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a]
+  have "convex hull S \<noteq> {}"
+    using hull_subset[of S convex] and assms(2) by auto
+  then obtain x where x: "x \<in> convex hull S" "\<forall>y\<in>convex hull S. norm (y - a) \<le> norm (x - a)"
+    using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \<open>finite S\<close>], of a]
     unfolding dist_commute[of a]
     unfolding dist_norm
     by auto
   show ?thesis
-  proof (cases "x \<in> s")
+  proof (cases "x \<in> S")
     case False
-    then obtain y where "y \<in> convex hull s" "norm (x - a) < norm (y - a)"
+    then obtain y where "y \<in> convex hull S" "norm (x - a) < norm (y - a)"
       using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1)
       by auto
     then show ?thesis
@@ -5435,82 +5123,82 @@
 qed
 
 lemma simplex_furthest_le_exists:
-  fixes s :: "('a::real_inner) set"
-  shows "finite s \<Longrightarrow> \<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm (x - a) \<le> norm (y - a)"
-  using simplex_furthest_le[of s] by (cases "s = {}") auto
+  fixes S :: "('a::real_inner) set"
+  shows "finite S \<Longrightarrow> \<forall>x\<in>(convex hull S). \<exists>y\<in>S. norm (x - a) \<le> norm (y - a)"
+  using simplex_furthest_le[of S] by (cases "S = {}") auto
 
 lemma simplex_extremal_le:
-  fixes s :: "'a::real_inner set"
-  assumes "finite s"
-    and "s \<noteq> {}"
-  shows "\<exists>u\<in>s. \<exists>v\<in>s. \<forall>x\<in>convex hull s. \<forall>y \<in> convex hull s. norm (x - y) \<le> norm (u - v)"
+  fixes S :: "'a::real_inner set"
+  assumes "finite S"
+    and "S \<noteq> {}"
+  shows "\<exists>u\<in>S. \<exists>v\<in>S. \<forall>x\<in>convex hull S. \<forall>y \<in> convex hull S. norm (x - y) \<le> norm (u - v)"
 proof -
-  have "convex hull s \<noteq> {}"
-    using hull_subset[of s convex] and assms(2) by auto
-  then obtain u v where obt: "u \<in> convex hull s" "v \<in> convex hull s"
-    "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull s. norm (x - y) \<le> norm (u - v)"
+  have "convex hull S \<noteq> {}"
+    using hull_subset[of S convex] and assms(2) by auto
+  then obtain u v where obt: "u \<in> convex hull S" "v \<in> convex hull S"
+    "\<forall>x\<in>convex hull S. \<forall>y\<in>convex hull S. norm (x - y) \<le> norm (u - v)"
     using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]]
     by (auto simp: dist_norm)
   then show ?thesis
-  proof (cases "u\<notin>s \<or> v\<notin>s", elim disjE)
-    assume "u \<notin> s"
-    then obtain y where "y \<in> convex hull s" "norm (u - v) < norm (y - v)"
+  proof (cases "u\<notin>S \<or> v\<notin>S", elim disjE)
+    assume "u \<notin> S"
+    then obtain y where "y \<in> convex hull S" "norm (u - v) < norm (y - v)"
       using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1)
       by auto
     then show ?thesis
       using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2)
       by auto
   next
-    assume "v \<notin> s"
-    then obtain y where "y \<in> convex hull s" "norm (v - u) < norm (y - u)"
+    assume "v \<notin> S"
+    then obtain y where "y \<in> convex hull S" "norm (v - u) < norm (y - u)"
       using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2)
       by auto
     then show ?thesis
       using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
-      by (auto simp add: norm_minus_commute)
+      by (auto simp: norm_minus_commute)
   qed auto
 qed
 
 lemma simplex_extremal_le_exists:
-  fixes s :: "'a::real_inner set"
-  shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s \<Longrightarrow>
-    \<exists>u\<in>s. \<exists>v\<in>s. norm (x - y) \<le> norm (u - v)"
-  using convex_hull_empty simplex_extremal_le[of s]
-  by(cases "s = {}") auto
+  fixes S :: "'a::real_inner set"
+  shows "finite S \<Longrightarrow> x \<in> convex hull S \<Longrightarrow> y \<in> convex hull S \<Longrightarrow>
+    \<exists>u\<in>S. \<exists>v\<in>S. norm (x - y) \<le> norm (u - v)"
+  using convex_hull_empty simplex_extremal_le[of S]
+  by(cases "S = {}") auto
 
 
 subsection \<open>Closest point of a convex set is unique, with a continuous projection\<close>
 
 definition%important closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
-  where "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))"
+  where "closest_point S a = (SOME x. x \<in> S \<and> (\<forall>y\<in>S. dist a x \<le> dist a y))"
 
 lemma closest_point_exists:
-  assumes "closed s"
-    and "s \<noteq> {}"
-  shows "closest_point s a \<in> s"
-    and "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y"
+  assumes "closed S"
+    and "S \<noteq> {}"
+  shows "closest_point S a \<in> S"
+    and "\<forall>y\<in>S. dist a (closest_point S a) \<le> dist a y"
   unfolding closest_point_def
   apply(rule_tac[!] someI2_ex)
   apply (auto intro: distance_attains_inf[OF assms(1,2), of a])
   done
 
-lemma closest_point_in_set: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s a \<in> s"
+lemma closest_point_in_set: "closed S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> closest_point S a \<in> S"
   by (meson closest_point_exists)
 
-lemma closest_point_le: "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x"
-  using closest_point_exists[of s] by auto
+lemma closest_point_le: "closed S \<Longrightarrow> x \<in> S \<Longrightarrow> dist a (closest_point S a) \<le> dist a x"
+  using closest_point_exists[of S] by auto
 
 lemma closest_point_self:
-  assumes "x \<in> s"
-  shows "closest_point s x = x"
+  assumes "x \<in> S"
+  shows "closest_point S x = x"
   unfolding closest_point_def
   apply (rule some1_equality, rule ex1I[of _ x])
   using assms
   apply auto
   done
 
-lemma closest_point_refl: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s x = x \<longleftrightarrow> x \<in> s"
-  using closest_point_in_set[of s x] closest_point_self[of x s]
+lemma closest_point_refl: "closed S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> closest_point S x = x \<longleftrightarrow> x \<in> S"
+  using closest_point_in_set[of S x] closest_point_self[of x S]
   by auto
 
 lemma closer_points_lemma:
@@ -5519,18 +5207,13 @@
 proof -
   have z: "inner z z > 0"
     unfolding inner_gt_zero_iff using assms by auto
+  have "norm (v *\<^sub>R z - y) < norm y"
+    if "0 < v" and "v \<le> inner y z / inner z z" for v
+    unfolding norm_lt using z assms that
+    by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>])
   then show ?thesis
-    using assms
-    apply (rule_tac x = "inner y z / inner z z" in exI)
-    apply rule
-    defer
-  proof rule+
-    fix v
-    assume "0 < v" and "v \<le> inner y z / inner z z"
-    then show "norm (v *\<^sub>R z - y) < norm y"
-      unfolding norm_lt using z and assms
-      by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>])
-  qed auto
+    using assms z
+    by (rule_tac x = "inner y z / inner z z" in exI) auto
 qed
 
 lemma closer_point_lemma:
@@ -5543,50 +5226,50 @@
   show ?thesis
     apply (rule_tac x="min u 1" in exI)
     using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close>
-    unfolding dist_norm by (auto simp add: norm_minus_commute field_simps)
+    unfolding dist_norm by (auto simp: norm_minus_commute field_simps)
 qed
 
 lemma any_closest_point_dot:
-  assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
+  assumes "convex S" "closed S" "x \<in> S" "y \<in> S" "\<forall>z\<in>S. dist a x \<le> dist a z"
   shows "inner (a - x) (y - x) \<le> 0"
 proof (rule ccontr)
   assume "\<not> ?thesis"
   then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a"
     using closer_point_lemma[of a x y] by auto
   let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y"
-  have "?z \<in> s"
+  have "?z \<in> S"
     using convexD_alt[OF assms(1,3,4), of u] using u by auto
   then show False
     using assms(5)[THEN bspec[where x="?z"]] and u(3)
-    by (auto simp add: dist_commute algebra_simps)
+    by (auto simp: dist_commute algebra_simps)
 qed
 
 lemma any_closest_point_unique:
   fixes x :: "'a::real_inner"
-  assumes "convex s" "closed s" "x \<in> s" "y \<in> s"
-    "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z"
+  assumes "convex S" "closed S" "x \<in> S" "y \<in> S"
+    "\<forall>z\<in>S. dist a x \<le> dist a z" "\<forall>z\<in>S. dist a y \<le> dist a z"
   shows "x = y"
   using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)]
   unfolding norm_pths(1) and norm_le_square
-  by (auto simp add: algebra_simps)
+  by (auto simp: algebra_simps)
 
 lemma closest_point_unique:
-  assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
-  shows "x = closest_point s a"
-  using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"]
+  assumes "convex S" "closed S" "x \<in> S" "\<forall>z\<in>S. dist a x \<le> dist a z"
+  shows "x = closest_point S a"
+  using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"]
   using closest_point_exists[OF assms(2)] and assms(3) by auto
 
 lemma closest_point_dot:
-  assumes "convex s" "closed s" "x \<in> s"
-  shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0"
+  assumes "convex S" "closed S" "x \<in> S"
+  shows "inner (a - closest_point S a) (x - closest_point S a) \<le> 0"
   apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)])
   using closest_point_exists[OF assms(2)] and assms(3)
   apply auto
   done
 
 lemma closest_point_lt:
-  assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a"
-  shows "dist a (closest_point s a) < dist a x"
+  assumes "convex S" "closed S" "x \<in> S" "x \<noteq> closest_point S a"
+  shows "dist a (closest_point S a) < dist a x"
   apply (rule ccontr)
   apply (rule_tac notE[OF assms(4)])
   apply (rule closest_point_unique[OF assms(1-3), of a])
@@ -5595,34 +5278,34 @@
   done
 
 lemma closest_point_lipschitz:
-  assumes "convex s"
-    and "closed s" "s \<noteq> {}"
-  shows "dist (closest_point s x) (closest_point s y) \<le> dist x y"
+  assumes "convex S"
+    and "closed S" "S \<noteq> {}"
+  shows "dist (closest_point S x) (closest_point S y) \<le> dist x y"
 proof -
-  have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0"
-    and "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0"
+  have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \<le> 0"
+    and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \<le> 0"
     apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)])
     using closest_point_exists[OF assms(2-3)]
     apply auto
     done
   then show ?thesis unfolding dist_norm and norm_le
-    using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"]
+    using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"]
     by (simp add: inner_add inner_diff inner_commute)
 qed
 
 lemma continuous_at_closest_point:
-  assumes "convex s"
-    and "closed s"
-    and "s \<noteq> {}"
-  shows "continuous (at x) (closest_point s)"
+  assumes "convex S"
+    and "closed S"
+    and "S \<noteq> {}"
+  shows "continuous (at x) (closest_point S)"
   unfolding continuous_at_eps_delta
   using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto
 
 lemma continuous_on_closest_point:
-  assumes "convex s"
-    and "closed s"
-    and "s \<noteq> {}"
-  shows "continuous_on t (closest_point s)"
+  assumes "convex S"
+    and "closed S"
+    and "S \<noteq> {}"
+  shows "continuous_on t (closest_point S)"
   by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms])
 
 proposition closest_point_in_rel_interior:
@@ -5647,7 +5330,7 @@
       by (simp add: y_def algebra_simps)
     then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)"
       by simp
-    also have "... < norm(x - closest_point S x)"
+    also have "\<dots> < norm(x - closest_point S x)"
       using clo_notx \<open>e > 0\<close>
       by (auto simp: mult_less_cancel_right2 divide_simps)
     finally have no_less: "norm (x - y) < norm (x - closest_point S x)" .
@@ -5673,121 +5356,84 @@
 
 lemma supporting_hyperplane_closed_point:
   fixes z :: "'a::{real_inner,heine_borel}"
-  assumes "convex s"
-    and "closed s"
-    and "s \<noteq> {}"
-    and "z \<notin> s"
-  shows "\<exists>a b. \<exists>y\<in>s. inner a z < b \<and> inner a y = b \<and> (\<forall>x\<in>s. inner a x \<ge> b)"
+  assumes "convex S"
+    and "closed S"
+    and "S \<noteq> {}"
+    and "z \<notin> S"
+  shows "\<exists>a b. \<exists>y\<in>S. inner a z < b \<and> inner a y = b \<and> (\<forall>x\<in>S. inner a x \<ge> b)"
 proof -
-  obtain y where "y \<in> s" and y: "\<forall>x\<in>s. dist z y \<le> dist z x"
+  obtain y where "y \<in> S" and y: "\<forall>x\<in>S. dist z y \<le> dist z x"
     by (metis distance_attains_inf[OF assms(2-3)])
   show ?thesis
-    apply (rule_tac x="y - z" in exI)
-    apply (rule_tac x="inner (y - z) y" in exI)
-    apply (rule_tac x=y in bexI)
-    apply rule
-    defer
-    apply rule
-    defer
-    apply rule
-    apply (rule ccontr)
-    using \<open>y \<in> s\<close>
-  proof -
-    show "inner (y - z) z < inner (y - z) y"
-      apply (subst diff_gt_0_iff_gt [symmetric])
-      unfolding inner_diff_right[symmetric] and inner_gt_zero_iff
-      using \<open>y\<in>s\<close> \<open>z\<notin>s\<close>
-      apply auto
-      done
-  next
-    fix x
-    assume "x \<in> s"
-    have *: "\<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
-      using assms(1)[unfolded convex_alt] and y and \<open>x\<in>s\<close> and \<open>y\<in>s\<close> by auto
-    assume "\<not> inner (y - z) y \<le> inner (y - z) x"
-    then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z"
-      using closer_point_lemma[of z y x] by (auto simp add: inner_diff)
-    then show False
-      using *[THEN spec[where x=v]] by (auto simp add: dist_commute algebra_simps)
-  qed auto
+  proof (intro exI bexI conjI ballI)
+    show "(y - z) \<bullet> z < (y - z) \<bullet> y"
+      by (metis \<open>y \<in> S\<close> assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq)
+    show "(y - z) \<bullet> y \<le> (y - z) \<bullet> x" if "x \<in> S" for x
+    proof (rule ccontr)
+      have *: "\<And>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
+        using assms(1)[unfolded convex_alt] and y and \<open>x\<in>S\<close> and \<open>y\<in>S\<close> by auto
+      assume "\<not> (y - z) \<bullet> y \<le> (y - z) \<bullet> x"
+      then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z"
+        using closer_point_lemma[of z y x] by (auto simp: inner_diff)
+      then show False
+        using *[of v] by (auto simp: dist_commute algebra_simps)
+    qed
+  qed (use \<open>y \<in> S\<close> in auto)
 qed
 
 lemma separating_hyperplane_closed_point:
   fixes z :: "'a::{real_inner,heine_borel}"
-  assumes "convex s"
-    and "closed s"
-    and "z \<notin> s"
-  shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)"
-proof (cases "s = {}")
+  assumes "convex S"
+    and "closed S"
+    and "z \<notin> S"
+  shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>S. inner a x > b)"
+proof (cases "S = {}")
   case True
   then show ?thesis
-    apply (rule_tac x="-z" in exI)
-    apply (rule_tac x=1 in exI)
-    using less_le_trans[OF _ inner_ge_zero[of z]]
-    apply auto
-    done
+    by (simp add: gt_ex)
 next
   case False
-  obtain y where "y \<in> s" and y: "\<forall>x\<in>s. dist z y \<le> dist z x"
+  obtain y where "y \<in> S" and y: "\<And>x. x \<in> S \<Longrightarrow> dist z y \<le> dist z x"
     by (metis distance_attains_inf[OF assms(2) False])
   show ?thesis
-    apply (rule_tac x="y - z" in exI)
-    apply (rule_tac x="inner (y - z) z + (norm (y - z))\<^sup>2 / 2" in exI)
-    apply rule
-    defer
-    apply rule
-  proof -
+  proof (intro exI conjI ballI)
+    show "(y - z) \<bullet> z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2"
+      using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto
+  next
     fix x
-    assume "x \<in> s"
-    have "\<not> 0 < inner (z - y) (x - y)"
-      apply (rule notI)
-      apply (drule closer_point_lemma)
+    assume "x \<in> S"
+    have "False" if *: "0 < inner (z - y) (x - y)"
     proof -
-      assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z"
-      then obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z"
-        by auto
-      then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]]
-        using assms(1)[unfolded convex_alt, THEN bspec[where x=y]]
-        using \<open>x\<in>s\<close> \<open>y\<in>s\<close> by (auto simp add: dist_commute algebra_simps)
+      obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z"
+        using * closer_point_lemma by blast
+      then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \<open>convex S\<close>]
+        using \<open>x\<in>S\<close> \<open>y\<in>S\<close> by (auto simp: dist_commute algebra_simps)
     qed
     moreover have "0 < (norm (y - z))\<^sup>2"
-      using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> by auto
+      using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto
     then have "0 < inner (y - z) (y - z)"
       unfolding power2_norm_eq_inner by simp
-    ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x"
-      unfolding power2_norm_eq_inner and not_less
-      by (auto simp add: field_simps inner_commute inner_diff)
-  qed (insert \<open>y\<in>s\<close> \<open>z\<notin>s\<close>, auto)
+    ultimately show "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> x"
+      by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff)
+  qed 
 qed
 
 lemma separating_hyperplane_closed_0:
-  assumes "convex (s::('a::euclidean_space) set)"
-    and "closed s"
-    and "0 \<notin> s"
-  shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)"
-proof (cases "s = {}")
+  assumes "convex (S::('a::euclidean_space) set)"
+    and "closed S"
+    and "0 \<notin> S"
+  shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>S. inner a x > b)"
+proof (cases "S = {}")
   case True
-  have "norm ((SOME i. i\<in>Basis)::'a) = 1" "(SOME i. i\<in>Basis) \<noteq> (0::'a)"
-    defer
-    apply (subst norm_le_zero_iff[symmetric])
-    apply (auto simp: SOME_Basis)
-    done
+  have "(SOME i. i\<in>Basis) \<noteq> (0::'a)"
+    by (metis Basis_zero SOME_Basis)
   then show ?thesis
-    apply (rule_tac x="SOME i. i\<in>Basis" in exI)
-    apply (rule_tac x=1 in exI)
-    using True using DIM_positive[where 'a='a]
-    apply auto
-    done
+    using True zero_less_one by blast
 next
   case False
   then show ?thesis
     using False using separating_hyperplane_closed_point[OF assms]
-    apply (elim exE)
-    unfolding inner_zero_right
-    apply (rule_tac x=a in exI)
-    apply (rule_tac x=b in exI)
-    apply auto
-    done
+    by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le)
 qed
 
 
@@ -5826,7 +5472,7 @@
     apply rule
     apply rule
     apply (erule_tac x="x - y" in ballE)
-    apply (auto simp add: inner_diff)
+    apply (auto simp: inner_diff)
     done
   define k where "k = (SUP x:T. a \<bullet> x)"
   show ?thesis
@@ -5876,8 +5522,7 @@
     by auto
   then show ?thesis
     apply (rule_tac x="-a" in exI)
-    apply (rule_tac x="-b" in exI)
-    apply auto
+    apply (rule_tac x="-b" in exI, auto)
     done
 qed
 
@@ -5885,13 +5530,13 @@
 subsubsection%unimportant \<open>General case without assuming closure and getting non-strict separation\<close>
 
 lemma separating_hyperplane_set_0:
-  assumes "convex s" "(0::'a::euclidean_space) \<notin> s"
-  shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>s. 0 \<le> inner a x)"
+  assumes "convex S" "(0::'a::euclidean_space) \<notin> S"
+  shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>S. 0 \<le> inner a x)"
 proof -
   let ?k = "\<lambda>c. {x::'a. 0 \<le> inner c x}"
-  have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` s" "finite f" for f
+  have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` S" "finite f" for f
   proof -
-    obtain c where c: "f = ?k ` c" "c \<subseteq> s" "finite c"
+    obtain c where c: "f = ?k ` c" "c \<subseteq> S" "finite c"
       using finite_subset_image[OF as(2,1)] by auto
     then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x"
       using separating_hyperplane_closed_0[OF convex_convex_hull, of c]
@@ -5902,50 +5547,50 @@
       apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI)
       using hull_subset[of c convex]
       unfolding subset_eq and inner_scaleR
-      by (auto simp add: inner_commute del: ballE elim!: ballE)
+      by (auto simp: inner_commute del: ballE elim!: ballE)
     then show "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}"
       unfolding c(1) frontier_cball sphere_def dist_norm by auto
   qed
-  have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` s)) \<noteq> {}"
+  have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` S)) \<noteq> {}"
     apply (rule compact_imp_fip)
     apply (rule compact_frontier[OF compact_cball])
     using * closed_halfspace_ge
     by auto
-  then obtain x where "norm x = 1" "\<forall>y\<in>s. x\<in>?k y"
+  then obtain x where "norm x = 1" "\<forall>y\<in>S. x\<in>?k y"
     unfolding frontier_cball dist_norm sphere_def by auto
   then show ?thesis
     by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one)
 qed
 
 lemma separating_hyperplane_sets:
-  fixes s t :: "'a::euclidean_space set"
-  assumes "convex s"
-    and "convex t"
-    and "s \<noteq> {}"
-    and "t \<noteq> {}"
-    and "s \<inter> t = {}"
-  shows "\<exists>a b. a \<noteq> 0 \<and> (\<forall>x\<in>s. inner a x \<le> b) \<and> (\<forall>x\<in>t. inner a x \<ge> b)"
+  fixes S T :: "'a::euclidean_space set"
+  assumes "convex S"
+    and "convex T"
+    and "S \<noteq> {}"
+    and "T \<noteq> {}"
+    and "S \<inter> T = {}"
+  shows "\<exists>a b. a \<noteq> 0 \<and> (\<forall>x\<in>S. inner a x \<le> b) \<and> (\<forall>x\<in>T. inner a x \<ge> b)"
 proof -
   from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
-  obtain a where "a \<noteq> 0" "\<forall>x\<in>{x - y |x y. x \<in> t \<and> y \<in> s}. 0 \<le> inner a x"
+  obtain a where "a \<noteq> 0" "\<forall>x\<in>{x - y |x y. x \<in> T \<and> y \<in> S}. 0 \<le> inner a x"
     using assms(3-5) by force
-  then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x"
-    by (force simp add: inner_diff)
-  then have bdd: "bdd_above (((\<bullet>) a)`s)"
-    using \<open>t \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *])
+  then have *: "\<And>x y. x \<in> T \<Longrightarrow> y \<in> S \<Longrightarrow> inner a y \<le> inner a x"
+    by (force simp: inner_diff)
+  then have bdd: "bdd_above (((\<bullet>) a)`S)"
+    using \<open>T \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *])
   show ?thesis
     using \<open>a\<noteq>0\<close>
-    by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"])
-       (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>s \<noteq> {}\<close> *)
+    by (intro exI[of _ a] exI[of _ "SUP x:S. a \<bullet> x"])
+       (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>S \<noteq> {}\<close> *)
 qed
 
 
 subsection%unimportant \<open>More convexity generalities\<close>
 
 lemma convex_closure [intro,simp]:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "convex s"
-  shows "convex (closure s)"
+  fixes S :: "'a::real_normed_vector set"
+  assumes "convex S"
+  shows "convex (closure S)"
   apply (rule convexI)
   apply (unfold closure_sequential, elim exE)
   apply (rule_tac x="\<lambda>n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI)
@@ -5955,65 +5600,58 @@
   done
 
 lemma convex_interior [intro,simp]:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "convex s"
-  shows "convex (interior s)"
+  fixes S :: "'a::real_normed_vector set"
+  assumes "convex S"
+  shows "convex (interior S)"
   unfolding convex_alt Ball_def mem_interior
-  apply (rule,rule,rule,rule,rule,rule)
-  apply (elim exE conjE)
-proof -
+proof clarify
   fix x y u
   assume u: "0 \<le> u" "u \<le> (1::real)"
   fix e d
-  assume ed: "ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e"
-  show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s"
-    apply (rule_tac x="min d e" in exI)
-    apply rule
-    unfolding subset_eq
-    defer
-    apply rule
-  proof -
+  assume ed: "ball x e \<subseteq> S" "ball y d \<subseteq> S" "0<d" "0<e"
+  show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> S"
+  proof (intro exI conjI subsetI)
     fix z
     assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
-    then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s"
+    then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> S"
       apply (rule_tac assms[unfolded convex_alt, rule_format])
       using ed(1,2) and u
       unfolding subset_eq mem_ball Ball_def dist_norm
-      apply (auto simp add: algebra_simps)
+      apply (auto simp: algebra_simps)
       done
-    then show "z \<in> s"
-      using u by (auto simp add: algebra_simps)
+    then show "z \<in> S"
+      using u by (auto simp: algebra_simps)
   qed(insert u ed(3-4), auto)
 qed
 
-lemma convex_hull_eq_empty[simp]: "convex hull s = {} \<longleftrightarrow> s = {}"
-  using hull_subset[of s convex] convex_hull_empty by auto
+lemma convex_hull_eq_empty[simp]: "convex hull S = {} \<longleftrightarrow> S = {}"
+  using hull_subset[of S convex] convex_hull_empty by auto
 
 
 subsection%unimportant \<open>Moving and scaling convex hulls\<close>
 
 lemma convex_hull_set_plus:
-  "convex hull (s + t) = convex hull s + convex hull t"
+  "convex hull (S + T) = convex hull S + convex hull T"
   unfolding set_plus_image
   apply (subst convex_hull_linear_image [symmetric])
   apply (simp add: linear_iff scaleR_right_distrib)
   apply (simp add: convex_hull_Times)
   done
 
-lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` t = {a} + t"
+lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` T = {a} + T"
   unfolding set_plus_def by auto
 
 lemma convex_hull_translation:
-  "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)"
+  "convex hull ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (convex hull S)"
   unfolding translation_eq_singleton_plus
   by (simp only: convex_hull_set_plus convex_hull_singleton)
 
 lemma convex_hull_scaling:
-  "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)"
-  by (simp add: convex_hull_linear_image)
+  "convex hull ((\<lambda>x. c *\<^sub>R x) ` S) = (\<lambda>x. c *\<^sub>R x) ` (convex hull S)"
+  using linear_scaleR by (rule convex_hull_linear_image [symmetric])
 
 lemma convex_hull_affinity:
-  "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)"
+  "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` S) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull S)"
   by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation)
 
 
@@ -6051,7 +5689,7 @@
       using assms mem_convex_alt[of S xx yy cx cy] x y by auto
     then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S"
       using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \<open>cx+cy>0\<close>
-      by (auto simp add: scaleR_right_distrib)
+      by (auto simp: scaleR_right_distrib)
     then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
       using x y by auto
   }
@@ -6091,8 +5729,7 @@
   fixes s :: "('a::euclidean_space) set"
   assumes "closed s" "convex s"
   shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
-  apply (rule set_eqI)
-  apply rule
+  apply (rule set_eqI, rule)
   unfolding Inter_iff Ball_def mem_Collect_eq
   apply (rule,rule,erule conjE)
 proof -
@@ -6105,8 +5742,7 @@
     apply (drule separating_hyperplane_closed_point[OF assms(2,1)])
     apply (erule exE)+
     apply (erule_tac x="-a" in allE)
-    apply (erule_tac x="-b" in allE)
-    apply auto
+    apply (erule_tac x="-b" in allE, auto)
     done
 qed auto
 
@@ -6124,7 +5760,7 @@
   then show ?thesis
     apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI)
     unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms(1), symmetric]
-    apply (auto simp add: Int_absorb1)
+    apply (auto simp: Int_absorb1)
     done
 qed
 
@@ -6179,8 +5815,7 @@
     next
       case False
       then have "sum u c \<le> sum (\<lambda>x. if x=v then u v else 0) c"
-        apply (rule_tac sum_mono)
-        apply auto
+        apply (rule_tac sum_mono, auto)
         done
       then show ?thesis
         unfolding sum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto
@@ -6190,20 +5825,18 @@
   then have *: "sum u {x\<in>c. u x > 0} > 0"
     unfolding less_le
     apply (rule_tac conjI)
-    apply (rule_tac sum_nonneg)
-    apply auto
+    apply (rule_tac sum_nonneg, auto)
     done
   moreover have "sum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = sum u c"
     "(\<Sum>x\<in>{x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>c. u x *\<^sub>R x)"
     using assms(1)
-    apply (rule_tac[!] sum.mono_neutral_left)
-    apply auto
+    apply (rule_tac[!] sum.mono_neutral_left, auto)
     done
   then have "sum u {x \<in> c. 0 < u x} = - sum u {x \<in> c. 0 > u x}"
     "(\<Sum>x\<in>{x \<in> c. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> c. 0 > u x}. u x *\<^sub>R x)"
     unfolding eq_neg_iff_add_eq_0
     using uv(1,4)
-    by (auto simp add: sum.union_inter_neutral[OF fin, symmetric])
+    by (auto simp: sum.union_inter_neutral[OF fin, symmetric])
   moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (sum u {x \<in> c. 0 < u x}) * - u x"
     apply rule
     apply (rule mult_nonneg_nonneg)
@@ -6215,7 +5848,7 @@
     apply (rule_tac x="{v \<in> c. u v < 0}" in exI)
     apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * - u y" in exI)
     using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def
-    apply (auto simp add: sum_negf sum_distrib_left[symmetric])
+    apply (auto simp: sum_negf sum_distrib_left[symmetric])
     done
   moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (sum u {x \<in> c. 0 < u x}) * u x"
     apply rule
@@ -6230,12 +5863,11 @@
     using assms(1)
     unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def
     using *
-    apply (auto simp add: sum_negf sum_distrib_left[symmetric])
+    apply (auto simp: sum_negf sum_distrib_left[symmetric])
     done
   ultimately show ?thesis
     apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI)
-    apply (rule_tac x="{v\<in>c. u v > 0}" in exI)
-    apply auto
+    apply (rule_tac x="{v\<in>c. u v > 0}" in exI, auto)
     done
 qed
 
@@ -6269,7 +5901,7 @@
     and "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
   shows "\<Inter>f \<noteq> {}"
   using assms
-proof (induct n arbitrary: f)
+proof (induction n arbitrary: f)
   case 0
   then show ?case by auto
 next
@@ -6277,46 +5909,39 @@
   have "finite f"
     using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite)
   show "\<Inter>f \<noteq> {}"
-    apply (cases "n = DIM('a)")
-    apply (rule Suc(5)[rule_format])
-    unfolding \<open>card f = Suc n\<close>
-  proof -
-    assume ng: "n \<noteq> DIM('a)"
-    then have "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})"
-      apply (rule_tac bchoice)
-      unfolding ex_in_conv
-      apply (rule, rule Suc(1)[rule_format])
-      unfolding card_Diff_singleton_if[OF \<open>finite f\<close>] \<open>card f = Suc n\<close>
-      defer
-      defer
-      apply (rule Suc(4)[rule_format])
-      defer
-      apply (rule Suc(5)[rule_format])
-      using Suc(3) \<open>finite f\<close>
-      apply auto
-      done
-    then obtain X where X: "\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
+  proof (cases "n = DIM('a)")
+    case True
+    then show ?thesis
+      by (simp add: Suc.prems(1) Suc.prems(4))
+  next
+    case False
+    have "\<Inter>(f - {s}) \<noteq> {}" if "s \<in> f" for s
+    proof (rule Suc.IH[rule_format])
+      show "card (f - {s}) = n"
+        by (simp add: Suc.prems(1) \<open>finite f\<close> that)
+      show "DIM('a) + 1 \<le> n"
+        using False Suc.prems(2) by linarith
+      show "\<And>t. \<lbrakk>t \<subseteq> f - {s}; card t = DIM('a) + 1\<rbrakk> \<Longrightarrow> \<Inter>t \<noteq> {}"
+        by (simp add: Suc.prems(4) subset_Diff_insert)
+    qed (use Suc in auto)
+    then have "\<forall>s\<in>f. \<exists>x. x \<in> \<Inter>(f - {s})"
+      by blast
+    then obtain X where X: "\<And>s. s\<in>f \<Longrightarrow> X s \<in> \<Inter>(f - {s})"
+      by metis
     show ?thesis
     proof (cases "inj_on X f")
       case False
-      then obtain s t where st: "s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t"
+      then obtain s t where "s\<noteq>t" and st: "s\<in>f" "t\<in>f" "X s = X t"
         unfolding inj_on_def by auto
       then have *: "\<Inter>f = \<Inter>(f - {s}) \<inter> \<Inter>(f - {t})" by auto
       show ?thesis
-        unfolding *
-        unfolding ex_in_conv[symmetric]
-        apply (rule_tac x="X s" in exI)
-        apply rule
-        apply (rule X[rule_format])
-        using X st
-        apply auto
-        done
+        by (metis "*" X disjoint_iff_not_equal st)
     next
       case True
       then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
         using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
         unfolding card_image[OF True] and \<open>card f = Suc n\<close>
-        using Suc(3) \<open>finite f\<close> and ng
+        using Suc(3) \<open>finite f\<close> and False
         by auto
       have "m \<subseteq> X ` f" "p \<subseteq> X ` f"
         using mp(2) by auto
@@ -6333,38 +5958,17 @@
         using inj_on_image_Int[OF True gh(3,4)]
         by auto
       have "convex hull (X ` h) \<subseteq> \<Inter>g" "convex hull (X ` g) \<subseteq> \<Inter>h"
-        apply (rule_tac [!] hull_minimal)
-        using Suc gh(3-4)
-        unfolding subset_eq
-        apply (rule_tac [2] convex_Inter, rule_tac [4] convex_Inter)
-        apply rule
-        prefer 3
-        apply rule
-      proof -
-        fix x
-        assume "x \<in> X ` g"
-        then obtain y where "y \<in> g" "x = X y"
-          unfolding image_iff ..
-        then show "x \<in> \<Inter>h"
-          using X[THEN bspec[where x=y]] using * f by auto
-      next
-        fix x
-        assume "x \<in> X ` h"
-        then obtain y where "y \<in> h" "x = X y"
-          unfolding image_iff ..
-        then show "x \<in> \<Inter>g"
-          using X[THEN bspec[where x=y]] using * f by auto
-      qed auto
+        by (rule hull_minimal; use X * f in \<open>auto simp: Suc.prems(3) convex_Inter\<close>)+
       then show ?thesis
         unfolding f using mp(3)[unfolded gh] by blast
     qed
-  qed auto
+  qed 
 qed
 
 theorem%important helly:
   fixes f :: "'a::euclidean_space set set"
   assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s"
-    and "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
+    and "\<And>t. \<lbrakk>t\<subseteq>f; card t = DIM('a) + 1\<rbrakk> \<Longrightarrow> \<Inter>t \<noteq> {}"
   shows "\<Inter>f \<noteq> {}"
   apply%unimportant (rule helly_induct)
   using assms
@@ -6374,104 +5978,109 @@
 
 subsection \<open>Epigraphs of convex functions\<close>
 
-definition%important "epigraph s (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
-
-lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y"
+definition%important "epigraph S (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> S \<and> f (fst xy) \<le> snd xy}"
+
+lemma mem_epigraph: "(x, y) \<in> epigraph S f \<longleftrightarrow> x \<in> S \<and> f x \<le> y"
   unfolding epigraph_def by auto
 
-lemma convex_epigraph: "convex (epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s"
-  unfolding convex_def convex_on_def
-  unfolding Ball_def split_paired_All epigraph_def
-  unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric]
-  apply safe
-  defer
-  apply (erule_tac x=x in allE)
-  apply (erule_tac x="f x" in allE)
-  apply safe
-  apply (erule_tac x=xa in allE)
-  apply (erule_tac x="f xa" in allE)
-  prefer 3
-  apply (rule_tac y="u * f a + v * f aa" in order_trans)
-  defer
-  apply (auto intro!:mult_left_mono add_mono)
-  done
-
-lemma convex_epigraphI: "convex_on s f \<Longrightarrow> convex s \<Longrightarrow> convex (epigraph s f)"
+lemma convex_epigraph: "convex (epigraph S f) \<longleftrightarrow> convex_on S f \<and> convex S"
+proof safe
+  assume L: "convex (epigraph S f)"
+  then show "convex_on S f"
+    by (auto simp: convex_def convex_on_def epigraph_def)
+  show "convex S"
+    using L
+    apply (clarsimp simp: convex_def convex_on_def epigraph_def)
+    apply (erule_tac x=x in allE)
+    apply (erule_tac x="f x" in allE, safe)
+    apply (erule_tac x=y in allE)
+    apply (erule_tac x="f y" in allE)
+    apply (auto simp: )
+    done
+next
+  assume "convex_on S f" "convex S"
+  then show "convex (epigraph S f)"
+    unfolding convex_def convex_on_def epigraph_def
+    apply safe
+     apply (rule_tac [2] y="u * f a + v * f aa" in order_trans)
+      apply (auto intro!:mult_left_mono add_mono)
+    done
+qed
+
+lemma convex_epigraphI: "convex_on S f \<Longrightarrow> convex S \<Longrightarrow> convex (epigraph S f)"
   unfolding convex_epigraph by auto
 
-lemma convex_epigraph_convex: "convex s \<Longrightarrow> convex_on s f \<longleftrightarrow> convex(epigraph s f)"
+lemma convex_epigraph_convex: "convex S \<Longrightarrow> convex_on S f \<longleftrightarrow> convex(epigraph S f)"
   by (simp add: convex_epigraph)
 
 
 subsubsection%unimportant \<open>Use this to derive general bound property of convex function\<close>
 
 lemma convex_on:
-  assumes "convex s"
-  shows "convex_on s f \<longleftrightarrow>
-    (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> sum u {1..k} = 1 \<longrightarrow>
-      f (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> sum (\<lambda>i. u i * f(x i)) {1..k})"
+  assumes "convex S"
+  shows "convex_on S f \<longleftrightarrow>
+    (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> S) \<and> sum u {1..k} = 1 \<longrightarrow>
+      f (sum (\<lambda>i. u i *\<^sub>R x i) {1..k}) \<le> sum (\<lambda>i. u i * f(x i)) {1..k})"
+
   unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq
   unfolding fst_sum snd_sum fst_scaleR snd_scaleR
   apply safe
-  apply (drule_tac x=k in spec)
-  apply (drule_tac x=u in spec)
-  apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec)
-  apply simp
-  using assms[unfolded convex]
-  apply simp
-  apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans)
-  defer
-  apply (rule sum_mono)
-  apply (erule_tac x=i in allE)
+    apply (drule_tac x=k in spec)
+    apply (drule_tac x=u in spec)
+    apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec)
+    apply simp
+  using assms[unfolded convex] apply simp
+  apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans, force)
+   apply (rule sum_mono)
+   apply (erule_tac x=i in allE)
   unfolding real_scaleR_def
-  apply (rule mult_left_mono)
-  using assms[unfolded convex]
-  apply auto
+   apply (rule mult_left_mono)
+  using assms[unfolded convex] apply auto
   done
 
 
 subsection%unimportant \<open>Convexity of general and special intervals\<close>
 
 lemma is_interval_convex:
-  fixes s :: "'a::euclidean_space set"
-  assumes "is_interval s"
-  shows "convex s"
+  fixes S :: "'a::euclidean_space set"
+  assumes "is_interval S"
+  shows "convex S"
 proof (rule convexI)
   fix x y and u v :: real
-  assume as: "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = 1"
+  assume as: "x \<in> S" "y \<in> S" "0 \<le> u" "0 \<le> v" "u + v = 1"
   then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0"
     by auto
   {
     fix a b
     assume "\<not> b \<le> u * a + v * b"
     then have "u * a < (1 - v) * b"
-      unfolding not_le using as(4) by (auto simp add: field_simps)
+      unfolding not_le using as(4) by (auto simp: field_simps)
     then have "a < b"
       unfolding * using as(4) *(2)
       apply (rule_tac mult_left_less_imp_less[of "1 - v"])
-      apply (auto simp add: field_simps)
+      apply (auto simp: field_simps)
       done
     then have "a \<le> u * a + v * b"
       unfolding * using as(4)
-      by (auto simp add: field_simps intro!:mult_right_mono)
+      by (auto simp: field_simps intro!:mult_right_mono)
   }
   moreover
   {
     fix a b
     assume "\<not> u * a + v * b \<le> a"
     then have "v * b > (1 - u) * a"
-      unfolding not_le using as(4) by (auto simp add: field_simps)
+      unfolding not_le using as(4) by (auto simp: field_simps)
     then have "a < b"
       unfolding * using as(4)
       apply (rule_tac mult_left_less_imp_less)
-      apply (auto simp add: field_simps)
+      apply (auto simp: field_simps)
       done
     then have "u * a + v * b \<le> b"
       unfolding **
       using **(2) as(3)
-      by (auto simp add: field_simps intro!:mult_right_mono)
+      by (auto simp: field_simps intro!:mult_right_mono)
   }
-  ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s"
+  ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> S"
     apply -
     apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)])
     using as(3-) DIM_positive[where 'a='a]
@@ -6480,8 +6089,8 @@
 qed
 
 lemma is_interval_connected:
-  fixes s :: "'a::euclidean_space set"
-  shows "is_interval s \<Longrightarrow> connected s"
+  fixes S :: "'a::euclidean_space set"
+  shows "is_interval S \<Longrightarrow> connected S"
   using is_interval_convex convex_connected by auto
 
 lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))"
@@ -6556,12 +6165,9 @@
   ultimately show False
     apply (rule_tac notE[OF as(1)[unfolded connected_def]])
     apply (rule_tac x = ?halfl in exI)
-    apply (rule_tac x = ?halfr in exI)
-    apply rule
-    apply (rule open_lessThan)
-    apply rule
-    apply (rule open_greaterThan)
-    apply auto
+    apply (rule_tac x = ?halfr in exI, rule)
+    apply (rule open_lessThan, rule)
+    apply (rule open_greaterThan, auto)
     done
 qed
 
@@ -6625,7 +6231,7 @@
   fixes f :: "real \<Rightarrow> 'a::euclidean_space"
   shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow>
     f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
-  by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on)
+  by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on)
 
 lemma ivt_decreasing_component_on_1:
   fixes f :: "real \<Rightarrow> 'a::euclidean_space"
@@ -6699,40 +6305,22 @@
   assumes "finite A" and "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)"
   using assms by (induct set: finite, simp, simp add: finite_set_plus)
 
-lemma set_sum_eq:
-  "finite A \<Longrightarrow> (\<Sum>i\<in>A. B i) = {\<Sum>i\<in>A. f i |f. \<forall>i\<in>A. f i \<in> B i}"
-  apply (induct set: finite)
-  apply simp
-  apply simp
-  apply (safe elim!: set_plus_elim)
-  apply (rule_tac x="fun_upd f x a" in exI)
-  apply simp
-  apply (rule_tac f="\<lambda>x. a + x" in arg_cong)
-  apply (rule sum.cong [OF refl])
-  apply clarsimp
-  apply fast
-  done
-
 lemma box_eq_set_sum_Basis:
   shows "{x. \<forall>i\<in>Basis. x\<bullet>i \<in> B i} = (\<Sum>i\<in>Basis. image (\<lambda>x. x *\<^sub>R i) (B i))"
-  apply (subst set_sum_eq [OF finite_Basis])
-  apply safe
+  apply (subst set_sum_alt [OF finite_Basis], safe)
   apply (fast intro: euclidean_representation [symmetric])
   apply (subst inner_sum_left)
+apply (rename_tac f)
   apply (subgoal_tac "(\<Sum>x\<in>Basis. f x \<bullet> i) = f i \<bullet> i")
   apply (drule (1) bspec)
   apply clarsimp
   apply (frule sum.remove [OF finite_Basis])
-  apply (erule trans)
-  apply simp
-  apply (rule sum.neutral)
-  apply clarsimp
+  apply (erule trans, simp)
+  apply (rule sum.neutral, clarsimp)
   apply (frule_tac x=i in bspec, assumption)
-  apply (drule_tac x=x in bspec, assumption)
-  apply clarsimp
+  apply (drule_tac x=x in bspec, assumption, clarsimp)
   apply (cut_tac u=x and v=i in inner_Basis, assumption+)
-  apply (rule ccontr)
-  apply simp
+  apply (rule ccontr, simp)
   done
 
 lemma convex_hull_set_sum:
@@ -6750,8 +6338,8 @@
   show "convex (cbox x y)"
     by (rule convex_box)
 next
-  fix s assume "{x, y} \<subseteq> s" and "convex s"
-  then show "cbox x y \<subseteq> s"
+  fix S assume "{x, y} \<subseteq> S" and "convex S"
+  then show "cbox x y \<subseteq> S"
     unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def
     by - (clarify, simp (no_asm_use), fast)
 qed
@@ -6782,74 +6370,53 @@
 text \<open>And this is a finite set of vertices.\<close>
 
 lemma unit_cube_convex_hull:
-  obtains s :: "'a::euclidean_space set"
-    where "finite s" and "cbox 0 (\<Sum>Basis) = convex hull s"
-  apply (rule that[of "{x::'a. \<forall>i\<in>Basis. x\<bullet>i=0 \<or> x\<bullet>i=1}"])
-  apply (rule finite_subset[of _ "(\<lambda>s. (\<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"])
-  prefer 3
-  apply (rule unit_interval_convex_hull)
-  apply rule
-  unfolding mem_Collect_eq
-proof -
-  fix x :: 'a
-  assume as: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1"
-  show "x \<in> (\<lambda>s. \<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i) ` Pow Basis"
-    apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
-    using as
-    apply (subst euclidean_eq_iff)
-    apply auto
-    done
-qed auto
+  obtains S :: "'a::euclidean_space set"
+  where "finite S" and "cbox 0 (\<Sum>Basis) = convex hull S"
+proof
+  show "finite {x::'a. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}"
+  proof (rule finite_subset, clarify)
+    show "finite ((\<lambda>S. \<Sum>i\<in>Basis. (if i \<in> S then 1 else 0) *\<^sub>R i) ` Pow Basis)"
+      using finite_Basis by blast
+    fix x :: 'a
+    assume as: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1"
+    show "x \<in> (\<lambda>S. \<Sum>i\<in>Basis. (if i\<in>S then 1 else 0) *\<^sub>R i) ` Pow Basis"
+      apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
+      using as
+       apply (subst euclidean_eq_iff, auto)
+      done
+  qed
+  show "cbox 0 One = convex hull {x. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}"
+    using unit_interval_convex_hull by blast
+qed 
 
 text \<open>Hence any cube (could do any nonempty interval).\<close>
 
 lemma cube_convex_hull:
   assumes "d > 0"
-  obtains s :: "'a::euclidean_space set" where
-    "finite s" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull s"
+  obtains S :: "'a::euclidean_space set" where
+    "finite S" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull S"
 proof -
-  let ?d = "(\<Sum>i\<in>Basis. d*\<^sub>Ri)::'a"
+  let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a"
   have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)"
-    apply (rule set_eqI, rule)
-    unfolding image_iff
-    defer
-    apply (erule bexE)
-  proof -
+  proof (intro set_eqI iffI)
     fix y
-    assume as: "y\<in>cbox (x - ?d) (x + ?d)"
+    assume "y \<in> cbox (x - ?d) (x + ?d)"
     then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)"
       using assms by (simp add: mem_box field_simps inner_simps)
-    with \<open>0 < d\<close> show "\<exists>z\<in>cbox 0 (\<Sum>Basis). y = x - ?d + (2 * d) *\<^sub>R z"
-      by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto
+    with \<open>0 < d\<close> show "y \<in> (\<lambda>y. x - sum (( *\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One"
+      by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"])
   next
-    fix y z
-    assume as: "z\<in>cbox 0 (\<Sum>Basis)" "y = x - ?d + (2*d) *\<^sub>R z"
-    have "\<And>i. i\<in>Basis \<Longrightarrow> 0 \<le> d * (z \<bullet> i) \<and> d * (z \<bullet> i) \<le> d"
-      using assms as(1)[unfolded mem_box]
-      apply (erule_tac x=i in ballE)
-      apply rule
-      prefer 2
-      apply (rule mult_right_le_one_le)
-      using assms
-      apply auto
-      done
+    fix y
+    assume "y \<in> (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One"
+    then obtain z where z: "z \<in> cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z"
+      by auto
     then show "y \<in> cbox (x - ?d) (x + ?d)"
-      unfolding as(2) mem_box
-      apply -
-      apply rule
-      using as(1)[unfolded mem_box]
-      apply (erule_tac x=i in ballE)
-      using assms
-      apply (auto simp: inner_simps)
-      done
+      using z assms by (auto simp: mem_box inner_simps)
   qed
-  obtain s where "finite s" "cbox 0 (\<Sum>Basis::'a) = convex hull s"
+  obtain S where "finite S" "cbox 0 (\<Sum>Basis::'a) = convex hull S"
     using unit_cube_convex_hull by auto
   then show ?thesis
-    apply (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"])
-    unfolding * and convex_hull_affinity
-    apply auto
-    done
+    by (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *)
 qed
 
 subsection%unimportant\<open>Representation of any interval as a finite convex hull\<close>
@@ -6877,13 +6444,13 @@
     next
       case False
       then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i"
-        by (auto simp add: field_simps)
+        by (auto simp: field_simps)
       from False have
           "min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (a \<bullet> i) else m i * (b \<bullet> i))"
           "max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (b \<bullet> i) else m i * (a \<bullet> i))"
         using a_le_b by (auto simp: min_def max_def mult_le_cancel_left)
       with False show ?thesis using a_le_b
-        unfolding * by (auto simp add: le_divide_eq divide_le_eq ac_simps)
+        unfolding * by (auto simp: le_divide_eq divide_le_eq ac_simps)
     qed
   qed
 qed simp
@@ -6895,7 +6462,7 @@
 lemma cbox_translation: "cbox (c + a) (c + b) = image (\<lambda>x. c + x) (cbox a b)"
   using image_affinity_cbox [of 1 c a b]
   using box_ne_empty [of "a+c" "b+c"]  box_ne_empty [of a b]
-  by (auto simp add: inner_left_distrib add.commute)
+  by (auto simp: inner_left_distrib add.commute)
 
 lemma cbox_image_unit_interval:
   fixes a :: "'a::euclidean_space"
@@ -6909,18 +6476,18 @@
 
 lemma closed_interval_as_convex_hull:
   fixes a :: "'a::euclidean_space"
-  obtains s where "finite s" "cbox a b = convex hull s"
+  obtains S where "finite S" "cbox a b = convex hull S"
 proof (cases "cbox a b = {}")
   case True with convex_hull_empty that show ?thesis
     by blast
 next
   case False
-  obtain s::"'a set" where "finite s" and eq: "cbox 0 One = convex hull s"
+  obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S"
     by (blast intro: unit_cube_convex_hull)
   have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)"
     by (rule linear_compose_sum) (auto simp: algebra_simps linearI)
-  have "finite ((+) a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` s)"
-    by (rule finite_imageI \<open>finite s\<close>)+
+  have "finite ((+) a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` S)"
+    by (rule finite_imageI \<open>finite S\<close>)+
   then show ?thesis
     apply (rule that)
     apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric])
@@ -6932,31 +6499,23 @@
 subsection%unimportant \<open>Bounded convex function on open set is continuous\<close>
 
 lemma convex_on_bounded_continuous:
-  fixes s :: "('a::real_normed_vector) set"
-  assumes "open s"
-    and "convex_on s f"
-    and "\<forall>x\<in>s. \<bar>f x\<bar> \<le> b"
-  shows "continuous_on s f"
+  fixes S :: "('a::real_normed_vector) set"
+  assumes "open S"
+    and "convex_on S f"
+    and "\<forall>x\<in>S. \<bar>f x\<bar> \<le> b"
+  shows "continuous_on S f"
   apply (rule continuous_at_imp_continuous_on)
   unfolding continuous_at_real_range
 proof (rule,rule,rule)
   fix x and e :: real
-  assume "x \<in> s" "e > 0"
+  assume "x \<in> S" "e > 0"
   define B where "B = \<bar>b\<bar> + 1"
-  have B: "0 < B" "\<And>x. x\<in>s \<Longrightarrow> \<bar>f x\<bar> \<le> B"
-    unfolding B_def
-    defer
-    apply (drule assms(3)[rule_format])
-    apply auto
-    done
-  obtain k where "k > 0" and k: "cball x k \<subseteq> s"
-    using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]]
-    using \<open>x\<in>s\<close> by auto
+  then have B:  "0 < B""\<And>x. x\<in>S \<Longrightarrow> \<bar>f x\<bar> \<le> B"
+    using assms(3) by auto 
+  obtain k where "k > 0" and k: "cball x k \<subseteq> S"
+    using \<open>x \<in> S\<close> assms(1) open_contains_cball_eq by blast
   show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
-    apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI)
-    apply rule
-    defer
-  proof (rule, rule)
+  proof (intro exI conjI allI impI)
     fix y
     assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)"
     show "\<bar>f y - f x\<bar> < e"
@@ -6965,79 +6524,63 @@
       define t where "t = k / norm (y - x)"
       have "2 < t" "0<t"
         unfolding t_def using as False and \<open>k>0\<close>
-        by (auto simp add:field_simps)
-      have "y \<in> s"
-        apply (rule k[unfolded subset_eq,rule_format])
+        by (auto simp:field_simps)
+      have "y \<in> S"
+        apply (rule k[THEN subsetD])
         unfolding mem_cball dist_norm
         apply (rule order_trans[of _ "2 * norm (x - y)"])
         using as
-        by (auto simp add: field_simps norm_minus_commute)
+        by (auto simp: field_simps norm_minus_commute)
       {
         define w where "w = x + t *\<^sub>R (y - x)"
-        have "w \<in> s"
-          unfolding w_def
-          apply (rule k[unfolded subset_eq,rule_format])
-          unfolding mem_cball dist_norm
-          unfolding t_def
-          using \<open>k>0\<close>
-          apply auto
-          done
+        have "w \<in> S"
+          using \<open>k>0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD])
         have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x"
-          by (auto simp add: algebra_simps)
+          by (auto simp: algebra_simps)
         also have "\<dots> = 0"
-          using \<open>t > 0\<close> by (auto simp add:field_simps)
+          using \<open>t > 0\<close> by (auto simp:field_simps)
         finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y"
           unfolding w_def using False and \<open>t > 0\<close>
-          by (auto simp add: algebra_simps)
-        have  "2 * B < e * t"
+          by (auto simp: algebra_simps)
+        have 2: "2 * B < e * t"
           unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
-          by (auto simp add:field_simps)
-        then have "(f w - f x) / t < e"
-          using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>x\<in>s\<close>]
-          using \<open>t > 0\<close> by (auto simp add:field_simps)
-        then have th1: "f y - f x < e"
-          apply -
-          apply (rule le_less_trans)
-          defer
-          apply assumption
+          by (auto simp:field_simps)
+        have "f y - f x \<le> (f w - f x) / t"
           using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
-          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> s\<close> \<open>w \<in> s\<close>
-          by (auto simp add:field_simps)
+          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> S\<close> \<open>w \<in> S\<close>
+          by (auto simp:field_simps)
+        also have "... < e"
+          using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>x\<in>S\<close>] 2 \<open>t > 0\<close> by (auto simp: field_simps)
+        finally have th1: "f y - f x < e" .
       }
       moreover
       {
         define w where "w = x - t *\<^sub>R (y - x)"
-        have "w \<in> s"
-          unfolding w_def
-          apply (rule k[unfolded subset_eq,rule_format])
-          unfolding mem_cball dist_norm
-          unfolding t_def
-          using \<open>k > 0\<close>
-          apply auto
-          done
+        have "w \<in> S"
+          using \<open>k > 0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD])
         have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x"
-          by (auto simp add: algebra_simps)
+          by (auto simp: algebra_simps)
         also have "\<dots> = x"
-          using \<open>t > 0\<close> by (auto simp add:field_simps)
+          using \<open>t > 0\<close> by (auto simp:field_simps)
         finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x"
           unfolding w_def using False and \<open>t > 0\<close>
-          by (auto simp add: algebra_simps)
+          by (auto simp: algebra_simps)
         have "2 * B < e * t"
           unfolding t_def
           using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
-          by (auto simp add:field_simps)
+          by (auto simp:field_simps)
         then have *: "(f w - f y) / t < e"
-          using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>y\<in>s\<close>]
+          using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>y\<in>S\<close>]
           using \<open>t > 0\<close>
-          by (auto simp add:field_simps)
+          by (auto simp:field_simps)
         have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y"
           using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
-          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> s\<close> \<open>w \<in> s\<close>
-          by (auto simp add:field_simps)
+          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> S\<close> \<open>w \<in> S\<close>
+          by (auto simp:field_simps)
         also have "\<dots> = (f w + t * f y) / (1 + t)"
-          using \<open>t > 0\<close> by (auto simp add: divide_simps)
+          using \<open>t > 0\<close> by (auto simp: divide_simps)
         also have "\<dots> < e + f y"
-          using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp add: field_simps)
+          using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp: field_simps)
         finally have "f x - f y < e" by auto
       }
       ultimately show ?thesis by auto
@@ -7062,13 +6605,13 @@
   have *: "x - (2 *\<^sub>R x - y) = y - x"
     by (simp add: scaleR_2)
   have z: "z \<in> cball x e"
-    using y unfolding z_def mem_cball dist_norm * by (auto simp add: norm_minus_commute)
+    using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute)
   have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x"
-    unfolding z_def by (auto simp add: algebra_simps)
+    unfolding z_def by (auto simp: algebra_simps)
   then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
     using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"]
     using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z]
-    by (auto simp add:field_simps)
+    by (auto simp:field_simps)
 next
   case False
   fix y
@@ -7121,12 +6664,7 @@
       by (simp add: dist_norm abs_le_iff algebra_simps)
     show "cball x d \<subseteq> convex hull c"
       unfolding 2
-      apply clarsimp
-      apply (simp only: dist_norm)
-      apply (subst inner_diff_left [symmetric])
-      apply simp
-      apply (erule (1) order_trans [OF Basis_le_norm])
-      done
+      by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le)
     have e': "e = (\<Sum>(i::'a)\<in>Basis. d)"
       by (simp add: d_def DIM_positive)
     show "convex hull c \<subseteq> cball x e"
@@ -7136,51 +6674,31 @@
       apply (rule order_trans [OF L2_set_le_sum])
       apply (rule zero_le_dist)
       unfolding e'
-      apply (rule sum_mono)
-      apply simp
+      apply (rule sum_mono, simp)
       done
   qed
   define k where "k = Max (f ` c)"
   have "convex_on (convex hull c) f"
     apply(rule convex_on_subset[OF assms(2)])
-    apply(rule subset_trans[OF _ e(1)])
-    apply(rule c1)
+    apply(rule subset_trans[OF c1 e(1)])
     done
   then have k: "\<forall>y\<in>convex hull c. f y \<le> k"
-    apply (rule_tac convex_on_convex_hull_bound)
-    apply assumption
-    unfolding k_def
-    apply (rule, rule Max_ge)
-    using c(1)
-    apply auto
-    done
-  have "d \<le> e"
-    unfolding d_def
-    apply (rule mult_imp_div_pos_le)
-    using \<open>e > 0\<close>
-    unfolding mult_le_cancel_left1
-    apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive)
-    done
+    apply (rule_tac convex_on_convex_hull_bound, assumption)
+    by (simp add: k_def c)
+  have "e \<le> e * real DIM('a)"
+    using e(2) real_of_nat_ge_one_iff by auto
+  then have "d \<le> e"
+    by (simp add: d_def divide_simps)
   then have dsube: "cball x d \<subseteq> cball x e"
     by (rule subset_cball)
   have conv: "convex_on (cball x d) f"
-    apply (rule convex_on_subset)
-    apply (rule convex_on_subset[OF assms(2)])
-    apply (rule e(1))
-    apply (rule dsube)
-    done
+    using \<open>convex_on (convex hull c) f\<close> c2 convex_on_subset by blast
   then have "\<forall>y\<in>cball x d. \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>"
-    apply (rule convex_bounds_lemma)
-    apply (rule ballI)
-    apply (rule k [rule_format])
-    apply (erule rev_subsetD)
-    apply (rule c2)
-    done
+    by (rule convex_bounds_lemma) (use c2 k in blast)
   then have "continuous_on (ball x d) f"
     apply (rule_tac convex_on_bounded_continuous)
     apply (rule open_ball, rule convex_on_subset[OF conv])
-    apply (rule ball_subset_cball)
-    apply force
+    apply (rule ball_subset_cball, force)
     done
   then show "continuous (at x) f"
     unfolding continuous_on_eq_continuous_at[OF open_ball]
--- a/src/HOL/Analysis/Derivative.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Derivative.thy	Thu May 03 15:07:14 2018 +0200
@@ -373,77 +373,76 @@
 
 lemma frechet_derivative_unique_within:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "(f has_derivative f') (at x within s)"
-    and "(f has_derivative f'') (at x within s)"
-    and "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> (x + d *\<^sub>R i) \<in> s"
+  assumes "(f has_derivative f') (at x within S)"
+    and "(f has_derivative f'') (at x within S)"
+    and "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> (x + d *\<^sub>R i) \<in> S"
   shows "f' = f''"
 proof -
   note as = assms(1,2)[unfolded has_derivative_def]
   then interpret f': bounded_linear f' by auto
   from as interpret f'': bounded_linear f'' by auto
-  have "x islimpt s" unfolding islimpt_approachable
+  have "x islimpt S" unfolding islimpt_approachable
   proof (rule, rule)
     fix e :: real
     assume "e > 0"
-    obtain d where "0 < \<bar>d\<bar>" and "\<bar>d\<bar> < e" and "x + d *\<^sub>R (SOME i. i \<in> Basis) \<in> s"
+    obtain d where "0 < \<bar>d\<bar>" and "\<bar>d\<bar> < e" and "x + d *\<^sub>R (SOME i. i \<in> Basis) \<in> S"
       using assms(3) SOME_Basis \<open>e>0\<close> by blast
-    then show "\<exists>x'\<in>s. x' \<noteq> x \<and> dist x' x < e"
+    then show "\<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e"
       apply (rule_tac x="x + d *\<^sub>R (SOME i. i \<in> Basis)" in bexI)
       unfolding dist_norm
       apply (auto simp: SOME_Basis nonzero_Basis)
       done
   qed
-  then have *: "netlimit (at x within s) = x"
+  then have *: "netlimit (at x within S) = x"
     apply (auto intro!: netlimit_within)
     by (metis trivial_limit_within)
   show ?thesis
-    apply (rule linear_eq_stdbasis)
-    unfolding linear_conv_bounded_linear
-    apply (rule as(1,2)[THEN conjunct1])+
-  proof (rule ccontr)
+  proof (rule linear_eq_stdbasis)
+    show "linear f'" "linear f''"
+      unfolding linear_conv_bounded_linear using as by auto
+  next
     fix i :: 'a
     assume i: "i \<in> Basis"
     define e where "e = norm (f' i - f'' i)"
-    assume "f' i \<noteq> f'' i"
-    then have "e > 0"
-      unfolding e_def by auto
-    obtain d where d:
-      "0 < d"
-      "(\<And>xa. xa\<in>s \<longrightarrow> 0 < dist xa x \<and> dist xa x < d \<longrightarrow>
-          dist ((f xa - f x - f' (xa - x)) /\<^sub>R norm (xa - x) -
-              (f xa - f x - f'' (xa - x)) /\<^sub>R norm (xa - x)) (0 - 0) < e)"
-      using tendsto_diff [OF as(1,2)[THEN conjunct2]]
-      unfolding * Lim_within
-      using \<open>e>0\<close> by blast
-    obtain c where c: "0 < \<bar>c\<bar>" "\<bar>c\<bar> < d \<and> x + c *\<^sub>R i \<in> s"
-      using assms(3) i d(1) by blast
-    have *: "norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R i)) =
+    show "f' i = f'' i"
+    proof (rule ccontr)
+      assume "f' i \<noteq> f'' i"
+      then have "e > 0"
+        unfolding e_def by auto
+      obtain d where d:
+        "0 < d"
+        "(\<And>y. y\<in>S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow>
+          dist ((f y - f x - f' (y - x)) /\<^sub>R norm (y - x) -
+              (f y - f x - f'' (y - x)) /\<^sub>R norm (y - x)) (0 - 0) < e)"
+        using tendsto_diff [OF as(1,2)[THEN conjunct2]]
+        unfolding * Lim_within
+        using \<open>e>0\<close> by blast
+      obtain c where c: "0 < \<bar>c\<bar>" "\<bar>c\<bar> < d \<and> x + c *\<^sub>R i \<in> S"
+        using assms(3) i d(1) by blast
+      have *: "norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R i)) =
         norm ((1 / \<bar>c\<bar>) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))"
-      unfolding scaleR_right_distrib by auto
-    also have "\<dots> = norm ((1 / \<bar>c\<bar>) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))"
-      unfolding f'.scaleR f''.scaleR
-      unfolding scaleR_right_distrib scaleR_minus_right
-      by auto
-    also have "\<dots> = e"
-      unfolding e_def
-      using c(1)
-      using norm_minus_cancel[of "f' i - f'' i"]
-      by auto
-    finally show False
-      using c
-      using d(2)[of "x + c *\<^sub>R i"]
-      unfolding dist_norm
-      unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff
-        scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib
-      using i
-      by (auto simp: inverse_eq_divide)
+        unfolding scaleR_right_distrib by auto
+      also have "\<dots> = norm ((1 / \<bar>c\<bar>) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))"
+        unfolding f'.scaleR f''.scaleR
+        unfolding scaleR_right_distrib scaleR_minus_right
+        by auto
+      also have "\<dots> = e"
+        unfolding e_def
+        using c(1)
+        using norm_minus_cancel[of "f' i - f'' i"]
+        by auto
+      finally show False
+        using c
+        using d(2)[of "x + c *\<^sub>R i"]
+        unfolding dist_norm
+        unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff
+          scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib
+        using i
+        by (auto simp: inverse_eq_divide)
+    qed
   qed
 qed
 
-lemma frechet_derivative_unique_at:
-  "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f'') (at x) \<Longrightarrow> f' = f''"
-  by (rule has_derivative_unique)
-
 lemma frechet_derivative_unique_within_closed_interval:
   fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
   assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"
@@ -506,12 +505,12 @@
   from assms(1) have *: "at x within box a b = at x"
     by (metis at_within_interior interior_open open_box)
   from assms(2,3) [unfolded *] show "f' = f''"
-    by (rule frechet_derivative_unique_at)
+    by (rule has_derivative_unique)
 qed
 
 lemma frechet_derivative_at:
   "(f has_derivative f') (at x) \<Longrightarrow> f' = frechet_derivative f (at x)"
-  apply (rule frechet_derivative_unique_at[of f])
+  apply (rule has_derivative_unique[of f])
   apply assumption
   unfolding frechet_derivative_works[symmetric]
   using differentiable_def
@@ -1197,13 +1196,13 @@
 
 lemma has_derivative_inverse_basic:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  assumes "(f has_derivative f') (at (g y))"
-    and "bounded_linear g'"
+  assumes derf: "(f has_derivative f') (at (g y))"
+    and ling': "bounded_linear g'"
     and "g' \<circ> f' = id"
-    and "continuous (at y) g"
-    and "open t"
-    and "y \<in> t"
-    and "\<forall>z\<in>t. f (g z) = z"
+    and contg: "continuous (at y) g"
+    and "open T"
+    and "y \<in> T"
+    and fg: "\<And>z. z \<in> T \<Longrightarrow> f (g z) = z"
   shows "(g has_derivative g') (at y)"
 proof -
   interpret f': bounded_linear f'
@@ -1214,70 +1213,48 @@
     using bounded_linear.pos_bounded[OF assms(2)] by blast
   have lem1: "\<forall>e>0. \<exists>d>0. \<forall>z.
     norm (z - y) < d \<longrightarrow> norm (g z - g y - g'(z - y)) \<le> e * norm (g z - g y)"
-  proof (rule, rule)
+  proof (intro allI impI)
     fix e :: real
     assume "e > 0"
     with C(1) have *: "e / C > 0" by auto
-    obtain d0 where d0:
-        "0 < d0"
-        "\<forall>ya. norm (ya - g y) < d0 \<longrightarrow> norm (f ya - f (g y) - f' (ya - g y)) \<le> e / C * norm (ya - g y)"
-      using assms(1)
-      unfolding has_derivative_at_alt
-      using * by blast
-    obtain d1 where d1:
-        "0 < d1"
-        "\<forall>x. 0 < dist x y \<and> dist x y < d1 \<longrightarrow> dist (g x) (g y) < d0"
-      using assms(4)
-      unfolding continuous_at Lim_at
-      using d0(1) by blast
-    obtain d2 where d2:
-        "0 < d2"
-        "\<forall>ya. dist ya y < d2 \<longrightarrow> ya \<in> t"
-      using assms(5)
-      unfolding open_dist
-      using assms(6) by blast
+    obtain d0 where  "0 < d0" and d0:
+        "\<And>u. norm (u - g y) < d0 \<Longrightarrow> norm (f u - f (g y) - f' (u - g y)) \<le> e / C * norm (u - g y)"
+      using derf * unfolding has_derivative_at_alt by blast
+    obtain d1 where "0 < d1" and d1: "\<And>x. \<lbrakk>0 < dist x y; dist x y < d1\<rbrakk> \<Longrightarrow> dist (g x) (g y) < d0"
+      using contg \<open>0 < d0\<close> unfolding continuous_at Lim_at by blast
+    obtain d2 where "0 < d2" and d2: "\<And>u. dist u y < d2 \<Longrightarrow> u \<in> T"
+      using \<open>open T\<close> \<open>y \<in> T\<close> unfolding open_dist by blast
     obtain d where d: "0 < d" "d < d1" "d < d2"
-      using real_lbound_gt_zero[OF d1(1) d2(1)] by blast
-    then show "\<exists>d>0. \<forall>z. norm (z - y) < d \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)"
-      apply (rule_tac x=d in exI)
-      apply rule
-      defer
-      apply rule
-      apply rule
-    proof -
+      using real_lbound_gt_zero[OF \<open>0 < d1\<close> \<open>0 < d2\<close>] by blast
+    show "\<exists>d>0. \<forall>z. norm (z - y) < d \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)"
+    proof (intro exI allI impI conjI)
       fix z
       assume as: "norm (z - y) < d"
-      then have "z \<in> t"
+      then have "z \<in> T"
         using d2 d unfolding dist_norm by auto
       have "norm (g z - g y - g' (z - y)) \<le> norm (g' (f (g z) - y - f' (g z - g y)))"
         unfolding g'.diff f'.diff
-        unfolding assms(3)[unfolded o_def id_def, THEN fun_cong]
-        unfolding assms(7)[rule_format,OF \<open>z\<in>t\<close>]
-        apply (subst norm_minus_cancel[symmetric])
-        apply auto
-        done
+        unfolding assms(3)[unfolded o_def id_def, THEN fun_cong] fg[OF \<open>z\<in>T\<close>]
+        by (simp add: norm_minus_commute)
       also have "\<dots> \<le> norm (f (g z) - y - f' (g z - g y)) * C"
         by (rule C(2))
       also have "\<dots> \<le> (e / C) * norm (g z - g y) * C"
-        apply (rule mult_right_mono)
-        apply (rule d0(2)[rule_format,unfolded assms(7)[rule_format,OF \<open>y\<in>t\<close>]])
-        apply (cases "z = y")
-        defer
-        apply (rule d1(2)[unfolded dist_norm,rule_format])
-        using as d C d0
-        apply auto
-        done
+      proof -
+        have "norm (g z - g y) < d0"
+          by (metis as cancel_comm_monoid_add_class.diff_cancel d(2) \<open>0 < d0\<close> d1 diff_gt_0_iff_gt diff_strict_mono dist_norm dist_self zero_less_dist_iff)
+        then show ?thesis
+          by (metis C(1) \<open>y \<in> T\<close> d0 fg real_mult_le_cancel_iff1)
+      qed
       also have "\<dots> \<le> e * norm (g z - g y)"
         using C by (auto simp add: field_simps)
       finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)"
         by simp
-    qed auto
+    qed (use d in auto)
   qed
   have *: "(0::real) < 1 / 2"
     by auto
-  obtain d where d:
-      "0 < d"
-      "\<forall>z. norm (z - y) < d \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> 1 / 2 * norm (g z - g y)"
+  obtain d where "0 < d" and d:
+      "\<And>z. norm (z - y) < d \<Longrightarrow> norm (g z - g y - g' (z - y)) \<le> 1/2 * norm (g z - g y)"
     using lem1 * by blast
   define B where "B = C * 2"
   have "B > 0"
@@ -1287,51 +1264,37 @@
     have "norm (g z - g y) \<le> norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))"
       by (rule norm_triangle_sub)
     also have "\<dots> \<le> norm (g' (z - y)) + 1 / 2 * norm (g z - g y)"
-      apply (rule add_left_mono)
-      using d and z
-      apply auto
-      done
+      by (rule add_left_mono) (use d z in auto)
     also have "\<dots> \<le> norm (z - y) * C + 1 / 2 * norm (g z - g y)"
-      apply (rule add_right_mono)
-      using C
-      apply auto
-      done
+      by (rule add_right_mono) (use C in auto)
     finally show "norm (g z - g y) \<le> B * norm (z - y)"
       unfolding B_def
       by (auto simp add: field_simps)
   qed
   show ?thesis
     unfolding has_derivative_at_alt
-    apply rule
-    apply (rule assms)
-    apply rule
-    apply rule
-  proof -
+  proof (intro conjI assms allI impI)
     fix e :: real
     assume "e > 0"
     then have *: "e / B > 0" by (metis \<open>B > 0\<close> divide_pos_pos)
-    obtain d' where d':
-        "0 < d'"
-        "\<forall>z. norm (z - y) < d' \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> e / B * norm (g z - g y)"
+    obtain d' where "0 < d'" and d':
+        "\<And>z. norm (z - y) < d' \<Longrightarrow> norm (g z - g y - g' (z - y)) \<le> e / B * norm (g z - g y)"
       using lem1 * by blast
     obtain k where k: "0 < k" "k < d" "k < d'"
-      using real_lbound_gt_zero[OF d(1) d'(1)] by blast
+      using real_lbound_gt_zero[OF \<open>0 < d\<close> \<open>0 < d'\<close>] by blast
     show "\<exists>d>0. \<forall>ya. norm (ya - y) < d \<longrightarrow> norm (g ya - g y - g' (ya - y)) \<le> e * norm (ya - y)"
-      apply (rule_tac x=k in exI)
-      apply auto
-    proof -
+    proof (intro exI allI impI conjI)
       fix z
       assume as: "norm (z - y) < k"
       then have "norm (g z - g y - g' (z - y)) \<le> e / B * norm(g z - g y)"
         using d' k by auto
       also have "\<dots> \<le> e * norm (z - y)"
         unfolding times_divide_eq_left pos_divide_le_eq[OF \<open>B>0\<close>]
-        using lem2[of z]
-        using k as using \<open>e > 0\<close>
+        using lem2[of z] k as \<open>e > 0\<close>
         by (auto simp add: field_simps)
       finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (z - y)"
         by simp
-    qed(insert k, auto)
+    qed (use k in auto)
   qed
 qed
 
@@ -1344,25 +1307,22 @@
     and "g' \<circ> f' = id"
     and "continuous (at (f x)) g"
     and "g (f x) = x"
-    and "open t"
-    and "f x \<in> t"
-    and "\<forall>y\<in>t. f (g y) = y"
+    and "open T"
+    and "f x \<in> T"
+    and "\<And>y. y \<in> T \<Longrightarrow> f (g y) = y"
   shows "(g has_derivative g') (at (f x))"
-  apply (rule has_derivative_inverse_basic)
-  using assms
-  apply auto
-  done
+  by (rule has_derivative_inverse_basic) (use assms in auto)
 
 text \<open>This is the version in Dieudonne', assuming continuity of f and g.\<close>
 
 lemma has_derivative_inverse_dieudonne:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  assumes "open s"
-    and "open (f ` s)"
-    and "continuous_on s f"
-    and "continuous_on (f ` s) g"
-    and "\<forall>x\<in>s. g (f x) = x"
-    and "x \<in> s"
+  assumes "open S"
+    and "open (f ` S)"
+    and "continuous_on S f"
+    and "continuous_on (f ` S) g"
+    and "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
+    and "x \<in> S"
     and "(f has_derivative f') (at x)"
     and "bounded_linear g'"
     and "g' \<circ> f' = id"
@@ -1377,11 +1337,11 @@
 
 lemma has_derivative_inverse:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  assumes "compact s"
-    and "x \<in> s"
-    and "f x \<in> interior (f ` s)"
-    and "continuous_on s f"
-    and "\<forall>y\<in>s. g (f y) = y"
+  assumes "compact S"
+    and "x \<in> S"
+    and fx: "f x \<in> interior (f ` S)"
+    and "continuous_on S f"
+    and "\<And>y. y \<in> S \<Longrightarrow> g (f y) = y"
     and "(f has_derivative f') (at x)"
     and "bounded_linear g'"
     and "g' \<circ> f' = id"
@@ -1389,20 +1349,17 @@
 proof -
   {
     fix y
-    assume "y \<in> interior (f ` s)"
-    then obtain x where "x \<in> s" and *: "y = f x"
-      unfolding image_iff
-      using interior_subset
-      by auto
+    assume "y \<in> interior (f ` S)"
+    then obtain x where "x \<in> S" and *: "y = f x"
+      by (meson imageE interior_subset subsetCE)
     have "f (g y) = y"
-      unfolding * and assms(5)[rule_format,OF \<open>x\<in>s\<close>] ..
+      unfolding * and assms(5)[rule_format,OF \<open>x\<in>S\<close>] ..
   } note * = this
   show ?thesis
-    apply (rule has_derivative_inverse_basic_x[OF assms(6-8)])
-    apply (rule continuous_on_interior[OF _ assms(3)])
-    apply (rule continuous_on_inv[OF assms(4,1)])
-    apply (rule assms(2,5) assms(5)[rule_format] open_interior assms(3))+
-    apply (metis *)
+    apply (rule has_derivative_inverse_basic_x[OF assms(6-8), where T = "interior (f ` S)"])
+    apply (rule continuous_on_interior[OF _ fx])
+    apply (rule continuous_on_inv)
+    apply (simp_all add: assms *)
     done
 qed
 
@@ -1411,13 +1368,13 @@
 
 lemma brouwer_surjective:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
-  assumes "compact t"
-    and "convex t"
-    and "t \<noteq> {}"
-    and "continuous_on t f"
-    and "\<forall>x\<in>s. \<forall>y\<in>t. x + (y - f y) \<in> t"
-    and "x \<in> s"
-  shows "\<exists>y\<in>t. f y = x"
+  assumes "compact T"
+    and "convex T"
+    and "T \<noteq> {}"
+    and "continuous_on T f"
+    and "\<And>x y. \<lbrakk>x\<in>S; y\<in>T\<rbrakk> \<Longrightarrow> x + (y - f y) \<in> T"
+    and "x \<in> S"
+  shows "\<exists>y\<in>T. f y = x"
 proof -
   have *: "\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y"
     by (auto simp add: algebra_simps)
@@ -1432,10 +1389,10 @@
 
 lemma brouwer_surjective_cball:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
-  assumes "e > 0"
-    and "continuous_on (cball a e) f"
-    and "\<forall>x\<in>s. \<forall>y\<in>cball a e. x + (y - f y) \<in> cball a e"
-    and "x \<in> s"
+  assumes "continuous_on (cball a e) f"
+    and "e > 0"
+    and "x \<in> S"
+    and "\<And>x y. \<lbrakk>x\<in>S; y\<in>cball a e\<rbrakk> \<Longrightarrow> x + (y - f y) \<in> cball a e"
   shows "\<exists>y\<in>cball a e. f y = x"
   apply (rule brouwer_surjective)
   apply (rule compact_cball convex_cball)+
@@ -1448,14 +1405,14 @@
 
 lemma sussmann_open_mapping:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
-  assumes "open s"
-    and "continuous_on s f"
-    and "x \<in> s"
+  assumes "open S"
+    and "continuous_on S f"
+    and "x \<in> S"
     and "(f has_derivative f') (at x)"
     and "bounded_linear g'" "f' \<circ> g' = id"
-    and "t \<subseteq> s"
-    and "x \<in> interior t"
-  shows "f x \<in> interior (f ` t)"
+    and "T \<subseteq> S"
+    and "x \<in> interior T"
+  shows "f x \<in> interior (f ` T)"
 proof -
   interpret f': bounded_linear f'
     using assms
@@ -1473,31 +1430,17 @@
     using assms(4)
     unfolding has_derivative_at_alt
     using * by blast
-  obtain e1 where e1: "0 < e1" "cball x e1 \<subseteq> t"
+  obtain e1 where e1: "0 < e1" "cball x e1 \<subseteq> T"
     using assms(8)
     unfolding mem_interior_cball
     by blast
   have *: "0 < e0 / B" "0 < e1 / B" using e0 e1 B by auto
   obtain e where e: "0 < e" "e < e0 / B" "e < e1 / B"
     using real_lbound_gt_zero[OF *] by blast
-  have "\<forall>z\<in>cball (f x) (e / 2). \<exists>y\<in>cball (f x) e. f (x + g' (y - f x)) = z"
-    apply rule
-    apply (rule brouwer_surjective_cball[where s="cball (f x) (e/2)"])
-    prefer 3
-    apply rule
-    apply rule
-  proof-
-    show "continuous_on (cball (f x) e) (\<lambda>y. f (x + g' (y - f x)))"
-      unfolding g'.diff
-      apply (rule continuous_on_compose[of _ _ f, unfolded o_def])
-      apply (rule continuous_intros linear_continuous_on[OF assms(5)])+
-      apply (rule continuous_on_subset[OF assms(2)])
-      apply rule
-      apply (unfold image_iff)
-      apply (erule bexE)
+  have lem: "\<exists>y\<in>cball (f x) e. f (x + g' (y - f x)) = z" if "z\<in>cball (f x) (e / 2)" for z
+  proof (rule brouwer_surjective_cball)
+    have z: "z \<in> S" if as: "y \<in>cball (f x) e" "z = x + (g' y - g' (f x))" for y z
     proof-
-      fix y z
-      assume as: "y \<in>cball (f x) e" "z = x + (g' y - g' (f x))"
       have "dist x z = norm (g' (f x) - g' y)"
         unfolding as(2) and dist_norm by auto
       also have "\<dots> \<le> norm (f x - y) * B"
@@ -1516,9 +1459,16 @@
       finally have "z \<in> cball x e1"
         unfolding mem_cball
         by force
-      then show "z \<in> s"
+      then show "z \<in> S"
         using e1 assms(7) by auto
     qed
+    show "continuous_on (cball (f x) e) (\<lambda>y. f (x + g' (y - f x)))"
+      unfolding g'.diff
+      apply (rule continuous_on_compose[of _ _ f, unfolded o_def])
+       apply (rule continuous_intros linear_continuous_on[OF assms(5)])+
+      apply (rule continuous_on_subset[OF assms(2)])
+      using z
+      by blast
   next
     fix y z
     assume as: "y \<in> cball (f x) (e / 2)" "z \<in> cball (f x) e"
@@ -1528,7 +1478,7 @@
       apply (rule mult_right_mono)
       using as(2)[unfolded mem_cball dist_norm] and B
       unfolding norm_minus_commute
-      apply auto
+       apply auto
       done
     also have "\<dots> < e0"
       using e and B
@@ -1563,7 +1513,7 @@
     finally show "y + (z - f (x + g' (z - f x))) \<in> cball (f x) e"
       unfolding mem_cball dist_norm
       by auto
-  qed (insert e, auto) note lem = this
+  qed (use e that in auto) 
   show ?thesis
     unfolding mem_interior
     apply (rule_tac x="e/2" in exI)
@@ -1589,13 +1539,13 @@
       done
     also have "\<dots> \<le> e1"
       using e B unfolding less_divide_eq by auto
-    finally have "x + g'(z - f x) \<in> t"
+    finally have "x + g'(z - f x) \<in> T"
       apply -
       apply (rule e1(2)[unfolded subset_eq,rule_format])
       unfolding mem_cball dist_norm
       apply auto
       done
-    then show "y \<in> f ` t"
+    then show "y \<in> f ` T"
       using z by auto
   qed (insert e, auto)
 qed
@@ -1728,9 +1678,9 @@
   fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
   assumes "a \<in> s"
       and "open s"
-      and "bounded_linear g'"
+      and bling: "bounded_linear g'"
       and "g' \<circ> f' a = id"
-      and "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative f' x) (at x)"
+      and derf: "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative f' x) (at x)"
       and "\<And>e. e > 0 \<Longrightarrow> \<exists>d>0. \<forall>x. dist a x < d \<longrightarrow> onorm (\<lambda>v. f' x v - f' a v) < e"
   obtains r where "r > 0" "ball a r \<subseteq> s" "inj_on f (ball a r)"
 proof -
@@ -1738,11 +1688,7 @@
     using assms by auto
   note f'g' = assms(4)[unfolded id_def o_def,THEN cong]
   have "g' (f' a (\<Sum>Basis)) = (\<Sum>Basis)" "(\<Sum>Basis) \<noteq> (0::'n)"
-    defer
-    apply (subst euclidean_eq_iff)
-    using f'g'
-    apply auto
-    done
+    using f'g' by auto
   then have *: "0 < onorm g'"
     unfolding onorm_pos_lt[OF assms(3)]
     by fastforce
@@ -1790,17 +1736,11 @@
         have *: "(\<lambda>v. v - g' (f' u v)) = g' \<circ> (\<lambda>w. f' a w - f' u w)"
           unfolding o_def and diff
           using f'g' by auto
+        have blin: "bounded_linear (f' a)"
+          using \<open>a \<in> s\<close> derf by blast
         show "(ph has_derivative (\<lambda>v. v - g' (f' u v))) (at u within ball a d)"
-          unfolding ph' *
-          apply (simp add: comp_def)
-          apply (rule bounded_linear.has_derivative[OF assms(3)])
-          apply (rule derivative_intros)
-          defer
-          apply (rule has_derivative_sub[where g'="\<lambda>x.0",unfolded diff_0_right])
-          apply (rule has_derivative_at_withinI)
-          using assms(5) and \<open>u \<in> s\<close> \<open>a \<in> s\<close>
-          apply (auto intro!: derivative_intros bounded_linear.has_derivative[of _ "\<lambda>x. x"] has_derivative_bounded_linear)
-          done
+          unfolding ph' * comp_def
+          by (rule \<open>u \<in> s\<close> derivative_eq_intros has_derivative_at_withinI [OF derf] bounded_linear.has_derivative [OF blin]  bounded_linear.has_derivative [OF bling] |simp)+
         have **: "bounded_linear (\<lambda>x. f' u x - f' a x)" "bounded_linear (\<lambda>x. f' a x - f' u x)"
           apply (rule_tac[!] bounded_linear_sub)
           apply (rule_tac[!] has_derivative_bounded_linear)
@@ -1874,21 +1814,20 @@
   qed
 qed
 
-lemma has_derivative_sequence_lipschitz:
+lemma has_derivative_sequence_Lipschitz:
   fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  assumes "convex s"
-    and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)"
-    and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
-  shows "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s.
+  assumes "convex S"
+    and "\<And>n x. x \<in> S \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within S)"
+    and "\<And>e. e > 0 \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
+    and "e > 0"
+  shows "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>y\<in>S.
     norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)"
-proof (rule, rule)
-  fix e :: real
-  assume "e > 0"
-  then have *: "2 * (1/2* e) = e" "1/2 * e >0"
-    by auto
-  obtain N where "\<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> 1 / 2 * e * norm h"
+proof -
+  have *: "2 * (1/2* e) = e" "1/2 * e >0"
+    using \<open>e>0\<close> by auto
+  obtain N where "\<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> 1 / 2 * e * norm h"
     using assms(3) *(2) by blast
-  then show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f m x - f n x - (f m y - f n y)) \<le> e * norm (x - y)"
+  then show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>y\<in>S. norm (f m x - f n x - (f m y - f n y)) \<le> e * norm (x - y)"
     apply (rule_tac x=N in exI)
     apply (rule has_derivative_sequence_lipschitz_lemma[where e="1/2 *e", unfolded *])
     using assms \<open>e > 0\<close>
@@ -1898,22 +1837,22 @@
 
 lemma has_derivative_sequence:
   fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach"
-  assumes "convex s"
-    and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)"
-    and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
-    and "x0 \<in> s"
+  assumes "convex S"
+    and "\<And>n x. x \<in> S \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within S)"
+    and "\<And>e. e > 0 \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
+    and "x0 \<in> S"
     and "((\<lambda>n. f n x0) \<longlongrightarrow> l) sequentially"
-  shows "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> (g has_derivative g'(x)) (at x within s)"
+  shows "\<exists>g. \<forall>x\<in>S. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> (g has_derivative g'(x)) (at x within S)"
 proof -
-  have lem1: "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s.
+  have lem1: "\<And>e. e > 0 \<Longrightarrow> \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>y\<in>S.
       norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)"
-    using assms(1,2,3) by (rule has_derivative_sequence_lipschitz)
-  have "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially"
+    using assms(1,2,3) by (rule has_derivative_sequence_Lipschitz)
+  have "\<exists>g. \<forall>x\<in>S. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially"
     apply (rule bchoice)
     unfolding convergent_eq_Cauchy
   proof
     fix x
-    assume "x \<in> s"
+    assume "x \<in> S"
     show "Cauchy (\<lambda>n. f n x)"
     proof (cases "x = x0")
       case True
@@ -1923,7 +1862,7 @@
       case False
       show ?thesis
         unfolding Cauchy_def
-      proof (rule, rule)
+      proof (intro allI impI)
         fix e :: real
         assume "e > 0"
         hence *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" using False by auto
@@ -1933,12 +1872,11 @@
           using *(1) by blast
         obtain N where N:
           "\<forall>m\<ge>N. \<forall>n\<ge>N.
-            \<forall>xa\<in>s. \<forall>y\<in>s. norm (f m xa - f n xa - (f m y - f n y)) \<le>
+            \<forall>xa\<in>S. \<forall>y\<in>S. norm (f m xa - f n xa - (f m y - f n y)) \<le>
               e / 2 / norm (x - x0) * norm (xa - y)"
         using lem1 *(2) by blast
         show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x) (f n x) < e"
-          apply (rule_tac x="max M N" in exI)
-        proof rule+
+        proof (intro exI allI impI)
           fix m n
           assume as: "max M N \<le>m" "max M N\<le>n"
           have "dist (f m x) (f n x) \<le>
@@ -1946,7 +1884,7 @@
             unfolding dist_norm
             by (rule norm_triangle_sub)
           also have "\<dots> \<le> norm (f m x0 - f n x0) + e / 2"
-            using N[rule_format,OF _ _ \<open>x\<in>s\<close> \<open>x0\<in>s\<close>, of m n] and as and False
+            using N[rule_format,OF _ _ \<open>x\<in>S\<close> \<open>x0\<in>S\<close>, of m n] and as and False
             by auto
           also have "\<dots> < e / 2 + e / 2"
             apply (rule add_strict_right_mono)
@@ -1960,27 +1898,24 @@
       qed
     qed
   qed
-  then obtain g where g: "\<forall>x\<in>s. (\<lambda>n. f n x) \<longlonglongrightarrow> g x" ..
-  have lem2: "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm ((f n x - f n y) - (g x - g y)) \<le> e * norm (x - y)"
+  then obtain g where g: "\<forall>x\<in>S. (\<lambda>n. f n x) \<longlonglongrightarrow> g x" ..
+  have lem2: "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>y\<in>S. norm ((f n x - f n y) - (g x - g y)) \<le> e * norm (x - y)"
   proof (rule, rule)
     fix e :: real
     assume *: "e > 0"
     obtain N where
-      N: "\<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f m x - f n x - (f m y - f n y)) \<le> e * norm (x - y)"
+      N: "\<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>y\<in>S. norm (f m x - f n x - (f m y - f n y)) \<le> e * norm (x - y)"
       using lem1 * by blast
-    show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)"
+    show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>y\<in>S. norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)"
       apply (rule_tac x=N in exI)
     proof rule+
       fix n x y
-      assume as: "N \<le> n" "x \<in> s" "y \<in> s"
+      assume as: "N \<le> n" "x \<in> S" "y \<in> S"
       have "((\<lambda>m. norm (f n x - f n y - (f m x - f m y))) \<longlongrightarrow> norm (f n x - f n y - (g x - g y))) sequentially"
         by (intro tendsto_intros g[rule_format] as)
       moreover have "eventually (\<lambda>m. norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)) sequentially"
         unfolding eventually_sequentially
-        apply (rule_tac x=N in exI)
-        apply rule
-        apply rule
-      proof -
+      proof (intro exI allI impI)
         fix m
         assume "N \<le> m"
         then show "norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)"
@@ -1991,11 +1926,11 @@
         by (simp add: tendsto_upperbound)
     qed
   qed
-  have "\<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> (g has_derivative g' x) (at x within s)"
+  have "\<forall>x\<in>S. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> (g has_derivative g' x) (at x within S)"
     unfolding has_derivative_within_alt2
   proof (intro ballI conjI)
     fix x
-    assume "x \<in> s"
+    assume "x \<in> S"
     then show "((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially"
       by (simp add: g)
     have lem3: "\<forall>u. ((\<lambda>n. f' n x u) \<longlongrightarrow> g' x u) sequentially"
@@ -2008,7 +1943,7 @@
       proof (cases "u = 0")
         case True
         have "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e * norm u) sequentially"
-          using assms(3)[folded eventually_sequentially] and \<open>0 < e\<close> and \<open>x \<in> s\<close>
+          using assms(3)[folded eventually_sequentially] and \<open>0 < e\<close> and \<open>x \<in> S\<close>
           by (fast elim: eventually_mono)
         then show ?thesis
           using \<open>u = 0\<close> and \<open>0 < e\<close> by (auto elim: eventually_mono)
@@ -2016,7 +1951,7 @@
         case False
         with \<open>0 < e\<close> have "0 < e / norm u" by simp
         then have "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e / norm u * norm u) sequentially"
-          using assms(3)[folded eventually_sequentially] and \<open>x \<in> s\<close>
+          using assms(3)[folded eventually_sequentially] and \<open>x \<in> S\<close>
           by (fast elim: eventually_mono)
         then show ?thesis
           using \<open>u \<noteq> 0\<close> by simp
@@ -2026,7 +1961,7 @@
     proof
       fix x' y z :: 'a
       fix c :: real
-      note lin = assms(2)[rule_format,OF \<open>x\<in>s\<close>,THEN has_derivative_bounded_linear]
+      note lin = assms(2)[rule_format,OF \<open>x\<in>S\<close>,THEN has_derivative_bounded_linear]
       show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'"
         apply (rule tendsto_unique[OF trivial_limit_sequentially])
         apply (rule lem3[rule_format])
@@ -2042,9 +1977,9 @@
         apply (rule lem3[rule_format])+
         done
       obtain N where N: "\<forall>h. norm (f' N x h - g' x h) \<le> 1 * norm h"
-        using assms(3) \<open>x \<in> s\<close> by (fast intro: zero_less_one)
+        using assms(3) \<open>x \<in> S\<close> by (fast intro: zero_less_one)
       have "bounded_linear (f' N x)"
-        using assms(2) \<open>x \<in> s\<close> by fast
+        using assms(2) \<open>x \<in> S\<close> by fast
       from bounded_linear.bounded [OF this]
       obtain K where K: "\<forall>h. norm (f' N x h) \<le> norm h * K" ..
       {
@@ -2060,36 +1995,36 @@
       }
       then show "\<exists>K. \<forall>h. norm (g' x h) \<le> norm h * K" by fast
     qed
-    show "\<forall>e>0. eventually (\<lambda>y. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)) (at x within s)"
+    show "\<forall>e>0. eventually (\<lambda>y. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)) (at x within S)"
     proof (rule, rule)
       fix e :: real
       assume "e > 0"
       then have *: "e / 3 > 0"
         by auto
-      obtain N1 where N1: "\<forall>n\<ge>N1. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e / 3 * norm h"
+      obtain N1 where N1: "\<forall>n\<ge>N1. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e / 3 * norm h"
         using assms(3) * by blast
       obtain N2 where
-          N2: "\<forall>n\<ge>N2. \<forall>x\<in>s. \<forall>y\<in>s. norm (f n x - f n y - (g x - g y)) \<le> e / 3 * norm (x - y)"
+          N2: "\<forall>n\<ge>N2. \<forall>x\<in>S. \<forall>y\<in>S. norm (f n x - f n y - (g x - g y)) \<le> e / 3 * norm (x - y)"
         using lem2 * by blast
       let ?N = "max N1 N2"
-      have "eventually (\<lambda>y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)) (at x within s)"
-        using assms(2)[unfolded has_derivative_within_alt2] and \<open>x \<in> s\<close> and * by fast
-      moreover have "eventually (\<lambda>y. y \<in> s) (at x within s)"
+      have "eventually (\<lambda>y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)) (at x within S)"
+        using assms(2)[unfolded has_derivative_within_alt2] and \<open>x \<in> S\<close> and * by fast
+      moreover have "eventually (\<lambda>y. y \<in> S) (at x within S)"
         unfolding eventually_at by (fast intro: zero_less_one)
-      ultimately show "\<forall>\<^sub>F y in at x within s. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)"
+      ultimately show "\<forall>\<^sub>F y in at x within S. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)"
       proof (rule eventually_elim2)
         fix y
-        assume "y \<in> s"
+        assume "y \<in> S"
         assume "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)"
         moreover have "norm (g y - g x - (f ?N y - f ?N x)) \<le> e / 3 * norm (y - x)"
-          using N2[rule_format, OF _ \<open>y \<in> s\<close> \<open>x \<in> s\<close>]
+          using N2[rule_format, OF _ \<open>y \<in> S\<close> \<open>x \<in> S\<close>]
           by (simp add: norm_minus_commute)
         ultimately have "norm (g y - g x - f' ?N x (y - x)) \<le> 2 * e / 3 * norm (y - x)"
           using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"]
           by (auto simp add: algebra_simps)
         moreover
         have " norm (f' ?N x (y - x) - g' x (y - x)) \<le> e / 3 * norm (y - x)"
-          using N1 \<open>x \<in> s\<close> by auto
+          using N1 \<open>x \<in> S\<close> by auto
         ultimately show "norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)"
           using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"]
           by (auto simp add: algebra_simps)
@@ -2103,77 +2038,63 @@
 
 lemma has_antiderivative_sequence:
   fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach"
-  assumes "convex s"
-    and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)"
-    and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
-  shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g' x) (at x within s)"
-proof (cases "s = {}")
+  assumes "convex S"
+    and der: "\<And>n x. x \<in> S \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within S)"
+    and no: "\<And>e. e > 0 \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
+  shows "\<exists>g. \<forall>x\<in>S. (g has_derivative g' x) (at x within S)"
+proof (cases "S = {}")
   case False
-  then obtain a where "a \<in> s"
+  then obtain a where "a \<in> S"
     by auto
-  have *: "\<And>P Q. \<exists>g. \<forall>x\<in>s. P g x \<and> Q g x \<Longrightarrow> \<exists>g. \<forall>x\<in>s. Q g x"
+  have *: "\<And>P Q. \<exists>g. \<forall>x\<in>S. P g x \<and> Q g x \<Longrightarrow> \<exists>g. \<forall>x\<in>S. Q g x"
     by auto
   show ?thesis
     apply (rule *)
-    apply (rule has_derivative_sequence[OF assms(1) _ assms(3), of "\<lambda>n x. f n x + (f 0 a - f n a)"])
-    apply (metis assms(2) has_derivative_add_const)
-    apply (rule \<open>a \<in> s\<close>)
+    apply (rule has_derivative_sequence [OF \<open>convex S\<close> _ no, of "\<lambda>n x. f n x + (f 0 a - f n a)"])
+       apply (metis assms(2) has_derivative_add_const)
+    using \<open>a \<in> S\<close> 
     apply auto
     done
 qed auto
 
 lemma has_antiderivative_limit:
   fixes g' :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'b::banach"
-  assumes "convex s"
-    and "\<forall>e>0. \<exists>f f'. \<forall>x\<in>s.
-      (f has_derivative (f' x)) (at x within s) \<and> (\<forall>h. norm (f' x h - g' x h) \<le> e * norm h)"
-  shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g' x) (at x within s)"
+  assumes "convex S"
+    and "\<And>e. e>0 \<Longrightarrow> \<exists>f f'. \<forall>x\<in>S.
+           (f has_derivative (f' x)) (at x within S) \<and> (\<forall>h. norm (f' x h - g' x h) \<le> e * norm h)"
+  shows "\<exists>g. \<forall>x\<in>S. (g has_derivative g' x) (at x within S)"
 proof -
-  have *: "\<forall>n. \<exists>f f'. \<forall>x\<in>s.
-    (f has_derivative (f' x)) (at x within s) \<and>
+  have *: "\<forall>n. \<exists>f f'. \<forall>x\<in>S.
+    (f has_derivative (f' x)) (at x within S) \<and>
     (\<forall>h. norm(f' x h - g' x h) \<le> inverse (real (Suc n)) * norm h)"
     by (simp add: assms(2))
   obtain f where
-    *: "\<forall>x. \<exists>f'. \<forall>xa\<in>s. (f x has_derivative f' xa) (at xa within s) \<and>
-      (\<forall>h. norm (f' xa h - g' xa h) \<le> inverse (real (Suc x)) * norm h)"
-    using *[THEN choice] ..
+    *: "\<And>x. \<exists>f'. \<forall>xa\<in>S. (f x has_derivative f' xa) (at xa within S) \<and>
+        (\<forall>h. norm (f' xa h - g' xa h) \<le> inverse (real (Suc x)) * norm h)"
+    using * by metis
   obtain f' where
-    f: "\<forall>x. \<forall>xa\<in>s. (f x has_derivative f' x xa) (at xa within s) \<and>
-      (\<forall>h. norm (f' x xa h - g' xa h) \<le> inverse (real (Suc x)) * norm h)"
-    using *[THEN choice] ..
+    f': "\<And>x. \<forall>z\<in>S. (f x has_derivative f' x z) (at z within S) \<and>
+            (\<forall>h. norm (f' x z h - g' z h) \<le> inverse (real (Suc x)) * norm h)"
+    using * by metis
   show ?thesis
-    apply (rule has_antiderivative_sequence[OF assms(1), of f f'])
-    defer
-    apply rule
-    apply rule
-  proof -
+  proof (rule has_antiderivative_sequence[OF \<open>convex S\<close>, of f f'])
     fix e :: real
     assume "e > 0"
     obtain N where N: "inverse (real (Suc N)) < e"
       using reals_Archimedean[OF \<open>e>0\<close>] ..
-    show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
-      apply (rule_tac x=N in exI)
-      apply rule
-      apply rule
-      apply rule
-      apply rule
-    proof -
+    show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
+    proof (intro exI allI ballI impI)
       fix n x h
-      assume n: "N \<le> n" and x: "x \<in> s"
+      assume n: "N \<le> n" and x: "x \<in> S"
       have *: "inverse (real (Suc n)) \<le> e"
         apply (rule order_trans[OF _ N[THEN less_imp_le]])
         using n
         apply (auto simp add: field_simps)
         done
       show "norm (f' n x h - g' x h) \<le> e * norm h"
-        using f[rule_format,THEN conjunct2, OF x, of n, THEN spec[where x=h]]
-        apply (rule order_trans)
-        using N *
-        apply (cases "h = 0")
-        apply auto
-        done
+        by (meson "*" mult_right_mono norm_ge_zero order.trans x f')
     qed
-  qed (insert f, auto)
+  qed (use f' in auto)
 qed
 
 
@@ -2181,12 +2102,12 @@
 
 lemma has_derivative_series:
   fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach"
-  assumes "convex s"
-    and "\<And>n x. x \<in> s \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within s)"
-    and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (sum (\<lambda>i. f' i x h) {..<n} - g' x h) \<le> e * norm h"
-    and "x \<in> s"
+  assumes "convex S"
+    and "\<And>n x. x \<in> S \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within S)"
+    and "\<And>e. e>0 \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm (sum (\<lambda>i. f' i x h) {..<n} - g' x h) \<le> e * norm h"
+    and "x \<in> S"
     and "(\<lambda>n. f n x) sums l"
-  shows "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums (g x) \<and> (g has_derivative g' x) (at x within s)"
+  shows "\<exists>g. \<forall>x\<in>S. (\<lambda>n. f n x) sums (g x) \<and> (g has_derivative g' x) (at x within S)"
   unfolding sums_def
   apply (rule has_derivative_sequence[OF assms(1) _ assms(3)])
   apply (metis assms(2) has_derivative_sum)
@@ -2197,20 +2118,19 @@
 
 lemma has_field_derivative_series:
   fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a"
-  assumes "convex s"
-  assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)"
-  assumes "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
-  assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)"
-  shows   "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (g has_field_derivative g' x) (at x within s)"
+  assumes "convex S"
+  assumes "\<And>n x. x \<in> S \<Longrightarrow> (f n has_field_derivative f' n x) (at x within S)"
+  assumes "uniform_limit S (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
+  assumes "x0 \<in> S" "summable (\<lambda>n. f n x0)"
+  shows   "\<exists>g. \<forall>x\<in>S. (\<lambda>n. f n x) sums g x \<and> (g has_field_derivative g' x) (at x within S)"
 unfolding has_field_derivative_def
 proof (rule has_derivative_series)
-  show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h"
-  proof (intro allI impI)
-    fix e :: real assume "e > 0"
-    with assms(3) obtain N where N: "\<And>n x. n \<ge> N \<Longrightarrow> x \<in> s \<Longrightarrow> norm ((\<Sum>i<n. f' i x) - g' x) < e"
+  show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" if "e > 0" for e
+  proof -
+    from that assms(3) obtain N where N: "\<And>n x. n \<ge> N \<Longrightarrow> x \<in> S \<Longrightarrow> norm ((\<Sum>i<n. f' i x) - g' x) < e"
       unfolding uniform_limit_iff eventually_at_top_linorder dist_norm by blast
     {
-      fix n :: nat and x h :: 'a assume nx: "n \<ge> N" "x \<in> s"
+      fix n :: nat and x h :: 'a assume nx: "n \<ge> N" "x \<in> S"
       have "norm ((\<Sum>i<n. f' i x * h) - g' x * h) = norm ((\<Sum>i<n. f' i x) - g' x) * norm h"
         by (simp add: norm_mult [symmetric] ring_distribs sum_distrib_right)
       also from N[OF nx] have "norm ((\<Sum>i<n. f' i x) - g' x) \<le> e" by simp
@@ -2218,55 +2138,55 @@
         by (intro mult_right_mono) simp_all
       finally have "norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" .
     }
-    thus "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" by blast
+    thus "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" by blast
   qed
-qed (insert assms, auto simp: has_field_derivative_def)
+qed (use assms in \<open>auto simp: has_field_derivative_def\<close>)
 
 lemma has_field_derivative_series':
   fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a"
-  assumes "convex s"
-  assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)"
-  assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)"
-  assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" "x \<in> interior s"
+  assumes "convex S"
+  assumes "\<And>n x. x \<in> S \<Longrightarrow> (f n has_field_derivative f' n x) (at x within S)"
+  assumes "uniformly_convergent_on S (\<lambda>n x. \<Sum>i<n. f' i x)"
+  assumes "x0 \<in> S" "summable (\<lambda>n. f n x0)" "x \<in> interior S"
   shows   "summable (\<lambda>n. f n x)" "((\<lambda>x. \<Sum>n. f n x) has_field_derivative (\<Sum>n. f' n x)) (at x)"
 proof -
-  from \<open>x \<in> interior s\<close> have "x \<in> s" using interior_subset by blast
+  from \<open>x \<in> interior S\<close> have "x \<in> S" using interior_subset by blast
   define g' where [abs_def]: "g' x = (\<Sum>i. f' i x)" for x
-  from assms(3) have "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
+  from assms(3) have "uniform_limit S (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
     by (simp add: uniformly_convergent_uniform_limit_iff suminf_eq_lim g'_def)
   from has_field_derivative_series[OF assms(1,2) this assms(4,5)] obtain g where g:
-    "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x"
-    "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast
-  from g(1)[OF \<open>x \<in> s\<close>] show "summable (\<lambda>n. f n x)" by (simp add: sums_iff)
-  from g(2)[OF \<open>x \<in> s\<close>] \<open>x \<in> interior s\<close> have "(g has_field_derivative g' x) (at x)"
-    by (simp add: at_within_interior[of x s])
+    "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>n. f n x) sums g x"
+    "\<And>x. x \<in> S \<Longrightarrow> (g has_field_derivative g' x) (at x within S)" by blast
+  from g(1)[OF \<open>x \<in> S\<close>] show "summable (\<lambda>n. f n x)" by (simp add: sums_iff)
+  from g(2)[OF \<open>x \<in> S\<close>] \<open>x \<in> interior S\<close> have "(g has_field_derivative g' x) (at x)"
+    by (simp add: at_within_interior[of x S])
   also have "(g has_field_derivative g' x) (at x) \<longleftrightarrow>
                 ((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' x) (at x)"
-    using eventually_nhds_in_nhd[OF \<open>x \<in> interior s\<close>] interior_subset[of s] g(1)
+    using eventually_nhds_in_nhd[OF \<open>x \<in> interior S\<close>] interior_subset[of S] g(1)
     by (intro DERIV_cong_ev) (auto elim!: eventually_mono simp: sums_iff)
   finally show "((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' x) (at x)" .
 qed
 
 lemma differentiable_series:
   fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a"
-  assumes "convex s" "open s"
-  assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
-  assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)"
-  assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" and x: "x \<in> s"
+  assumes "convex S" "open S"
+  assumes "\<And>n x. x \<in> S \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
+  assumes "uniformly_convergent_on S (\<lambda>n x. \<Sum>i<n. f' i x)"
+  assumes "x0 \<in> S" "summable (\<lambda>n. f n x0)" and x: "x \<in> S"
   shows   "summable (\<lambda>n. f n x)" and "(\<lambda>x. \<Sum>n. f n x) differentiable (at x)"
 proof -
-  from assms(4) obtain g' where A: "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
+  from assms(4) obtain g' where A: "uniform_limit S (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
     unfolding uniformly_convergent_on_def by blast
-  from x and \<open>open s\<close> have s: "at x within s = at x" by (rule at_within_open)
-  have "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (g has_field_derivative g' x) (at x within s)"
-    by (intro has_field_derivative_series[of s f f' g' x0] assms A has_field_derivative_at_within)
-  then obtain g where g: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x"
-    "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast
+  from x and \<open>open S\<close> have S: "at x within S = at x" by (rule at_within_open)
+  have "\<exists>g. \<forall>x\<in>S. (\<lambda>n. f n x) sums g x \<and> (g has_field_derivative g' x) (at x within S)"
+    by (intro has_field_derivative_series[of S f f' g' x0] assms A has_field_derivative_at_within)
+  then obtain g where g: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>n. f n x) sums g x"
+    "\<And>x. x \<in> S \<Longrightarrow> (g has_field_derivative g' x) (at x within S)" by blast
   from g[OF x] show "summable (\<lambda>n. f n x)" by (auto simp: summable_def)
   from g(2)[OF x] have g': "(g has_derivative ( * ) (g' x)) (at x)"
-    by (simp add: has_field_derivative_def s)
+    by (simp add: has_field_derivative_def S)
   have "((\<lambda>x. \<Sum>n. f n x) has_derivative ( * ) (g' x)) (at x)"
-    by (rule has_derivative_transform_within_open[OF g' \<open>open s\<close> x])
+    by (rule has_derivative_transform_within_open[OF g' \<open>open S\<close> x])
        (insert g, auto simp: sums_iff)
   thus "(\<lambda>x. \<Sum>n. f n x) differentiable (at x)" unfolding differentiable_def
     by (auto simp: summable_def differentiable_def has_field_derivative_def)
@@ -2274,32 +2194,32 @@
 
 lemma differentiable_series':
   fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a"
-  assumes "convex s" "open s"
-  assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
-  assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)"
-  assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)"
+  assumes "convex S" "open S"
+  assumes "\<And>n x. x \<in> S \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
+  assumes "uniformly_convergent_on S (\<lambda>n x. \<Sum>i<n. f' i x)"
+  assumes "x0 \<in> S" "summable (\<lambda>n. f n x0)"
   shows   "(\<lambda>x. \<Sum>n. f n x) differentiable (at x0)"
-  using differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+
+  using differentiable_series[OF assms, of x0] \<open>x0 \<in> S\<close> by blast+
 
 text \<open>Considering derivative @{typ "real \<Rightarrow> 'b::real_normed_vector"} as a vector.\<close>
 
 definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)"
 
 lemma vector_derivative_unique_within:
-  assumes not_bot: "at x within s \<noteq> bot"
-    and f': "(f has_vector_derivative f') (at x within s)"
-    and f'': "(f has_vector_derivative f'') (at x within s)"
+  assumes not_bot: "at x within S \<noteq> bot"
+    and f': "(f has_vector_derivative f') (at x within S)"
+    and f'': "(f has_vector_derivative f'') (at x within S)"
   shows "f' = f''"
 proof -
   have "(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')"
   proof (rule frechet_derivative_unique_within)
-    show "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R i \<in> s"
+    show "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R i \<in> S"
     proof clarsimp
       fix e :: real assume "0 < e"
-      with islimpt_approachable_real[of x s] not_bot
-      obtain x' where "x' \<in> s" "x' \<noteq> x" "\<bar>x' - x\<bar> < e"
+      with islimpt_approachable_real[of x S] not_bot
+      obtain x' where "x' \<in> S" "x' \<noteq> x" "\<bar>x' - x\<bar> < e"
         by (auto simp add: trivial_limit_within)
-      then show "\<exists>d. d \<noteq> 0 \<and> \<bar>d\<bar> < e \<and> x + d \<in> s"
+      then show "\<exists>d. d \<noteq> 0 \<and> \<bar>d\<bar> < e \<and> x + d \<in> S"
         by (intro exI[of _ "x' - x"]) auto
     qed
   qed (insert f' f'', auto simp: has_vector_derivative_def)
@@ -2329,8 +2249,8 @@
 qed (auto simp: vector_derivative_def has_vector_derivative_def differentiable_def)
 
 lemma vector_derivative_within:
-  assumes not_bot: "at x within s \<noteq> bot" and y: "(f has_vector_derivative y) (at x within s)"
-  shows "vector_derivative f (at x within s) = y"
+  assumes not_bot: "at x within S \<noteq> bot" and y: "(f has_vector_derivative y) (at x within S)"
+  shows "vector_derivative f (at x within S) = y"
   using y
   by (intro vector_derivative_unique_within[OF not_bot vector_derivative_works[THEN iffD1] y])
      (auto simp: differentiable_def has_vector_derivative_def)
@@ -2359,11 +2279,11 @@
   by (metis has_field_derivative_def has_real_derivative)
 
 lemma has_vector_derivative_cong_ev:
-  assumes *: "eventually (\<lambda>x. x \<in> s \<longrightarrow> f x = g x) (nhds x)" "f x = g x"
-  shows "(f has_vector_derivative f') (at x within s) = (g has_vector_derivative f') (at x within s)"
+  assumes *: "eventually (\<lambda>x. x \<in> S \<longrightarrow> f x = g x) (nhds x)" "f x = g x"
+  shows "(f has_vector_derivative f') (at x within S) = (g has_vector_derivative f') (at x within S)"
   unfolding has_vector_derivative_def has_derivative_def
   using *
-  apply (cases "at x within s \<noteq> bot")
+  apply (cases "at x within S \<noteq> bot")
   apply (intro refl conj_cong filterlim_cong)
   apply (auto simp: netlimit_within eventually_at_filter elim: eventually_mono)
   done
--- a/src/HOL/Analysis/Determinants.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Determinants.thy	Thu May 03 15:07:14 2018 +0200
@@ -476,21 +476,26 @@
   assumes x: "x \<in> vec.span {row j A |j. j \<noteq> i}"
   shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A"
   using x
-proof (induction rule: vec.span_induct_alt)
-  case 1
+proof (induction rule: span_induct_alt)
+  case base
+  {
+    fix k
+    have "(if k = i then row i A + 0 else row k A) = row k A"
+      by simp
+  }
   then show ?case
-    by (rule cong[of det, OF refl]) (vector row_def)
+    apply -
+    apply (rule cong[of det, OF refl])
+    apply (vector row_def)
+    done
 next
-  case (2 c z y)
-  note zS = 2(1)
-    and Py = 2(2)
-  let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)"
-  let ?P = "\<lambda>x. ?d (row i A + x) = det A"
-  from zS obtain j where j: "z = row j A" "i \<noteq> j"
+  case (step c z y)
+  then obtain j where j: "z = row j A" "i \<noteq> j"
     by blast
   let ?w = "row i A + y"
   have th0: "row i A + (c*s z + y) = ?w + c*s z"
     by vector
+  let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)"
   have thz: "?d z = 0"
     apply (rule det_identical_rows[OF j(2)])
     using j
@@ -498,10 +503,11 @@
     done
   have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)"
     unfolding th0 ..
+  then have "?d (row i A + (c*s z + y)) = det A"
+    unfolding thz step.IH det_row_mul[of i] det_row_add[of i] by simp
   then show ?case
-    unfolding thz Py det_row_mul[of i] det_row_add[of i]
-    by simp
-qed
+    unfolding scalar_mult_eq_scaleR .
+qed 
 
 lemma matrix_id [simp]: "det (matrix id) = 1"
   by (simp add: matrix_id_mat_1)
@@ -834,26 +840,16 @@
       unfolding invertible_right_inverse
       unfolding matrix_right_invertible_independent_rows
       by blast
-    have *: "\<And>(a::'a^'n) b. a + b = 0 \<Longrightarrow> -a = b"
-      apply (drule_tac f="(+) (- a)" in cong[OF refl])
-      apply (simp only: ab_left_minus add.assoc[symmetric])
-      apply simp
-      done
-    from c ci
     have thr0: "- row i A = sum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
-      unfolding sum.remove[OF fU iU] sum_cmul
-      apply -
-      apply (rule vector_mul_lcancel_imp[OF ci])
-      apply (auto simp add: field_simps)
-      unfolding *
-      apply rule
-      done
-    have thr: "- row i A \<in> vec.span {row j A| j. j \<noteq> i}"
+      unfolding sum_cmul
+      using c ci   
+      by (auto simp add: sum.remove[OF fU iU] eq_vector_fraction_iff add_eq_0_iff)
+    have thr: "- row i A \<in> span {row j A| j. j \<noteq> i}"
       unfolding thr0
       apply (rule vec.span_sum)
       apply simp
-      apply (rule vec.span_scale[folded scalar_mult_eq_scaleR])+
-      apply (rule vec.span_base)
+      apply (rule span_mul)
+      apply (rule span_superset)
       apply auto
       done
     let ?B = "(\<chi> k. if k = i then 0 else row k A) :: 'a^'n^'n"
--- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy	Thu May 03 15:07:14 2018 +0200
@@ -244,7 +244,7 @@
       finally have "?\<mu> E - 2*e \<le> ?\<mu> (E - (E \<inter> \<Union>(snd`(p - s))))"
         using \<open>0 < e\<close> \<open>E \<in> sets ?L\<close> tagged_division_ofD(1)[OF p(1)]
         by (subst emeasure_Diff)
-           (auto simp: ennreal_plus[symmetric] top_unique simp del: ennreal_plus
+           (auto simp: top_unique reorient: ennreal_plus
                  intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L)
       also have "\<dots> \<le> ?\<mu> (\<Union>x\<in>p \<inter> s. snd x)"
       proof (safe intro!: emeasure_mono subsetI)
@@ -481,7 +481,7 @@
     by (simp add: nn_integral_add)
   with add obtain a b where "0 \<le> a" "0 \<le> b" "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ennreal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal b" "r = a + b"
     by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases)
-       (auto simp: add_top nn_integral_add top_add ennreal_plus[symmetric] simp del: ennreal_plus)
+       (auto simp: add_top nn_integral_add top_add reorient: ennreal_plus)
   with add show ?case
     by (auto intro!: has_integral_add)
 next
@@ -1325,6 +1325,13 @@
     by (simp add: negligible_interval box_eq_empty algebra_simps divide_simps mult_le_0_iff)
 qed
 
+lemma negligible_convex_interior:
+   "convex S \<Longrightarrow> (negligible S \<longleftrightarrow> interior S = {})"
+  apply safe
+  apply (metis interior_subset negligible_subset open_interior open_not_negligible)
+   apply (metis Diff_empty closure_subset frontier_def negligible_convex_frontier negligible_subset)
+  done
+
 lemma measure_eq_0_null_sets: "S \<in> null_sets M \<Longrightarrow> measure M S = 0"
   by (auto simp: measure_def null_sets_def)
 
--- a/src/HOL/Analysis/FPS_Convergence.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/FPS_Convergence.thy	Thu May 03 15:07:14 2018 +0200
@@ -271,7 +271,7 @@
 lemma fps_conv_radius_uminus [simp]:
   "fps_conv_radius (-f) = fps_conv_radius f"
   using fps_conv_radius_cmult_left[of "-1" f]
-  by (simp add: fps_const_neg [symmetric] del: fps_const_neg)
+  by (simp reorient: fps_const_neg)
 
 lemma fps_conv_radius_add: "fps_conv_radius (f + g) \<ge> min (fps_conv_radius f) (fps_conv_radius g)"
   unfolding fps_conv_radius_def using conv_radius_add_ge[of "fps_nth f" "fps_nth g"]
--- a/src/HOL/Analysis/Fashoda_Theorem.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Fashoda_Theorem.thy	Thu May 03 15:07:14 2018 +0200
@@ -109,21 +109,16 @@
   have lem1: "\<forall>z::real^2. infnorm (negatex z) = infnorm z"
     unfolding negatex_def infnorm_2 vector_2 by auto
   have lem2: "\<forall>z. z \<noteq> 0 \<longrightarrow> infnorm (sqprojection z) = 1"
-    unfolding sqprojection_def
-    unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR]
-    unfolding abs_inverse real_abs_infnorm
-    apply (subst infnorm_eq_0[symmetric])
-    apply auto
-    done
+    unfolding sqprojection_def infnorm_mul[unfolded scalar_mult_eq_scaleR]
+    by (simp add: real_abs_infnorm infnorm_eq_0)
   let ?F = "\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w"
-  have *: "\<And>i. (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 = {-1 .. 1}"
-    apply (rule set_eqI)
-    unfolding image_iff Bex_def mem_box_cart interval_cbox_cart
-    apply rule
-    defer
-    apply (rule_tac x="vec x" in exI)
-    apply auto
-    done
+  have *: "\<And>i. (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 = {-1..1}"
+  proof 
+    show "(\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 \<subseteq> {-1..1}" for i
+      by (auto simp: mem_box_cart)
+    show "{-1..1} \<subseteq> (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1" for i
+      by (clarsimp simp: image_iff mem_box_cart Bex_def) (metis (no_types, hide_lams) vec_component)
+  qed
   {
     fix x
     assume "x \<in> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w) ` (cbox (- 1) (1::real^2))"
@@ -136,37 +131,19 @@
       unfolding mem_box_cart atLeastAtMost_iff
       by auto
   } note x0 = this
-  have 21: "\<And>i::2. i \<noteq> 1 \<Longrightarrow> i = 2"
-    using UNIV_2 by auto
   have 1: "box (- 1) (1::real^2) \<noteq> {}"
     unfolding interval_eq_empty_cart by auto
-  have 2: "continuous_on (cbox (- 1) 1) (negatex \<circ> sqprojection \<circ> ?F)"
+  have "negatex (x + y) $ i = (negatex x + negatex y) $ i \<and> negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i"
+    for i x y c
+    using exhaust_2 [of i] by (auto simp: negatex_def)
+  then have "bounded_linear negatex"
+    by (simp add: bounded_linearI' vec_eq_iff)
+  then have 2: "continuous_on (cbox (- 1) 1) (negatex \<circ> sqprojection \<circ> ?F)"
     apply (intro continuous_intros continuous_on_component)
-    unfolding *
-    apply (rule assms)+
-    apply (subst sqprojection_def)
-    apply (intro continuous_intros)
-    apply (simp add: infnorm_eq_0 x0)
-    apply (rule linear_continuous_on)
-  proof -
-    show "bounded_linear negatex"
-      apply (rule bounded_linearI')
-      unfolding vec_eq_iff
-    proof (rule_tac[!] allI)
-      fix i :: 2
-      fix x y :: "real^2"
-      fix c :: real
-      show "negatex (x + y) $ i =
-        (negatex x + negatex y) $ i" "negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i"
-        apply -
-        apply (case_tac[!] "i\<noteq>1")
-        prefer 3
-        apply (drule_tac[1-2] 21)
-        unfolding negatex_def
-        apply (auto simp add:vector_2)
-        done
-    qed
-  qed
+    unfolding * sqprojection_def
+    apply (intro assms continuous_intros)+
+     apply (simp_all add: infnorm_eq_0 x0 linear_continuous_on)
+    done
   have 3: "(negatex \<circ> sqprojection \<circ> ?F) ` cbox (-1) 1 \<subseteq> cbox (-1) 1"
     unfolding subset_eq
   proof (rule, goal_cases)
@@ -176,29 +153,19 @@
         "x = (negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) y"
       unfolding image_iff ..
     have "?F y \<noteq> 0"
-      apply (rule x0)
-      using y(1)
-      apply auto
-      done
+      by (rule x0) (use y in auto)
     then have *: "infnorm (sqprojection (?F y)) = 1"
       unfolding y o_def
       by - (rule lem2[rule_format])
-    have "infnorm x = 1"
+    have inf1: "infnorm x = 1"
       unfolding *[symmetric] y o_def
       by (rule lem1[rule_format])
-    then show "x \<in> cbox (-1) 1"
+    show "x \<in> cbox (-1) 1"
       unfolding mem_box_cart interval_cbox_cart infnorm_2
-      apply -
-      apply rule
-    proof -
+    proof 
       fix i
-      assume "max \<bar>x $ 1\<bar> \<bar>x $ 2\<bar> = 1"
-      then show "(- 1) $ i \<le> x $ i \<and> x $ i \<le> 1 $ i"
-        apply (cases "i = 1")
-        defer
-        apply (drule 21)
-        apply auto
-        done
+      show "(- 1) $ i \<le> x $ i \<and> x $ i \<le> 1 $ i"
+        using exhaust_2 [of i] inf1 by (auto simp: infnorm_2)
     qed
   qed
   obtain x :: "real^2" where x:
@@ -211,10 +178,7 @@
     apply blast
     done
   have "?F x \<noteq> 0"
-    apply (rule x0)
-    using x(1)
-    apply auto
-    done
+    by (rule x0) (use x in auto)
   then have *: "infnorm (sqprojection (?F x)) = 1"
     unfolding o_def
     by (rule lem2[rule_format])
@@ -223,109 +187,73 @@
     unfolding *[symmetric] o_def
     apply (rule lem1[rule_format])
     done
-  have "\<forall>x i. x \<noteq> 0 \<longrightarrow> (0 < (sqprojection x)$i \<longleftrightarrow> 0 < x$i)"
-    and "\<forall>x i. x \<noteq> 0 \<longrightarrow> ((sqprojection x)$i < 0 \<longleftrightarrow> x$i < 0)"
-    apply -
-    apply (rule_tac[!] allI impI)+
+  have iff: "0 < sqprojection x$i \<longleftrightarrow> 0 < x$i" "sqprojection x$i < 0 \<longleftrightarrow> x$i < 0" if "x \<noteq> 0" for x i
   proof -
-    fix x :: "real^2"
-    fix i :: 2
-    assume x: "x \<noteq> 0"
     have "inverse (infnorm x) > 0"
-      using x[unfolded infnorm_pos_lt[symmetric]] by auto
+      by (simp add: infnorm_pos_lt that)
     then show "(0 < sqprojection x $ i) = (0 < x $ i)"
       and "(sqprojection x $ i < 0) = (x $ i < 0)"
       unfolding sqprojection_def vector_component_simps vector_scaleR_component real_scaleR_def
       unfolding zero_less_mult_iff mult_less_0_iff
       by (auto simp add: field_simps)
   qed
-  note lem3 = this[rule_format]
   have x1: "x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}"
     using x(1) unfolding mem_box_cart by auto
   then have nz: "f (x $ 1) - g (x $ 2) \<noteq> 0"
-    unfolding right_minus_eq
-    apply -
-    apply (rule as)
-    apply auto
-    done
-  have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1"
+    using as by auto
+  consider "x $ 1 = -1" | "x $ 1 = 1" | "x $ 2 = -1" | "x $ 2 = 1"
     using nx unfolding infnorm_eq_1_2 by auto
   then show False
-  proof -
-    fix P Q R S
-    presume "P \<or> Q \<or> R \<or> S"
-      and "P \<Longrightarrow> False"
-      and "Q \<Longrightarrow> False"
-      and "R \<Longrightarrow> False"
-      and "S \<Longrightarrow> False"
-    then show False by auto
-  next
-    assume as: "x$1 = 1"
-    then have *: "f (x $ 1) $ 1 = 1"
-      using assms(6) by auto
-    have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0"
-      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]]
-      unfolding as negatex_def vector_2
-      by auto
-    moreover
-    from x1 have "g (x $ 2) \<in> cbox (-1) 1"
-      using assms(2) by blast
-    ultimately show False
-      unfolding lem3[OF nz] vector_component_simps * mem_box_cart
-      using not_le by auto
-  next
-    assume as: "x$1 = -1"
+  proof cases
+    case 1
     then have *: "f (x $ 1) $ 1 = - 1"
       using assms(5) by auto
     have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0"
       using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]]
-      unfolding as negatex_def vector_2
-      by auto
+      by (auto simp: negatex_def 1)
     moreover
     from x1 have "g (x $ 2) \<in> cbox (-1) 1"
       using assms(2) by blast
     ultimately show False
-      unfolding lem3[OF nz] vector_component_simps * mem_box_cart
+      unfolding iff[OF nz] vector_component_simps * mem_box_cart
       using not_le by auto
   next
-    assume as: "x$2 = 1"
+    case 2
+    then have *: "f (x $ 1) $ 1 = 1"
+      using assms(6) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0"
+      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] 2
+      by (auto simp: negatex_def)
+    moreover have "g (x $ 2) \<in> cbox (-1) 1"
+      using assms(2) x1 by blast
+    ultimately show False
+      unfolding iff[OF nz] vector_component_simps * mem_box_cart
+      using not_le by auto
+  next
+    case 3
+    then have *: "g (x $ 2) $ 2 = - 1"
+      using assms(7) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0"
+      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] 3 by (auto simp: negatex_def)
+    moreover
+    from x1 have "f (x $ 1) \<in> cbox (-1) 1"
+      using assms(1) by blast
+    ultimately show False
+      unfolding iff[OF nz] vector_component_simps * mem_box_cart
+      by (erule_tac x=2 in allE) auto
+  next
+    case 4
     then have *: "g (x $ 2) $ 2 = 1"
       using assms(8) by auto
     have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0"
-      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]]
-      unfolding as negatex_def vector_2
-      by auto
+      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] 4 by (auto simp: negatex_def)
     moreover
     from x1 have "f (x $ 1) \<in> cbox (-1) 1"
-      apply -
-      apply (rule assms(1)[unfolded subset_eq,rule_format])
-      apply auto
-      done
+      using assms(1) by blast
     ultimately show False
-      unfolding lem3[OF nz] vector_component_simps * mem_box_cart
-      apply (erule_tac x=2 in allE)
-      apply auto
-      done
-  next
-    assume as: "x$2 = -1"
-    then have *: "g (x $ 2) $ 2 = - 1"
-      using assms(7) by auto
-    have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0"
-      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]]
-      unfolding as negatex_def vector_2
-      by auto
-    moreover
-    from x1 have "f (x $ 1) \<in> cbox (-1) 1"
-      apply -
-      apply (rule assms(1)[unfolded subset_eq,rule_format])
-      apply auto
-      done
-    ultimately show False
-      unfolding lem3[OF nz] vector_component_simps * mem_box_cart
-      apply (erule_tac x=2 in allE)
-      apply auto
-      done
-  qed auto
+      unfolding iff[OF nz] vector_component_simps * mem_box_cart
+      by (erule_tac x=2 in allE) auto
+  qed 
 qed
 
 lemma fashoda_unit_path:
@@ -533,14 +461,12 @@
     using as
     by auto
   show thesis
-    apply (rule_tac z="interval_bij (- 1,1) (a,b) z" in that)
-    apply (subst zf)
-    defer
-    apply (subst zg)
-    unfolding o_def interval_bij_bij_cart[OF *] path_image_def
-    using zf(1) zg(1)
-    apply auto
-    done
+  proof (rule_tac z="interval_bij (- 1,1) (a,b) z" in that)
+    show "interval_bij (- 1, 1) (a, b) z \<in> path_image f"
+      using zf by (simp add: interval_bij_bij_cart[OF *] path_image_def)
+    show "interval_bij (- 1, 1) (a, b) z \<in> path_image g"
+      using zg by (simp add: interval_bij_bij_cart[OF *] path_image_def)
+  qed
 qed
 
 
@@ -714,8 +640,8 @@
   fixes a :: "real^2"
   assumes "path f"
     and "path g"
-    and "path_image f \<subseteq> cbox a b"
-    and "path_image g \<subseteq> cbox a b"
+    and paf: "path_image f \<subseteq> cbox a b"
+    and pag: "path_image g \<subseteq> cbox a b"
     and "(pathstart f)$2 = a$2"
     and "(pathfinish f)$2 = a$2"
     and "(pathstart g)$2 = a$2"
@@ -776,30 +702,9 @@
   proof -
     show "path ?P1" and "path ?P2"
       using assms by auto
-    have "path_image ?P1 \<subseteq> cbox ?a ?b"
-      unfolding P1P2 path_image_linepath
-      apply (rule Un_least)+
-      defer 3
-      apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format])
-      unfolding mem_box_cart forall_2 vector_2
-      using ab startfin abab assms(3)
-      using assms(9-)
-      unfolding assms
-      apply (auto simp add: field_simps box_def)
-      done
-    then show "path_image ?P1 \<subseteq> cbox ?a ?b" .
-    have "path_image ?P2 \<subseteq> cbox ?a ?b"
-      unfolding P1P2 path_image_linepath
-      apply (rule Un_least)+
-      defer 2
-      apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format])
-      unfolding mem_box_cart forall_2 vector_2
-      using ab startfin abab assms(4)
-      using assms(9-)
-      unfolding assms
-      apply (auto simp add: field_simps box_def)
-      done
-    then show "path_image ?P2 \<subseteq> cbox ?a ?b" .
+    show "path_image ?P1 \<subseteq> cbox ?a ?b" "path_image ?P2 \<subseteq> cbox ?a ?b"
+      unfolding P1P2 path_image_linepath using startfin paf pag
+      by (auto simp: mem_box_cart segment_horizontal segment_vertical forall_2)
     show "a $ 1 - 2 = a $ 1 - 2"
       and "b $ 1 + 2 = b $ 1 + 2"
       and "pathstart g $ 2 - 3 = a $ 2 - 3"
@@ -808,8 +713,7 @@
   qed
   note z=this[unfolded P1P2 path_image_linepath]
   show thesis
-    apply (rule that[of z])
-  proof -
+  proof (rule that[of z])
     have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or>
       z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or>
       z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or>
@@ -838,32 +742,26 @@
         using prems(2) using assms ab
         by (auto simp add: field_simps)
       ultimately have *: "z$2 = a$2 - 2"
-        using prems(1)
-        by auto
+        using prems(1) by auto
       have "z$1 \<noteq> pathfinish g$1"
-        using prems(2)
-        using assms ab
+        using prems(2) assms ab
         by (auto simp add: field_simps *)
       moreover have "pathstart g \<in> cbox a b"
         using assms(4) pathstart_in_path_image[of g]
         by auto
       note this[unfolded mem_box_cart forall_2]
       then have "z$1 \<noteq> pathstart g$1"
-        using prems(1)
-        using assms ab
+        using prems(1) assms ab
         by (auto simp add: field_simps *)
       ultimately have "a $ 2 - 1 \<le> z $ 2 \<and> z $ 2 \<le> b $ 2 + 3 \<or> b $ 2 + 3 \<le> z $ 2 \<and> z $ 2 \<le> a $ 2 - 1"
-        using prems(2)
-        unfolding * assms
-        by (auto simp add: field_simps)
+        using prems(2)  unfolding * assms by (auto simp add: field_simps)
       then show False
         unfolding * using ab by auto
     qed
     then have "z \<in> path_image f \<or> z \<in> path_image g"
       using z unfolding Un_iff by blast
     then have z': "z \<in> cbox a b"
-      using assms(3-4)
-      by auto
+      using assms(3-4) by auto
     have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow>
       z = pathstart f \<or> z = pathfinish f"
       unfolding vec_eq_iff forall_2 assms
@@ -871,11 +769,7 @@
     with z' show "z \<in> path_image f"
       using z(1)
       unfolding Un_iff mem_box_cart forall_2
-      apply -
-      apply (simp only: segment_vertical segment_horizontal vector_2)
-      unfolding assms
-      apply auto
-      done
+      by (simp only: segment_vertical segment_horizontal vector_2) (auto simp: assms)
     have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow>
       z = pathstart g \<or> z = pathfinish g"
       unfolding vec_eq_iff forall_2 assms
@@ -883,244 +777,8 @@
     with z' show "z \<in> path_image g"
       using z(2)
       unfolding Un_iff mem_box_cart forall_2
-      apply (simp only: segment_vertical segment_horizontal vector_2)
-      unfolding assms
-      apply auto
-      done
+      by (simp only: segment_vertical segment_horizontal vector_2) (auto simp: assms)
   qed
 qed
 
-(** The Following still needs to be translated. Maybe I will do that later.
-
-(* ------------------------------------------------------------------------- *)
-(* Complement in dimension N >= 2 of set homeomorphic to any interval in     *)
-(* any dimension is (path-)connected. This naively generalizes the argument  *)
-(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer        *)
-(* fixed point theorem", American Mathematical Monthly 1984.                 *)
-(* ------------------------------------------------------------------------- *)
-
-let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove
- (`!p:real^M->real^N a b.
-        ~(interval[a,b] = {}) /\
-        p continuous_on interval[a,b] /\
-        (!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y)
-        ==> ?f. f continuous_on (:real^N) /\
-                IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\
-                (!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`,
-  REPEAT STRIP_TAC THEN
-  FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN
-  DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN
-  SUBGOAL_THEN `(q:real^N->real^M) continuous_on
-                (IMAGE p (interval[a:real^M,b]))`
-  ASSUME_TAC THENL
-   [MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL];
-    ALL_TAC] THEN
-  MP_TAC(ISPECL [`q:real^N->real^M`;
-                 `IMAGE (p:real^M->real^N)
-                 (interval[a,b])`;
-                 `a:real^M`; `b:real^M`]
-        TIETZE_CLOSED_INTERVAL) THEN
-  ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE;
-               COMPACT_IMP_CLOSED] THEN
-  ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
-  DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN
-  EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN
-  REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN
-  CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN
-  MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN
-  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ]
-        CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);;
-
-let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove
- (`!s:real^N->bool a b:real^M.
-        s homeomorphic (interval[a,b])
-        ==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`,
-  REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN
-  REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
-  MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN
-  DISCH_TAC THEN
-  SUBGOAL_THEN
-   `!x y. x IN interval[a,b] /\ y IN interval[a,b] /\
-          (p:real^M->real^N) x = p y ==> x = y`
-  ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN
-  FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN
-  DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN
-  ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN
-  ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV;
-                  NOT_BOUNDED_UNIV] THEN
-  ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN
-  X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN
-  SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
-  SUBGOAL_THEN `bounded((path_component s c) UNION
-                        (IMAGE (p:real^M->real^N) (interval[a,b])))`
-  MP_TAC THENL
-   [ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED;
-                 COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
-    ALL_TAC] THEN
-  DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN
-  REWRITE_TAC[UNION_SUBSET] THEN
-  DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN
-  MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`]
-    RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN
-  ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN
-  DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN
-  DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC
-   (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN
-  REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN
-  ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN
-  SUBGOAL_THEN
-    `(q:real^N->real^N) continuous_on
-     (closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))`
-  MP_TAC THENL
-   [EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN
-    REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN
-    REPEAT CONJ_TAC THENL
-     [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN
-      ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;
-                   COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
-      ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV];
-      ALL_TAC] THEN
-    X_GEN_TAC `z:real^N` THEN
-    REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN
-    STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN
-    MP_TAC(ISPECL
-     [`path_component s (z:real^N)`; `path_component s (c:real^N)`]
-     OPEN_INTER_CLOSURE_EQ_EMPTY) THEN
-    ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL
-     [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN
-      ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;
-                   COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
-      REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN
-      DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN
-      GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN
-      REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]];
-    ALL_TAC] THEN
-  SUBGOAL_THEN
-   `closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) =
-    (:real^N)`
-  SUBST1_TAC THENL
-   [MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN
-    REWRITE_TAC[CLOSURE_SUBSET];
-    DISCH_TAC] THEN
-  MP_TAC(ISPECL
-   [`(\x. &2 % c - x) o
-     (\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`;
-    `cball(c:real^N,B)`]
-    BROUWER) THEN
-  REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN
-  ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN
-  SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL
-   [X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN
-    REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN
-    ASM SET_TAC[PATH_COMPONENT_REFL_EQ];
-    ALL_TAC] THEN
-  REPEAT CONJ_TAC THENL
-   [MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN
-    SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL
-     [ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_MUL THEN
-    SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN
-    REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN
-    MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN
-    ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN
-    SUBGOAL_THEN
-     `(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)`
-    SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN
-    ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST;
-                 CONTINUOUS_ON_LIFT_NORM];
-    REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN
-    X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
-    REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN
-    REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN
-    ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN
-    ASM_REAL_ARITH_TAC;
-    REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN
-    REWRITE_TAC[IN_CBALL; o_THM; dist] THEN
-    X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
-    REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN
-    ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL
-     [MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN
-      REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN
-      ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN
-      ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN
-      UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN
-      REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB];
-      EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN
-      REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN
-      ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN
-      SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL
-       [ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN
-      ASM_REWRITE_TAC[] THEN
-      MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN
-      ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);;
-
-let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove
- (`!s:real^N->bool a b:real^M.
-        2 <= dimindex(:N) /\ s homeomorphic interval[a,b]
-        ==> path_connected((:real^N) DIFF s)`,
-  REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN
-  FIRST_ASSUM(MP_TAC o MATCH_MP
-    UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN
-  ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN
-  ABBREV_TAC `t = (:real^N) DIFF s` THEN
-  DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN
-  STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN
-  REWRITE_TAC[COMPACT_INTERVAL] THEN
-  DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN
-  REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN
-  X_GEN_TAC `B:real` THEN STRIP_TAC THEN
-  SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\
-                (?v:real^N. v IN path_component t y /\ B < norm(v))`
-  STRIP_ASSUME_TAC THENL
-   [ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN
-  MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN
-  CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN
-  MATCH_MP_TAC PATH_COMPONENT_SYM THEN
-  MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN
-  CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN
-  MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN
-  EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL
-   [EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE
-     `s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN
-    ASM_REWRITE_TAC[SUBSET; IN_CBALL_0];
-    MP_TAC(ISPEC `cball(vec 0:real^N,B)`
-       PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN
-    ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN
-    REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN
-    DISCH_THEN MATCH_MP_TAC THEN
-    ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);;
-
-(* ------------------------------------------------------------------------- *)
-(* In particular, apply all these to the special case of an arc.             *)
-(* ------------------------------------------------------------------------- *)
-
-let RETRACTION_ARC = prove
- (`!p. arc p
-       ==> ?f. f continuous_on (:real^N) /\
-               IMAGE f (:real^N) SUBSET path_image p /\
-               (!x. x IN path_image p ==> f x = x)`,
-  REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN
-  MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN
-  ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);;
-
-let PATH_CONNECTED_ARC_COMPLEMENT = prove
- (`!p. 2 <= dimindex(:N) /\ arc p
-       ==> path_connected((:real^N) DIFF path_image p)`,
-  REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN
-  MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`]
-    PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN
-  ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN
-  ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN
-  MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN
-  EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);;
-
-let CONNECTED_ARC_COMPLEMENT = prove
- (`!p. 2 <= dimindex(:N) /\ arc p
-       ==> connected((:real^N) DIFF path_image p)`,
-  SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *)
-
 end
--- a/src/HOL/Analysis/Finite_Cartesian_Product.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy	Thu May 03 15:07:14 2018 +0200
@@ -623,8 +623,6 @@
   unfolding norm_vec_def
   by (rule L2_set_mono) (auto simp: assms)
 
-lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero)
-
 lemma component_le_norm_cart: "\<bar>x$i\<bar> \<le> norm x"
   apply (simp add: norm_vec_def)
   apply (rule member_le_L2_set, simp_all)
@@ -974,12 +972,6 @@
 lemma vector_mul_rcancel[simp]: "a *s x = b *s x \<longleftrightarrow> (a::'a::field) = b \<or> x = 0"
   by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib)
 
-lemma vector_mul_lcancel_imp: "a \<noteq> (0::'a::field) ==>  a *s x = a *s y ==> (x = y)"
-  by (metis vector_mul_lcancel)
-
-lemma vector_mul_rcancel_imp: "x \<noteq> 0 \<Longrightarrow> (a::'a::field) *s x = b *s x ==> a = b"
-  by (metis vector_mul_rcancel)
-
 lemma scalar_mult_eq_scaleR [abs_def]: "c *s x = c *\<^sub>R x"
   unfolding scaleR_vec_def vector_scalar_mult_def by simp
 
@@ -1041,6 +1033,12 @@
 definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
 definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
 
+lemma times0_left [simp]: "(0::'a::semiring_1^'n^'m) ** (A::'a ^'p^'n) = 0" 
+  by (simp add: matrix_matrix_mult_def zero_vec_def)
+
+lemma times0_right [simp]: "(A::'a::semiring_1^'n^'m) ** (0::'a ^'p^'n) = 0" 
+  by (simp add: matrix_matrix_mult_def zero_vec_def)
+
 lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
 lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
   by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps)
@@ -1076,6 +1074,16 @@
   apply simp
   done
 
+lemma scalar_matrix_assoc:
+  fixes A :: "('a::real_algebra_1)^'m^'n"
+  shows "k *\<^sub>R (A ** B) = (k *\<^sub>R A) ** B"
+  by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff scaleR_sum_right)
+
+lemma matrix_scalar_ac:
+  fixes A :: "('a::real_algebra_1)^'m^'n"
+  shows "A ** (k *\<^sub>R B) = k *\<^sub>R A ** B"
+  by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff)
+
 lemma matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"
   apply (vector matrix_vector_mult_def mat_def)
   apply (simp add: if_distrib if_distribR sum.delta' cong del: if_weak_cong)
@@ -1098,7 +1106,7 @@
     sum.delta[OF finite] cong del: if_weak_cong)
   done
 
-lemma matrix_vector_mul_component: "((A::real^_^_) *v x)$k = inner (A$k) x"
+lemma matrix_vector_mul_component: "(A *v x)$k = inner (A$k) x"
   by (simp add: matrix_vector_mult_def inner_vec_def)
 
 lemma dot_lmul_matrix: "inner ((x::real ^_) v* A) y = inner x (A *v y)"
@@ -1125,6 +1133,13 @@
 lemma columns_transpose [simp]: "columns(transpose A) = rows A"
   by (metis transpose_transpose rows_transpose)
 
+lemma transpose_scalar: "transpose (k *\<^sub>R A) = k *\<^sub>R transpose A"
+  unfolding transpose_def
+  by (simp add: vec_eq_iff)
+
+lemma transpose_iff [iff]: "transpose A = transpose B \<longleftrightarrow> A = B"
+  by (metis transpose_transpose)
+
 lemma matrix_mult_sum:
   "(A::'a::comm_semiring_1^'n^'m) *v x = sum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
   by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute)
@@ -1148,11 +1163,15 @@
 lemma matrix_scaleR: "(matrix (( *\<^sub>R) r)) = mat r"
   by (simp add: mat_def matrix_def axis_def if_distrib cong: if_cong)
 
-lemma matrix_vector_mul_linear[intro, simp]: "linear (\<lambda>x. A *v (x::real ^ _))"
-  by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff
-      field_simps sum_distrib_left sum.distrib)
+lemma matrix_vector_mul_linear[intro, simp]: "linear (\<lambda>x. A *v (x::'a::real_algebra_1 ^ _))"
+  by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff field_simps sum_distrib_left sum.distrib)
 
-lemma matrix_vector_mult_add_distrib [algebra_simps]:
+lemma vector_matrix_left_distrib [algebra_simps]:
+  shows "(x + y) v* A = x v* A + y v* A"
+  unfolding vector_matrix_mult_def
+  by (simp add: algebra_simps sum.distrib vec_eq_iff)
+
+lemma matrix_vector_right_distrib [algebra_simps]:
   "A *v (x + y) = A *v x + A *v y"
   by (vector matrix_vector_mult_def sum.distrib distrib_left)
 
@@ -1200,4 +1219,70 @@
   shows "inj (( *v) A)"
   by (metis assms inj_on_inverseI invertible_def matrix_vector_mul_assoc matrix_vector_mul_lid)
 
+lemma scalar_invertible:
+  fixes A :: "('a::real_algebra_1)^'m^'n"
+  assumes "k \<noteq> 0" and "invertible A"
+  shows "invertible (k *\<^sub>R A)"
+proof -
+  obtain A' where "A ** A' = mat 1" and "A' ** A = mat 1"
+    using assms unfolding invertible_def by auto
+  with `k \<noteq> 0`
+  have "(k *\<^sub>R A) ** ((1/k) *\<^sub>R A') = mat 1" "((1/k) *\<^sub>R A') ** (k *\<^sub>R A) = mat 1"
+    by (simp_all add: assms matrix_scalar_ac)
+  thus "invertible (k *\<^sub>R A)"
+    unfolding invertible_def by auto
+qed
+
+lemma scalar_invertible_iff:
+  fixes A :: "('a::real_algebra_1)^'m^'n"
+  assumes "k \<noteq> 0" and "invertible A"
+  shows "invertible (k *\<^sub>R A) \<longleftrightarrow> k \<noteq> 0 \<and> invertible A"
+  by (simp add: assms scalar_invertible)
+
+lemma vector_transpose_matrix [simp]: "x v* transpose A = A *v x"
+  unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
+  by simp
+
+lemma transpose_matrix_vector [simp]: "transpose A *v x = x v* A"
+  unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
+  by simp
+
+lemma vector_scalar_commute:
+  fixes A :: "'a::{field}^'m^'n"
+  shows "A *v (c *s x) = c *s (A *v x)"
+  by (simp add: vector_scalar_mult_def matrix_vector_mult_def mult_ac sum_distrib_left)
+
+lemma scalar_vector_matrix_assoc:
+  fixes k :: "'a::{field}" and x :: "'a::{field}^'n" and A :: "'a^'m^'n"
+  shows "(k *s x) v* A = k *s (x v* A)"
+  by (metis transpose_matrix_vector vector_scalar_commute)
+ 
+lemma vector_matrix_mult_0 [simp]: "0 v* A = 0"
+  unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
+
+lemma vector_matrix_mult_0_right [simp]: "x v* 0 = 0"
+  unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
+
+lemma vector_matrix_mul_rid [simp]:
+  fixes v :: "('a::semiring_1)^'n"
+  shows "v v* mat 1 = v"
+  by (metis matrix_vector_mul_lid transpose_mat vector_transpose_matrix)
+
+lemma scaleR_vector_matrix_assoc:
+  fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
+  shows "(k *\<^sub>R x) v* A = k *\<^sub>R (x v* A)"
+  by (metis matrix_vector_mult_scaleR transpose_matrix_vector)
+
+lemma vector_scaleR_matrix_ac:
+  fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
+  shows "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)"
+proof -
+  have "x v* (k *\<^sub>R A) = (k *\<^sub>R x) v* A"
+    unfolding vector_matrix_mult_def
+    by (simp add: algebra_simps)
+  with scaleR_vector_matrix_assoc
+  show "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)"
+    by auto
+qed
+
 end
--- a/src/HOL/Analysis/Gamma_Function.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Gamma_Function.thy	Thu May 03 15:07:14 2018 +0200
@@ -1149,7 +1149,7 @@
 
 lemma Gamma_fact: "Gamma (1 + of_nat n) = fact n"
   by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff
-        of_nat_Suc [symmetric] del: of_nat_Suc)
+        reorient: of_nat_Suc)
 
 lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)"
   by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc,
@@ -2430,7 +2430,7 @@
           inverse ((- 1) ^ n * fact n :: 'a)"
     by (intro tendsto_intros rGamma_zeros) simp_all
   also have "inverse ((- 1) ^ n * fact n) = ?c"
-    by (simp_all add: field_simps power_mult_distrib [symmetric] del: power_mult_distrib)
+    by (simp_all add: field_simps reorient: power_mult_distrib)
   finally show "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow> ?c" .
 qed
 
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Thu May 03 15:07:14 2018 +0200
@@ -128,20 +128,13 @@
   assumes "k \<in> Basis"
   shows "content (cbox a b) = content(cbox a b \<inter> {x. x\<bullet>k \<le> c}) + content(cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
   \<comment> \<open>Prove using measure theory\<close>
-proof cases
+proof (cases "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i")
+  case True
+  have 1: "\<And>X Y Z. (\<Prod>i\<in>Basis. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>Basis-{k}. Z i (Y i))"
+    by (simp add: if_distrib prod.delta_remove assms)
   note simps = interval_split[OF assms] content_cbox_cases
-  have *: "Basis = insert k (Basis - {k})" "\<And>x. finite (Basis-{x})" "\<And>x. x\<notin>Basis-{x}"
-    using assms by auto
-  have *: "\<And>X Y Z. (\<Prod>i\<in>Basis. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>Basis-{k}. Z i (Y i))"
-    "(\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i) = (\<Prod>i\<in>Basis-{k}. b\<bullet>i - a\<bullet>i) * (b\<bullet>k - a\<bullet>k)"
-    apply (subst *(1))
-    defer
-    apply (subst *(1))
-    unfolding prod.insert[OF *(2-)]
-    apply auto
-    done
-  assume as: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
-  moreover
+  have 2: "(\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i) = (\<Prod>i\<in>Basis-{k}. b\<bullet>i - a\<bullet>i) * (b\<bullet>k - a\<bullet>k)"
+    by (metis (no_types, lifting) assms finite_Basis mult.commute prod.remove)
   have "\<And>x. min (b \<bullet> k) c = max (a \<bullet> k) c \<Longrightarrow>
     x * (b\<bullet>k - a\<bullet>k) = x * (max (a \<bullet> k) c - a \<bullet> k) + x * (b \<bullet> k - max (a \<bullet> k) c)"
     by  (auto simp add: field_simps)
@@ -152,17 +145,12 @@
       (\<Prod>i\<in>Basis. b \<bullet> i - (if i = k then max (a \<bullet> k) c else a \<bullet> i))"
     by (auto intro!: prod.cong)
   have "\<not> a \<bullet> k \<le> c \<Longrightarrow> \<not> c \<le> b \<bullet> k \<Longrightarrow> False"
-    unfolding not_le
-    using as[unfolded ,rule_format,of k] assms
-    by auto
+    unfolding not_le using True assms by auto
   ultimately show ?thesis
-    using assms
-    unfolding simps **
-    unfolding *(1)[of "\<lambda>i x. b\<bullet>i - x"] *(1)[of "\<lambda>i x. x - a\<bullet>i"]
-    unfolding *(2)
+    using assms unfolding simps ** 1[of "\<lambda>i x. b\<bullet>i - x"] 1[of "\<lambda>i x. x - a\<bullet>i"] 2
     by auto
 next
-  assume "\<not> (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)"
+  case False
   then have "cbox a b = {}"
     unfolding box_eq_empty by (auto simp: not_le)
   then show ?thesis
@@ -784,13 +772,13 @@
   fixes f :: "_ \<Rightarrow> 'b :: real_normed_field"
   assumes "f integrable_on s"
   shows "(\<lambda>x. f x / of_real c) integrable_on s"
-by (simp add: integrable_on_cmult_right divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse)
+by (simp add: integrable_on_cmult_right divide_inverse assms reorient: of_real_inverse)
 
 lemma integrable_on_cdivide_iff [simp]:
   fixes f :: "_ \<Rightarrow> 'b :: real_normed_field"
   assumes "c \<noteq> 0"
   shows "(\<lambda>x. f x / of_real c) integrable_on s \<longleftrightarrow> f integrable_on s"
-by (simp add: divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse)
+by (simp add: divide_inverse assms reorient: of_real_inverse)
 
 lemma has_integral_null [intro]: "content(cbox a b) = 0 \<Longrightarrow> (f has_integral 0) (cbox a b)"
   unfolding has_integral_cbox
@@ -2585,24 +2573,22 @@
       norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content (cbox a b)))"
 proof (cases "content (cbox a b) = 0")
   case True
-  show ?thesis
+  have "\<And>e p. p tagged_division_of cbox a b \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)) \<le> e * content (cbox a b)"
+    unfolding sum_content_null[OF True] True by force
+  moreover have "i = 0" 
+    if "\<And>e. e > 0 \<Longrightarrow> \<exists>d. gauge d \<and>
+              (\<forall>p. p tagged_division_of cbox a b \<and>
+                   d fine p \<longrightarrow>
+                   norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) \<le> e * content (cbox a b))"
+    using that [of 1]
+    by (force simp add: True sum_content_null[OF True] intro: fine_division_exists[of _ a b])
+  ultimately show ?thesis
     unfolding has_integral_null_eq[OF True]
-    apply safe
-    apply (rule, rule, rule gauge_trivial, safe)
-    unfolding sum_content_null[OF True] True
-    defer
-    apply (erule_tac x=1 in allE)
-    apply safe
-    defer
-    apply (rule fine_division_exists[of _ a b])
-    apply assumption
-    apply (erule_tac x=p in allE)
-    unfolding sum_content_null[OF True]
-    apply auto
-    done
+    by (force simp add: )
 next
   case False
-  note F = this[unfolded content_lt_nz[symmetric]]
+  then have F: "0 < content (cbox a b)"
+    using zero_less_measure_iff by blast
   let ?P = "\<lambda>e opp. \<exists>d. gauge d \<and>
     (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> opp (norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i)) e)"
   show ?thesis
@@ -2610,28 +2596,18 @@
   proof safe
     fix e :: real
     assume e: "e > 0"
-    {
-      assume "\<forall>e>0. ?P e (<)"
+    { assume "\<forall>e>0. ?P e (<)"
       then show "?P (e * content (cbox a b)) (\<le>)"
-        apply (erule_tac x="e * content (cbox a b)" in allE)
-        apply (erule impE)
-        defer
-        apply (erule exE,rule_tac x=d in exI)
-        using F e
-        apply (auto simp add:field_simps)
-        done
-    }
-    {
-      assume "\<forall>e>0. ?P (e * content (cbox a b)) (\<le>)"
+        apply (rule allE [where x="e * content (cbox a b)"])
+        apply (elim impE ex_forward conj_forward)
+        using F e apply (auto simp add: algebra_simps)
+        done }
+    { assume "\<forall>e>0. ?P (e * content (cbox a b)) (\<le>)"
       then show "?P e (<)"
-        apply (erule_tac x="e/2 / content (cbox a b)" in allE)
-        apply (erule impE)
-        defer
-        apply (erule exE,rule_tac x=d in exI)
-        using F e
-        apply (auto simp add: field_simps)
-        done
-    }
+        apply (rule allE [where x="e/2 / content (cbox a b)"])
+        apply (elim impE ex_forward conj_forward)
+        using F e apply (auto simp add: algebra_simps)
+        done }
   qed
 qed
 
@@ -2995,19 +2971,22 @@
 
 lemma integrable_subinterval:
   fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on cbox a b"
-    and "cbox c d \<subseteq> cbox a b"
+  assumes f: "f integrable_on cbox a b"
+    and cd: "cbox c d \<subseteq> cbox a b"
   shows "f integrable_on cbox c d"
 proof -
   interpret operative conj True "\<lambda>i. f integrable_on i"
     using order_refl by (rule operative_integrableI)
   show ?thesis
-    apply (cases "cbox c d = {}")
-     defer
-     apply (rule partial_division_extend_1[OF assms(2)],assumption)
-    using division [symmetric] assms(1)
-     apply (auto simp: comm_monoid_set_F_and)
-    done
+  proof (cases "cbox c d = {}")
+    case True
+    then show ?thesis
+      using division [symmetric] f by (auto simp: comm_monoid_set_F_and)
+  next
+    case False
+    then show ?thesis
+      by (metis cd comm_monoid_set_F_and division division_of_finite f partial_division_extend_1)
+  qed
 qed
 
 lemma integrable_subinterval_real:
@@ -3431,7 +3410,7 @@
       by (simp add: inner_simps field_simps)
     ultimately show ?thesis using False
       by (simp add: image_affinity_cbox content_cbox'
-        prod.distrib[symmetric] prod_constant[symmetric] inner_diff_left del: prod_constant)
+        prod.distrib[symmetric] inner_diff_left reorient: prod_constant)
   qed
 qed
 
@@ -4261,31 +4240,29 @@
 lemma has_derivative_zero_unique_strong_interval:
   fixes f :: "real \<Rightarrow> 'a::banach"
   assumes "finite k"
-    and "continuous_on {a..b} f"
+    and contf: "continuous_on {a..b} f"
     and "f a = y"
-    and "\<forall>x\<in>({a..b} - k). (f has_derivative (\<lambda>h. 0)) (at x within {a..b})" "x \<in> {a..b}"
+    and fder: "\<And>x. x \<in> {a..b} - k \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within {a..b})"
+    and x: "x \<in> {a..b}"
   shows "f x = y"
 proof -
-  have ab: "a \<le> b"
+  have "a \<le> b" "a \<le> x"
     using assms by auto
-  have *: "a \<le> x"
-    using assms(5) by auto
   have "((\<lambda>x. 0::'a) has_integral f x - f a) {a..x}"
-    apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *])
-    apply (rule continuous_on_subset[OF assms(2)])
-    defer
-    apply safe
-    unfolding has_vector_derivative_def
-    apply (subst has_derivative_within_open[symmetric])
-    apply assumption
-    apply (rule open_greaterThanLessThan)
-    apply (rule has_derivative_within_subset[where s="{a..b}"])
-    using assms(4) assms(5)
-    apply (auto simp: mem_box)
-    done
-  note this[unfolded *]
-  note has_integral_unique[OF has_integral_0 this]
-  then show ?thesis
+  proof (rule fundamental_theorem_of_calculus_interior_strong[OF \<open>finite k\<close> \<open>a \<le> x\<close>]; clarify?)
+    have "{a..x} \<subseteq> {a..b}"
+      using x by auto
+    then show "continuous_on {a..x} f"
+      by (rule continuous_on_subset[OF contf])
+    show "(f has_vector_derivative 0) (at z)" if z: "z \<in> {a<..<x}" and notin: "z \<notin> k" for z
+      unfolding has_vector_derivative_def
+    proof (simp add: has_derivative_within_open[OF z, symmetric])
+      show "(f has_derivative (\<lambda>x. 0)) (at z within {a<..<x})"
+        by (rule has_derivative_within_subset [OF fder]) (use x z notin in auto)
+    qed
+  qed
+  from has_integral_unique[OF has_integral_0 this]
+  show ?thesis
     unfolding assms by auto
 qed
 
@@ -4297,7 +4274,7 @@
   assumes "convex S" "finite K"
     and contf: "continuous_on S f"
     and "c \<in> S" "f c = y"
-    and derf: "\<And>x. x \<in> (S - K) \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within S)"
+    and derf: "\<And>x. x \<in> S - K \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within S)"
     and "x \<in> S"
   shows "f x = y"
 proof (cases "x = c")
@@ -4307,8 +4284,10 @@
   case False
   let ?\<phi> = "\<lambda>u. (1 - u) *\<^sub>R c + u *\<^sub>R x"
   have contf': "continuous_on {0 ..1} (f \<circ> ?\<phi>)"
-    apply (rule continuous_intros continuous_on_subset[OF contf])+
-    using \<open>convex S\<close> \<open>x \<in> S\<close> \<open>c \<in> S\<close> by (auto simp add: convex_alt algebra_simps)
+  proof (rule continuous_intros continuous_on_subset[OF contf])+
+    show "(\<lambda>u. (1 - u) *\<^sub>R c + u *\<^sub>R x) ` {0..1} \<subseteq> S"
+      using \<open>convex S\<close> \<open>x \<in> S\<close> \<open>c \<in> S\<close> by (auto simp add: convex_alt algebra_simps)
+  qed
   have "t = u" if "?\<phi> t = ?\<phi> u" for t u
   proof -
     from that have "(t - u) *\<^sub>R x = (t - u) *\<^sub>R c"
@@ -4353,7 +4332,7 @@
     and contf: "continuous_on S f"
     and "c \<in> S"
     and "f c = y"
-    and derf: "\<And>x. x \<in> (S - K) \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within S)"
+    and derf: "\<And>x. x \<in> S - K \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within S)"
     and "x \<in> S"
   shows "f x = y"
 proof -
@@ -4486,10 +4465,7 @@
   note has_integral_restrict_open_subinterval[OF assms]
   note * = has_integral_spike[OF negligible_frontier_interval _ this]
   show ?thesis
-    apply (rule *[of c d])
-    using box_subset_cbox[of c d]
-    apply auto
-    done
+    by (rule *[of c d]) (use box_subset_cbox[of c d] in auto)
 qed
 
 lemma has_integral_restrict_closed_subintervals_eq:
--- a/src/HOL/Analysis/Homeomorphism.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Homeomorphism.thy	Thu May 03 15:07:14 2018 +0200
@@ -881,9 +881,8 @@
     using aff_dim_convex_Int_open [OF \<open>convex U\<close> open_ball]
     by (fastforce simp add: Int_commute)
   have "rel_frontier S homeomorphic rel_frontier (ball z 1 \<inter> U)"
-    apply (rule homeomorphic_rel_frontiers_convex_bounded_sets)
-    apply (auto simp: \<open>affine U\<close> affine_imp_convex convex_Int affdS assms)
-    done
+    by (rule homeomorphic_rel_frontiers_convex_bounded_sets)
+       (auto simp: \<open>affine U\<close> affine_imp_convex convex_Int affdS assms)
   also have "... = sphere z 1 \<inter> U"
     using convex_affine_rel_frontier_Int [of "ball z 1" U]
     by (simp add: \<open>affine U\<close> bne)
@@ -903,9 +902,8 @@
       by (force simp: h [symmetric] image_comp o_def kh)
   qed (auto intro: continuous_on_subset hcon kcon simp: kh hk)
   also have "... homeomorphic T"
-    apply (rule homeomorphic_punctured_affine_sphere_affine)
-    using a him
-    by (auto simp: affS affdS \<open>affine T\<close> \<open>affine U\<close> \<open>z \<in> U\<close>)
+    by (rule homeomorphic_punctured_affine_sphere_affine)
+       (use a him in \<open>auto simp: affS affdS \<open>affine T\<close> \<open>affine U\<close> \<open>z \<in> U\<close>\<close>)
   finally show ?thesis .
 qed
 
@@ -965,7 +963,7 @@
   then obtain f g where fg: "homeomorphism (sphere 0 1 - {i}) {x. i \<bullet> x = 0} f g"
     by (force simp: homeomorphic_def)
   have "h ` (+) (- a) ` S \<subseteq> T"
-    using heq span_clauses(1) span_linear_image by blast
+    using heq span_superset span_linear_image by blast
   then have "g ` h ` (+) (- a) ` S \<subseteq> g ` {x. i \<bullet> x = 0}"
     using Tsub by (simp add: image_mono)
   also have "... \<subseteq> sphere 0 1 - {i}"
@@ -989,8 +987,8 @@
     apply (simp add: homeomorphic_def homeomorphism_def)
     apply (rule_tac x="g \<circ> h" in exI)
     apply (rule_tac x="k \<circ> f" in exI)
-    apply (auto simp: ghcont kfcont span_clauses(1) homeomorphism_apply2 [OF fg] image_comp)
-    apply (force simp: o_def homeomorphism_apply2 [OF fg] span_clauses(1))
+    apply (auto simp: ghcont kfcont span_superset homeomorphism_apply2 [OF fg] image_comp)
+    apply (force simp: o_def homeomorphism_apply2 [OF fg] span_superset)
     done
   finally have Shom: "S homeomorphic (g \<circ> h) ` (+) (- a) ` S" .
   show ?thesis
--- a/src/HOL/Analysis/Infinite_Products.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Infinite_Products.thy	Thu May 03 15:07:14 2018 +0200
@@ -1,6 +1,5 @@
-(*
-  File:      HOL/Analysis/Infinite_Product.thy
-  Author:    Manuel Eberl
+(*File:      HOL/Analysis/Infinite_Product.thy
+  Author:    Manuel Eberl & LC Paulson
 
   Basic results about convergence and absolute convergence of infinite products
   and their connection to summability.
@@ -9,7 +8,7 @@
 theory Infinite_Products
   imports Complex_Main
 begin
-
+    
 lemma sum_le_prod:
   fixes f :: "'a \<Rightarrow> 'b :: linordered_semidom"
   assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0"
@@ -51,8 +50,27 @@
     by (rule tendsto_eq_intros refl | simp)+
 qed auto
 
+definition gen_has_prod :: "[nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}, nat, 'a] \<Rightarrow> bool" 
+  where "gen_has_prod f M p \<equiv> (\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> p \<and> p \<noteq> 0"
+
+text\<open>The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\<close>
+definition has_prod :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "has'_prod" 80)
+  where "f has_prod p \<equiv> gen_has_prod f 0 p \<or> (\<exists>i q. p = 0 \<and> f i = 0 \<and> gen_has_prod f (Suc i) q)"
+
 definition convergent_prod :: "(nat \<Rightarrow> 'a :: {t2_space,comm_semiring_1}) \<Rightarrow> bool" where
-  "convergent_prod f \<longleftrightarrow> (\<exists>M L. (\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> L \<and> L \<noteq> 0)"
+  "convergent_prod f \<equiv> \<exists>M p. gen_has_prod f M p"
+
+definition prodinf :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a"
+    (binder "\<Prod>" 10)
+  where "prodinf f = (THE p. f has_prod p)"
+
+lemmas prod_defs = gen_has_prod_def has_prod_def convergent_prod_def prodinf_def
+
+lemma has_prod_subst[trans]: "f = g \<Longrightarrow> g has_prod z \<Longrightarrow> f has_prod z"
+  by simp
+
+lemma has_prod_cong: "(\<And>n. f n = g n) \<Longrightarrow> f has_prod c \<longleftrightarrow> g has_prod c"
+  by presburger
 
 lemma convergent_prod_altdef:
   fixes f :: "nat \<Rightarrow> 'a :: {t2_space,comm_semiring_1}"
@@ -60,7 +78,7 @@
 proof
   assume "convergent_prod f"
   then obtain M L where *: "(\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> L" "L \<noteq> 0"
-    by (auto simp: convergent_prod_def)
+    by (auto simp: prod_defs)
   have "f i \<noteq> 0" if "i \<ge> M" for i
   proof
     assume "f i = 0"
@@ -79,7 +97,7 @@
   qed
   with * show "(\<exists>M L. (\<forall>n\<ge>M. f n \<noteq> 0) \<and> (\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> L \<and> L \<noteq> 0)" 
     by blast
-qed (auto simp: convergent_prod_def)
+qed (auto simp: prod_defs)
 
 definition abs_convergent_prod :: "(nat \<Rightarrow> _) \<Rightarrow> bool" where
   "abs_convergent_prod f \<longleftrightarrow> convergent_prod (\<lambda>i. 1 + norm (f i - 1))"
@@ -101,12 +119,12 @@
     qed
   qed (use L in simp_all)
   hence "L \<noteq> 0" by auto
-  with L show ?thesis unfolding abs_convergent_prod_def convergent_prod_def
+  with L show ?thesis unfolding abs_convergent_prod_def prod_defs
     by (intro exI[of _ "0::nat"] exI[of _ L]) auto
 qed
 
 lemma
-  fixes f :: "nat \<Rightarrow> 'a :: {real_normed_div_algebra,idom}"
+  fixes f :: "nat \<Rightarrow> 'a :: {topological_semigroup_mult,t2_space,idom}"
   assumes "convergent_prod f"
   shows   convergent_prod_imp_convergent: "convergent (\<lambda>n. \<Prod>i\<le>n. f i)"
     and   convergent_prod_to_zero_iff:    "(\<lambda>n. \<Prod>i\<le>n. f i) \<longlonglongrightarrow> 0 \<longleftrightarrow> (\<exists>i. f i = 0)"
@@ -141,8 +159,30 @@
   qed
 qed
 
+lemma convergent_prod_iff_nz_lim:
+  fixes f :: "nat \<Rightarrow> 'a :: {topological_semigroup_mult,t2_space,idom}"
+  assumes "\<And>i. f i \<noteq> 0"
+  shows "convergent_prod f \<longleftrightarrow> (\<exists>L. (\<lambda>n. \<Prod>i\<le>n. f i) \<longlonglongrightarrow> L \<and> L \<noteq> 0)"
+    (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?lhs then show ?rhs
+    using assms convergentD convergent_prod_imp_convergent convergent_prod_to_zero_iff by blast
+next
+  assume ?rhs then show ?lhs
+    unfolding prod_defs
+    by (rule_tac x="0" in exI) (auto simp: )
+qed
+
+lemma convergent_prod_iff_convergent: 
+  fixes f :: "nat \<Rightarrow> 'a :: {topological_semigroup_mult,t2_space,idom}"
+  assumes "\<And>i. f i \<noteq> 0"
+  shows "convergent_prod f \<longleftrightarrow> convergent (\<lambda>n. \<Prod>i\<le>n. f i) \<and> lim (\<lambda>n. \<Prod>i\<le>n. f i) \<noteq> 0"
+  by (force simp add: convergent_prod_iff_nz_lim assms convergent_def limI)
+
+
 lemma abs_convergent_prod_altdef:
-  "abs_convergent_prod f \<longleftrightarrow> convergent (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))"
+  fixes f :: "nat \<Rightarrow> 'a :: {one,real_normed_vector}"
+  shows  "abs_convergent_prod f \<longleftrightarrow> convergent (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))"
 proof
   assume "abs_convergent_prod f"
   thus "convergent (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))"
@@ -180,7 +220,7 @@
   also have "norm (1::'a) = 1" by simp
   also note insert.IH
   also have "(\<Prod>n\<in>A. 1 + norm (f n)) - 1 + norm (f x) * (\<Prod>x\<in>A. 1 + norm (f x)) =
-               (\<Prod>n\<in>insert x A. 1 + norm (f n)) - 1"
+             (\<Prod>n\<in>insert x A. 1 + norm (f n)) - 1"
     using insert.hyps by (simp add: algebra_simps)
   finally show ?case by - (simp_all add: mult_left_mono)
 qed simp_all
@@ -297,8 +337,9 @@
   shows   "convergent_prod f"
 proof -
   from assms obtain M L where "(\<lambda>n. \<Prod>k\<le>n. f (k + (M + m))) \<longlonglongrightarrow> L" "L \<noteq> 0"
-    by (auto simp: convergent_prod_def add.assoc)
-  thus "convergent_prod f" unfolding convergent_prod_def by blast
+    by (auto simp: prod_defs add.assoc)
+  thus "convergent_prod f" 
+    unfolding prod_defs by blast
 qed
 
 lemma abs_convergent_prod_offset:
@@ -334,7 +375,8 @@
     by (intro tendsto_divide tendsto_const) auto
   hence "(\<lambda>n. \<Prod>k\<le>n. f (k + M + m)) \<longlonglongrightarrow> L / C" by simp
   moreover from \<open>L \<noteq> 0\<close> have "L / C \<noteq> 0" by simp
-  ultimately show ?thesis unfolding convergent_prod_def by blast
+  ultimately show ?thesis 
+    unfolding prod_defs by blast
 qed
 
 lemma abs_convergent_prod_ignore_initial_segment:
@@ -343,11 +385,6 @@
   using assms unfolding abs_convergent_prod_def 
   by (rule convergent_prod_ignore_initial_segment)
 
-lemma summable_LIMSEQ': 
-  assumes "summable (f::nat\<Rightarrow>'a::{t2_space,comm_monoid_add})"
-  shows   "(\<lambda>n. \<Sum>i\<le>n. f i) \<longlonglongrightarrow> suminf f"
-  using assms sums_def_le by blast
-
 lemma abs_convergent_prod_imp_convergent_prod:
   fixes f :: "nat \<Rightarrow> 'a :: {real_normed_div_algebra,complete_space,comm_ring_1}"
   assumes "abs_convergent_prod f"
@@ -473,7 +510,98 @@
     qed simp_all
     thus False by simp
   qed
-  with L show ?thesis by (auto simp: convergent_prod_def)
+  with L show ?thesis by (auto simp: prod_defs)
+qed
+
+lemma convergent_prod_offset_0:
+  fixes f :: "nat \<Rightarrow> 'a :: {idom,topological_semigroup_mult,t2_space}"
+  assumes "convergent_prod f" "\<And>i. f i \<noteq> 0"
+  shows "\<exists>p. gen_has_prod f 0 p"
+  using assms
+  unfolding convergent_prod_def
+proof (clarsimp simp: prod_defs)
+  fix M p
+  assume "(\<lambda>n. \<Prod>i\<le>n. f (i + M)) \<longlonglongrightarrow> p" "p \<noteq> 0"
+  then have "(\<lambda>n. prod f {..<M} * (\<Prod>i\<le>n. f (i + M))) \<longlonglongrightarrow> prod f {..<M} * p"
+    by (metis tendsto_mult_left)
+  moreover have "prod f {..<M} * (\<Prod>i\<le>n. f (i + M)) = prod f {..n+M}" for n
+  proof -
+    have "{..n+M} = {..<M} \<union> {M..n+M}"
+      by auto
+    then have "prod f {..n+M} = prod f {..<M} * prod f {M..n+M}"
+      by simp (subst prod.union_disjoint; force)
+    also have "... = prod f {..<M} * (\<Prod>i\<le>n. f (i + M))"
+      by (metis (mono_tags, lifting) add.left_neutral atMost_atLeast0 prod_shift_bounds_cl_nat_ivl)
+    finally show ?thesis by metis
+  qed
+  ultimately have "(\<lambda>n. prod f {..n}) \<longlonglongrightarrow> prod f {..<M} * p"
+    by (auto intro: LIMSEQ_offset [where k=M])
+  then show "\<exists>p. (\<lambda>n. prod f {..n}) \<longlonglongrightarrow> p \<and> p \<noteq> 0"
+    using \<open>p \<noteq> 0\<close> assms
+    by (rule_tac x="prod f {..<M} * p" in exI) auto
+qed
+
+lemma prodinf_eq_lim:
+  fixes f :: "nat \<Rightarrow> 'a :: {idom,topological_semigroup_mult,t2_space}"
+  assumes "convergent_prod f" "\<And>i. f i \<noteq> 0"
+  shows "prodinf f = lim (\<lambda>n. \<Prod>i\<le>n. f i)"
+  using assms convergent_prod_offset_0 [OF assms]
+  by (simp add: prod_defs lim_def) (metis (no_types) assms(1) convergent_prod_to_zero_iff)
+
+lemma has_prod_one[simp, intro]: "(\<lambda>n. 1) has_prod 1"
+  unfolding prod_defs by auto
+
+lemma convergent_prod_one[simp, intro]: "convergent_prod (\<lambda>n. 1)"
+  unfolding prod_defs by auto
+
+lemma prodinf_cong: "(\<And>n. f n = g n) \<Longrightarrow> prodinf f = prodinf g"
+  by presburger
+
+lemma convergent_prod_cong:
+  fixes f g :: "nat \<Rightarrow> 'a::{field,topological_semigroup_mult,t2_space}"
+  assumes ev: "eventually (\<lambda>x. f x = g x) sequentially" and f: "\<And>i. f i \<noteq> 0" and g: "\<And>i. g i \<noteq> 0"
+  shows "convergent_prod f = convergent_prod g"
+proof -
+  from assms obtain N where N: "\<forall>n\<ge>N. f n = g n"
+    by (auto simp: eventually_at_top_linorder)
+  define C where "C = (\<Prod>k<N. f k / g k)"
+  with g have "C \<noteq> 0"
+    by (simp add: f)
+  have *: "eventually (\<lambda>n. prod f {..n} = C * prod g {..n}) sequentially"
+    using eventually_ge_at_top[of N]
+  proof eventually_elim
+    case (elim n)
+    then have "{..n} = {..<N} \<union> {N..n}"
+      by auto
+    also have "prod f ... = prod f {..<N} * prod f {N..n}"
+      by (intro prod.union_disjoint) auto
+    also from N have "prod f {N..n} = prod g {N..n}"
+      by (intro prod.cong) simp_all
+    also have "prod f {..<N} * prod g {N..n} = C * (prod g {..<N} * prod g {N..n})"
+      unfolding C_def by (simp add: g prod_dividef)
+    also have "prod g {..<N} * prod g {N..n} = prod g ({..<N} \<union> {N..n})"
+      by (intro prod.union_disjoint [symmetric]) auto
+    also from elim have "{..<N} \<union> {N..n} = {..n}"
+      by auto                                                                    
+    finally show "prod f {..n} = C * prod g {..n}" .
+  qed
+  then have cong: "convergent (\<lambda>n. prod f {..n}) = convergent (\<lambda>n. C * prod g {..n})"
+    by (rule convergent_cong)
+  show ?thesis
+  proof
+    assume cf: "convergent_prod f"
+    then have "\<not> (\<lambda>n. prod g {..n}) \<longlonglongrightarrow> 0"
+      using tendsto_mult_left * convergent_prod_to_zero_iff f filterlim_cong by fastforce
+    then show "convergent_prod g"
+      by (metis convergent_mult_const_iff \<open>C \<noteq> 0\<close> cong cf convergent_LIMSEQ_iff convergent_prod_iff_convergent convergent_prod_imp_convergent g)
+  next
+    assume cg: "convergent_prod g"
+    have "\<exists>a. C * a \<noteq> 0 \<and> (\<lambda>n. prod g {..n}) \<longlonglongrightarrow> a"
+      by (metis (no_types) \<open>C \<noteq> 0\<close> cg convergent_prod_iff_nz_lim divide_eq_0_iff g nonzero_mult_div_cancel_right)
+    then show "convergent_prod f"
+      using "*" tendsto_mult_left filterlim_cong
+      by (fastforce simp add: convergent_prod_iff_nz_lim f)
+  qed
 qed
 
 end
--- a/src/HOL/Analysis/Inner_Product.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Inner_Product.thy	Thu May 03 15:07:14 2018 +0200
@@ -177,9 +177,6 @@
   using isCont_power[OF continuous_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2]
   by (force simp add: power2_eq_square)
 
-lemma norm_eq_0_dot: "norm x = 0 \<longleftrightarrow> inner x x = (0::real)"
-  by simp (* TODO: delete *)
-
 lemma norm_triangle_sub:
   fixes x y :: "'a::real_normed_vector"
   shows "norm x \<le> norm y + norm (x - y)"
--- a/src/HOL/Analysis/Interval_Integral.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Interval_Integral.thy	Thu May 03 15:07:14 2018 +0200
@@ -109,8 +109,7 @@
   from ereal_incseq_approx[OF this] guess X .
   then show thesis
     apply (intro that[of "\<lambda>i. - X i"])
-    apply (auto simp add: uminus_ereal.simps[symmetric] decseq_def incseq_def
-                simp del: uminus_ereal.simps)
+    apply (auto simp add: decseq_def incseq_def reorient: uminus_ereal.simps)
     apply (metis ereal_minus_less_minus ereal_uminus_uminus ereal_Lim_uminus)+
     done
 qed
@@ -143,7 +142,7 @@
     fix x i assume "l i \<le> x" "x \<le> u i"
     with \<open>a < ereal (l i)\<close> \<open>ereal (u i) < b\<close>
     show "a < ereal x" "ereal x < b"
-      by (auto simp del: ereal_less_eq(3) simp add: ereal_less_eq(3)[symmetric])
+      by (auto reorient: ereal_less_eq(3))
   qed
   show thesis
     by (intro that) fact+
@@ -304,8 +303,8 @@
   have "LBINT x:{x. - x \<in> einterval a b}. f (- x) = LBINT x:einterval (- b) (- a). f (- x)"
     unfolding interval_lebesgue_integrable_def set_lebesgue_integral_def
     apply (rule Bochner_Integration.integral_cong [OF refl])
-    by (auto simp: einterval_iff ereal_uminus_le_reorder ereal_uminus_less_reorder not_less uminus_ereal.simps[symmetric]
-             simp del: uminus_ereal.simps
+    by (auto simp: einterval_iff ereal_uminus_le_reorder ereal_uminus_less_reorder not_less
+             reorient: uminus_ereal.simps
              split: split_indicator)
   then show ?case
     unfolding interval_lebesgue_integral_def 
@@ -676,7 +675,7 @@
     fix i show "set_integrable lborel {l i .. u i} f"
       using \<open>a < l i\<close> \<open>u i < b\<close> unfolding set_integrable_def
       by (intro borel_integrable_compact f continuous_at_imp_continuous_on compact_Icc ballI)
-         (auto simp del: ereal_less_eq simp add: ereal_less_eq(3)[symmetric])
+         (auto reorient: ereal_less_eq)
   qed
   have 2: "set_borel_measurable lborel (einterval a b) f"
     unfolding set_borel_measurable_def
--- a/src/HOL/Analysis/Lebesgue_Measure.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Lebesgue_Measure.thy	Thu May 03 15:07:14 2018 +0200
@@ -283,7 +283,7 @@
     also have "... \<le> (\<Sum>i\<le>n. F (r i) - F (l i)) + epsilon"
       using finS Sbound Sprop by (auto intro!: add_right_mono sum_mono2)
     finally have "ennreal (F b - F a) \<le> (\<Sum>i\<le>n. ennreal (F (r i) - F (l i))) + epsilon"
-      using egt0 by (simp add: ennreal_plus[symmetric] sum_nonneg del: ennreal_plus)
+      using egt0 by (simp add: sum_nonneg reorient: ennreal_plus)
     then show "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i))) + (epsilon :: real)"
       by (rule order_trans) (auto intro!: add_mono sum_le_suminf simp del: sum_ennreal)
   qed
--- a/src/HOL/Analysis/Linear_Algebra.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Linear_Algebra.thy	Thu May 03 15:07:14 2018 +0200
@@ -27,22 +27,8 @@
   show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scale)
 qed
 
-lemma bounded_linearI:
-  assumes "\<And>x y. f (x + y) = f x + f y"
-    and "\<And>r x. f (r *\<^sub>R x) = r *\<^sub>R f x"
-    and "\<And>x. norm (f x) \<le> norm x * K"
-  shows "bounded_linear f"
-  using assms by (rule bounded_linear_intro) (* FIXME: duplicate *)
-
-lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> (UNIV::'a::finite set)}"
-proof -
-  have eq: "{f x |x. x\<in> UNIV} = f ` UNIV"
-    by auto
-  show ?thesis unfolding eq
-    apply (rule finite_imageI)
-    apply (rule finite)
-    done
-qed
+lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x \<in> (UNIV::'a::finite set)}"
+  using finite finite_image_set by blast
 
 
 subsection%unimportant \<open>More interesting properties of the norm.\<close>
@@ -123,10 +109,8 @@
 lemma sum_group:
   assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"
   shows "sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) T = sum g S"
-  apply (subst sum_image_gen[OF fS, of g f])
-  apply (rule sum.mono_neutral_right[OF fT fST])
-  apply (auto intro: sum.neutral)
-  done
+  unfolding sum_image_gen[OF fS, of g f]
+  by (auto intro: sum.neutral sum.mono_neutral_right[OF fT fST])
 
 lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z"
 proof
@@ -351,12 +335,7 @@
   assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n"
     and "\<forall>n \<ge> m. e n \<le> e m"
   shows "\<forall>n \<ge> m. d n < e m"
-  using assms
-  apply auto
-  apply (erule_tac x="n" in allE)
-  apply (erule_tac x="n" in allE)
-  apply auto
-  done
+  using assms by force
 
 lemma infinite_enumerate:
   assumes fS: "infinite S"
@@ -468,10 +447,7 @@
 next
   case False
   with y x1 show ?thesis
-    apply auto
-    apply (rule exI[where x=1])
-    apply auto
-    done
+    by (metis less_le_trans not_less power_one_right)
 qed
 
 lemma forall_pos_mono:
@@ -522,11 +498,7 @@
     proof -
       from Basis_le_norm[OF that, of x]
       show "norm (?g i) \<le> norm (f i) * norm x"
-        unfolding norm_scaleR
-        apply (subst mult.commute)
-        apply (rule mult_mono)
-        apply (auto simp add: field_simps)
-        done
+        unfolding norm_scaleR  by (metis mult.commute mult_left_mono norm_ge_zero)
     qed
     from sum_norm_le[of _ ?g, OF th]
     show "norm (f x) \<le> ?B * norm x"
@@ -611,23 +583,17 @@
   fix x :: 'm
   fix y :: 'n
   have "norm (h x y) = norm (h (sum (\<lambda>i. (x \<bullet> i) *\<^sub>R i) Basis) (sum (\<lambda>i. (y \<bullet> i) *\<^sub>R i) Basis))"
-    apply (subst euclidean_representation[where 'a='m])
-    apply (subst euclidean_representation[where 'a='n])
-    apply rule
-    done
+    by (simp add: euclidean_representation)
   also have "\<dots> = norm (sum (\<lambda> (i,j). h ((x \<bullet> i) *\<^sub>R i) ((y \<bullet> j) *\<^sub>R j)) (Basis \<times> Basis))"
     unfolding bilinear_sum[OF bh] ..
   finally have th: "norm (h x y) = \<dots>" .
-  show "norm (h x y) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)) * norm x * norm y"
-    apply (auto simp add: sum_distrib_right th sum.cartesian_product)
-    apply (rule sum_norm_le)
-    apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh]
-      field_simps simp del: scaleR_scaleR)
-    apply (rule mult_mono)
-    apply (auto simp add: zero_le_mult_iff Basis_le_norm)
-    apply (rule mult_mono)
-    apply (auto simp add: zero_le_mult_iff Basis_le_norm)
-    done
+  have "\<And>i j. \<lbrakk>i \<in> Basis; j \<in> Basis\<rbrakk>
+           \<Longrightarrow> \<bar>x \<bullet> i\<bar> * (\<bar>y \<bullet> j\<bar> * norm (h i j)) \<le> norm x * (norm y * norm (h i j))"
+    by (auto simp add: zero_le_mult_iff Basis_le_norm mult_mono)
+  then show "norm (h x y) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)) * norm x * norm y"
+    unfolding sum_distrib_right th sum.cartesian_product
+    by (clarsimp simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh]
+      field_simps simp del: scaleR_scaleR intro!: sum_norm_le)
 qed
 
 lemma bilinear_conv_bounded_bilinear:
@@ -645,15 +611,9 @@
     show "h x (y + z) = h x y + h x z"
       using \<open>bilinear h\<close> unfolding bilinear_def linear_iff by simp
   next
-    fix r x y
-    show "h (scaleR r x) y = scaleR r (h x y)"
+    show "h (scaleR r x) y = scaleR r (h x y)" "h x (scaleR r y) = scaleR r (h x y)" for r x y
       using \<open>bilinear h\<close> unfolding bilinear_def linear_iff
-      by simp
-  next
-    fix r x y
-    show "h x (scaleR r y) = scaleR r (h x y)"
-      using \<open>bilinear h\<close> unfolding bilinear_def linear_iff
-      by simp
+      by simp_all
   next
     have "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
       using \<open>bilinear h\<close> by (rule bilinear_bounded)
@@ -803,7 +763,7 @@
 proof -
   from basis_exists[of V] obtain B where
     B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "card B = dim V"
-    by blast
+    by force
   from B have fB: "finite B" "card B = dim V"
     using independent_bound by auto
   from basis_orthogonal[OF fB(1)] obtain C where
@@ -855,8 +815,8 @@
     done
   with a have a0:"?a  \<noteq> 0"
     by auto
-  have "\<forall>x\<in>span B. ?a \<bullet> x = 0"
-  proof (rule span_induct')
+  have "?a \<bullet> x = 0" if "x\<in>span B" for x
+  proof (rule span_induct [OF that])
     show "subspace {x. ?a \<bullet> x = 0}"
       by (auto simp add: subspace_def inner_add)
   next
@@ -879,9 +839,9 @@
           intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])
         done
     }
-    then show "\<forall>x \<in> B. ?a \<bullet> x = 0"
-      by blast
-  qed
+    then show "?a \<bullet> x = 0" if "x \<in> B" for x
+      using that by blast
+    qed
   with a0 show ?thesis
     unfolding sSB by (auto intro: exI[where x="?a"])
 qed
@@ -927,8 +887,9 @@
     and bg: "bilinear g"
     and SB: "S \<subseteq> span B"
     and TC: "T \<subseteq> span C"
-    and fg: "\<forall>x\<in> B. \<forall>y\<in> C. f x y = g x y"
-  shows "\<forall>x\<in>S. \<forall>y\<in>T. f x y = g x y "
+    and "x\<in>S" "y\<in>T"
+    and fg: "\<And>x y. \<lbrakk>x \<in> B; y\<in> C\<rbrakk> \<Longrightarrow> f x y = g x y"
+  shows "f x y = g x y"
 proof -
   let ?P = "{x. \<forall>y\<in> span C. f x y = g x y}"
   from bf bg have sp: "subspace ?P"
@@ -936,27 +897,25 @@
     by (auto simp add: span_zero bilinear_lzero[OF bf] bilinear_lzero[OF bg]
         span_add Ball_def
       intro: bilinear_ladd[OF bf])
-
-  have "\<forall>x \<in> span B. \<forall>y\<in> span C. f x y = g x y"
-    apply (rule span_induct' [OF _ sp])
-    apply (rule ballI)
-    apply (rule span_induct')
-    apply (simp add: fg)
+  have sfg: "\<And>x. x \<in> B \<Longrightarrow> subspace {a. f x a = g x a}"
     apply (auto simp add: subspace_def)
     using bf bg unfolding bilinear_def linear_iff
       apply (auto simp add: span_zero bilinear_rzero[OF bf] bilinear_rzero[OF bg]
         span_add Ball_def
       intro: bilinear_ladd[OF bf])
     done
+  have "\<forall>y\<in> span C. f x y = g x y" if "x \<in> span B" for x
+    apply (rule span_induct [OF that sp])
+    using fg sfg span_induct by blast
   then show ?thesis
-    using SB TC by auto
+    using SB TC assms by auto
 qed
 
 lemma bilinear_eq_stdbasis:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> _"
   assumes bf: "bilinear f"
     and bg: "bilinear g"
-    and fg: "\<forall>i\<in>Basis. \<forall>j\<in>Basis. f i j = g i j"
+    and fg: "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> f i j = g i j"
   shows "f = g"
   using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis] fg] by blast
 
@@ -1010,28 +969,21 @@
   by (simp add: infnorm_eq_0)
 
 lemma infnorm_neg: "infnorm (- x) = infnorm x"
-  unfolding infnorm_def
-  apply (rule cong[of "Sup" "Sup"])
-  apply blast
-  apply auto
-  done
+  unfolding infnorm_def by simp
 
 lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)"
-proof -
-  have "y - x = - (x - y)" by simp
-  then show ?thesis
-    by (metis infnorm_neg)
-qed
+  by (metis infnorm_neg minus_diff_eq)
 
-lemma real_abs_sub_infnorm: "\<bar>infnorm x - infnorm y\<bar> \<le> infnorm (x - y)"
+lemma absdiff_infnorm: "\<bar>infnorm x - infnorm y\<bar> \<le> infnorm (x - y)"
 proof -
-  have th: "\<And>(nx::real) n ny. nx \<le> n + ny \<Longrightarrow> ny \<le> n + nx \<Longrightarrow> \<bar>nx - ny\<bar> \<le> n"
+  have *: "\<And>(nx::real) n ny. nx \<le> n + ny \<Longrightarrow> ny \<le> n + nx \<Longrightarrow> \<bar>nx - ny\<bar> \<le> n"
     by arith
-  from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
-  have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"
-    "infnorm y \<le> infnorm (x - y) + infnorm x"
-    by (simp_all add: field_simps infnorm_neg)
-  from th[OF ths] show ?thesis .
+  show ?thesis
+  proof (rule *)
+    from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
+    show "infnorm x \<le> infnorm (x - y) + infnorm y" "infnorm y \<le> infnorm (x - y) + infnorm x"
+      by (simp_all add: field_simps infnorm_neg)
+  qed
 qed
 
 lemma real_abs_infnorm: "\<bar>infnorm x\<bar> = infnorm x"
@@ -1046,8 +998,7 @@
   unfolding infnorm_Max
 proof (safe intro!: Max_eqI)
   let ?B = "(\<lambda>i. \<bar>x \<bullet> i\<bar>) ` Basis"
-  {
-    fix b :: 'a
+  { fix b :: 'a
     assume "b \<in> Basis"
     then show "\<bar>a *\<^sub>R x \<bullet> b\<bar> \<le> \<bar>a\<bar> * Max ?B"
       by (simp add: abs_mult mult_left_mono)
@@ -1073,27 +1024,17 @@
 lemma norm_le_infnorm:
   fixes x :: "'a::euclidean_space"
   shows "norm x \<le> sqrt DIM('a) * infnorm x"
-proof -
-  let ?d = "DIM('a)"
-  have "real ?d \<ge> 0"
-    by simp
-  then have d2: "(sqrt (real ?d))\<^sup>2 = real ?d"
-    by (auto intro: real_sqrt_pow2)
-  have th: "sqrt (real ?d) * infnorm x \<ge> 0"
+  unfolding norm_eq_sqrt_inner id_def 
+proof (rule real_le_lsqrt[OF inner_ge_zero])
+  show "sqrt DIM('a) * infnorm x \<ge> 0"
     by (simp add: zero_le_mult_iff infnorm_pos_le)
-  have th1: "x \<bullet> x \<le> (sqrt (real ?d) * infnorm x)\<^sup>2"
-    unfolding power_mult_distrib d2
-    apply (subst euclidean_inner)
-    apply (subst power2_abs[symmetric])
-    apply (rule order_trans[OF sum_bounded_above[where K="\<bar>infnorm x\<bar>\<^sup>2"]])
-    apply (auto simp add: power2_eq_square[symmetric])
-    apply (subst power2_abs[symmetric])
-    apply (rule power_mono)
-    apply (auto simp: infnorm_Max)
-    done
-  from real_le_lsqrt[OF inner_ge_zero th th1]
-  show ?thesis
-    unfolding norm_eq_sqrt_inner id_def .
+  have "x \<bullet> x \<le> (\<Sum>b\<in>Basis. x \<bullet> b * (x \<bullet> b))"
+    by (metis euclidean_inner order_refl)
+  also have "... \<le> DIM('a) * \<bar>infnorm x\<bar>\<^sup>2"
+    by (rule sum_bounded_above) (metis Basis_le_infnorm abs_le_square_iff power2_eq_square real_abs_infnorm)
+  also have "... \<le> (sqrt DIM('a) * infnorm x)\<^sup>2"
+    by (simp add: power_mult_distrib)
+  finally show "x \<bullet> x \<le> (sqrt DIM('a) * infnorm x)\<^sup>2" .
 qed
 
 lemma tendsto_infnorm [tendsto_intros]:
@@ -1103,46 +1044,30 @@
   fix r :: real
   assume "r > 0"
   then show "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (infnorm x - infnorm a) < r"
-    by (metis real_norm_def le_less_trans real_abs_sub_infnorm infnorm_le_norm)
+    by (metis real_norm_def le_less_trans absdiff_infnorm infnorm_le_norm)
 qed
 
 text \<open>Equality in Cauchy-Schwarz and triangle inequalities.\<close>
 
 lemma norm_cauchy_schwarz_eq: "x \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x"
   (is "?lhs \<longleftrightarrow> ?rhs")
-proof -
-  {
-    assume h: "x = 0"
-    then have ?thesis by simp
-  }
-  moreover
-  {
-    assume h: "y = 0"
-    then have ?thesis by simp
-  }
-  moreover
-  {
-    assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
-    from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"]
-    have "?rhs \<longleftrightarrow>
+proof (cases "x=0")
+  case True
+  then show ?thesis 
+    by auto
+next
+  case False
+  from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"]
+  have "?rhs \<longleftrightarrow>
       (norm y * (norm y * norm x * norm x - norm x * (x \<bullet> y)) -
         norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) =  0)"
-      using x y
-      unfolding inner_simps
-      unfolding power2_norm_eq_inner[symmetric] power2_eq_square right_minus_eq
-      apply (simp add: inner_commute)
-      apply (simp add: field_simps)
-      apply metis
-      done
-    also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y
-      by (simp add: field_simps inner_commute)
-    also have "\<dots> \<longleftrightarrow> ?lhs" using x y
-      apply simp
-      apply metis
-      done
-    finally have ?thesis by blast
-  }
-  ultimately show ?thesis by blast
+    using False unfolding inner_simps
+    by (auto simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)
+  also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" 
+    using False  by (simp add: field_simps inner_commute)
+  also have "\<dots> \<longleftrightarrow> ?lhs" 
+    using False by auto
+  finally show ?thesis by metis
 qed
 
 lemma norm_cauchy_schwarz_abs_eq:
@@ -1154,7 +1079,7 @@
     by arith
   have "?rhs \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm (- x) *\<^sub>R y = norm y *\<^sub>R (- x)"
     by simp
-  also have "\<dots> \<longleftrightarrow>(x \<bullet> y = norm x * norm y \<or> (- x) \<bullet> y = norm x * norm y)"
+  also have "\<dots> \<longleftrightarrow> (x \<bullet> y = norm x * norm y \<or> (- x) \<bullet> y = norm x * norm y)"
     unfolding norm_cauchy_schwarz_eq[symmetric]
     unfolding norm_minus_cancel norm_scaleR ..
   also have "\<dots> \<longleftrightarrow> ?lhs"
@@ -1166,33 +1091,21 @@
 lemma norm_triangle_eq:
   fixes x y :: "'a::real_inner"
   shows "norm (x + y) = norm x + norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x"
-proof -
-  {
-    assume x: "x = 0 \<or> y = 0"
-    then have ?thesis
-      by (cases "x = 0") simp_all
-  }
-  moreover
-  {
-    assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
-    then have "norm x \<noteq> 0" "norm y \<noteq> 0"
-      by simp_all
-    then have n: "norm x > 0" "norm y > 0"
-      using norm_ge_zero[of x] norm_ge_zero[of y] by arith+
-    have th: "\<And>(a::real) b c. a + b + c \<noteq> 0 \<Longrightarrow> a = b + c \<longleftrightarrow> a\<^sup>2 = (b + c)\<^sup>2"
-      by algebra
-    have "norm (x + y) = norm x + norm y \<longleftrightarrow> (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2"
-      apply (rule th)
-      using n norm_ge_zero[of "x + y"]
-      apply arith
-      done
-    also have "\<dots> \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x"
-      unfolding norm_cauchy_schwarz_eq[symmetric]
-      unfolding power2_norm_eq_inner inner_simps
-      by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)
-    finally have ?thesis .
-  }
-  ultimately show ?thesis by blast
+proof (cases "x = 0 \<or> y = 0")
+  case True
+  then show ?thesis 
+    by force
+next
+  case False
+  then have n: "norm x > 0" "norm y > 0"
+    by auto
+  have "norm (x + y) = norm x + norm y \<longleftrightarrow> (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2"
+    by simp
+  also have "\<dots> \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x"
+    unfolding norm_cauchy_schwarz_eq[symmetric]
+    unfolding power2_norm_eq_inner inner_simps
+    by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)
+  finally show ?thesis .
 qed
 
 
@@ -1251,81 +1164,67 @@
 lemma collinear_2 [iff]: "collinear {x, y}"
   apply (simp add: collinear_def)
   apply (rule exI[where x="x - y"])
-  apply auto
-  apply (rule exI[where x=1], simp)
-  apply (rule exI[where x="- 1"], simp)
-  done
+  by (metis minus_diff_eq scaleR_left.minus scaleR_one)
 
 lemma collinear_lemma: "collinear {0, x, y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *\<^sub>R x)"
   (is "?lhs \<longleftrightarrow> ?rhs")
-proof -
-  {
-    assume "x = 0 \<or> y = 0"
-    then have ?thesis
-      by (cases "x = 0") (simp_all add: collinear_2 insert_commute)
-  }
-  moreover
-  {
-    assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
-    have ?thesis
-    proof
-      assume h: "?lhs"
-      then obtain u where u: "\<forall> x\<in> {0,x,y}. \<forall>y\<in> {0,x,y}. \<exists>c. x - y = c *\<^sub>R u"
-        unfolding collinear_def by blast
-      from u[rule_format, of x 0] u[rule_format, of y 0]
-      obtain cx and cy where
-        cx: "x = cx *\<^sub>R u" and cy: "y = cy *\<^sub>R u"
-        by auto
-      from cx x have cx0: "cx \<noteq> 0" by auto
-      from cy y have cy0: "cy \<noteq> 0" by auto
-      let ?d = "cy / cx"
-      from cx cy cx0 have "y = ?d *\<^sub>R x"
-        by simp
-      then show ?rhs using x y by blast
-    next
-      assume h: "?rhs"
-      then obtain c where c: "y = c *\<^sub>R x"
-        using x y by blast
-      show ?lhs
-        unfolding collinear_def c
-        apply (rule exI[where x=x])
-        apply auto
-        apply (rule exI[where x="- 1"], simp)
-        apply (rule exI[where x= "-c"], simp)
+proof (cases "x = 0 \<or> y = 0")
+  case True
+  then show ?thesis
+    by (auto simp: insert_commute)
+next
+  case False
+  show ?thesis 
+  proof
+    assume h: "?lhs"
+    then obtain u where u: "\<forall> x\<in> {0,x,y}. \<forall>y\<in> {0,x,y}. \<exists>c. x - y = c *\<^sub>R u"
+      unfolding collinear_def by blast
+    from u[rule_format, of x 0] u[rule_format, of y 0]
+    obtain cx and cy where
+      cx: "x = cx *\<^sub>R u" and cy: "y = cy *\<^sub>R u"
+      by auto
+    from cx cy False have cx0: "cx \<noteq> 0" and cy0: "cy \<noteq> 0" by auto
+    let ?d = "cy / cx"
+    from cx cy cx0 have "y = ?d *\<^sub>R x"
+      by simp
+    then show ?rhs using False by blast
+  next
+    assume h: "?rhs"
+    then obtain c where c: "y = c *\<^sub>R x"
+      using False by blast
+    show ?lhs
+      unfolding collinear_def c
+      apply (rule exI[where x=x])
+      apply auto
+          apply (rule exI[where x="- 1"], simp)
+         apply (rule exI[where x= "-c"], simp)
         apply (rule exI[where x=1], simp)
-        apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib)
-        apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib)
-        done
-    qed
-  }
-  ultimately show ?thesis by blast
+       apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib)
+      apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib)
+      done
+  qed
 qed
 
 lemma norm_cauchy_schwarz_equal: "\<bar>x \<bullet> y\<bar> = norm x * norm y \<longleftrightarrow> collinear {0, x, y}"
-  unfolding norm_cauchy_schwarz_abs_eq
-  apply (cases "x=0", simp_all)
-  apply (cases "y=0", simp_all add: insert_commute)
-  unfolding collinear_lemma
-  apply simp
-  apply (subgoal_tac "norm x \<noteq> 0")
-  apply (subgoal_tac "norm y \<noteq> 0")
-  apply (rule iffI)
-  apply (cases "norm x *\<^sub>R y = norm y *\<^sub>R x")
-  apply (rule exI[where x="(1/norm x) * norm y"])
-  apply (drule sym)
-  unfolding scaleR_scaleR[symmetric]
-  apply (simp add: field_simps)
-  apply (rule exI[where x="(1/norm x) * - norm y"])
-  apply clarify
-  apply (drule sym)
-  unfolding scaleR_scaleR[symmetric]
-  apply (simp add: field_simps)
-  apply (erule exE)
-  apply (erule ssubst)
-  unfolding scaleR_scaleR
-  unfolding norm_scaleR
-  apply (subgoal_tac "norm x * c = \<bar>c\<bar> * norm x \<or> norm x * c = - \<bar>c\<bar> * norm x")
-  apply (auto simp add: field_simps)
-  done
+proof (cases "x=0")
+  case True
+  then show ?thesis
+    by (auto simp: insert_commute)
+next
+  case False
+  then have nnz: "norm x \<noteq> 0"
+    by auto
+  show ?thesis
+  proof
+    assume "\<bar>x \<bullet> y\<bar> = norm x * norm y"
+    then show "collinear {0, x, y}"
+      unfolding norm_cauchy_schwarz_abs_eq collinear_lemma 
+      by (meson eq_vector_fraction_iff nnz)
+  next
+    assume "collinear {0, x, y}"
+    with False show "\<bar>x \<bullet> y\<bar> = norm x * norm y"
+      unfolding norm_cauchy_schwarz_abs_eq collinear_lemma  by (auto simp: abs_if)
+  qed
+qed
 
 end
--- a/src/HOL/Analysis/Measure_Space.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Measure_Space.thy	Thu May 03 15:07:14 2018 +0200
@@ -1622,7 +1622,7 @@
     using emeasure_subadditive[OF measurable] fin
     apply simp
     apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure)
-    apply (auto simp add: ennreal_plus[symmetric] simp del: ennreal_plus)
+    apply (auto reorient: ennreal_plus)
     done
 qed
 
--- a/src/HOL/Analysis/Polytope.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Polytope.thy	Thu May 03 15:07:14 2018 +0200
@@ -1188,7 +1188,7 @@
     qed
     then have "dim (S \<inter> {x. a \<bullet> x = 0}) < n"
       by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff
-           inf_le1 \<open>dim S = n\<close> not_le rel_interior_subset span_0 span_clauses(1))
+           inf_le1 \<open>dim S = n\<close> not_le rel_interior_subset span_0 span_superset)
     then have "0 \<in> convex hull {x. x extreme_point_of (S \<inter> {x. a \<bullet> x = 0})}"
       by (rule less.IH) (auto simp: co less.prems)
     then show ?thesis
--- a/src/HOL/Analysis/Regularity.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Regularity.thy	Thu May 03 15:07:14 2018 +0200
@@ -107,7 +107,7 @@
     finally have "measure M (space M) \<le> measure M K + e"
       using \<open>0 < e\<close> by simp
     hence "emeasure M (space M) \<le> emeasure M K + e"
-      using \<open>0 < e\<close> by (simp add: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] del: ennreal_plus)
+      using \<open>0 < e\<close> by (simp add: emeasure_eq_measure reorient: ennreal_plus)
     moreover have "compact K"
       unfolding compact_eq_totally_bounded
     proof safe
@@ -139,9 +139,9 @@
       also have "\<dots> \<le> measure M (space M) - measure M K"
         by (simp add: emeasure_eq_measure sU sb finite_measure_mono)
       also have "\<dots> \<le> e"
-        using K \<open>0 < e\<close> by (simp add: emeasure_eq_measure ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
+        using K \<open>0 < e\<close> by (simp add: emeasure_eq_measure reorient: ennreal_plus)
       finally have "emeasure M A \<le> emeasure M (A \<inter> K) + ennreal e"
-        using \<open>0<e\<close> by (simp add: emeasure_eq_measure algebra_simps ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
+        using \<open>0<e\<close> by (simp add: emeasure_eq_measure algebra_simps reorient: ennreal_plus)
       moreover have "A \<inter> K \<subseteq> A" "compact (A \<inter> K)" using \<open>closed A\<close> \<open>compact K\<close> by auto
       ultimately show "\<exists>K \<subseteq> A. compact K \<and> emeasure M A \<le> emeasure M K + ennreal e"
         by blast
@@ -301,7 +301,7 @@
       have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (D i)) + e/2" using n0 by simp
       also have "(\<Sum>i<n0. measure M (D i)) \<le> (\<Sum>i<n0. measure M (K i) + e/(2*Suc n0))"
         using K \<open>0 < e\<close>
-        by (auto intro: sum_mono simp: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] simp del: ennreal_plus)
+        by (auto intro: sum_mono simp: emeasure_eq_measure reorient: ennreal_plus)
       also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))"
         by (simp add: sum.distrib)
       also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) +  e / 2" using \<open>0 < e\<close>
@@ -310,7 +310,7 @@
       have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (K i)) + e / 2 + e / 2"
         by auto
       hence "M (\<Union>i. D i) < M ?K + e"
-        using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure measure_nonneg sum_nonneg ennreal_less_iff ennreal_plus[symmetric] simp del: ennreal_plus)
+        using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure sum_nonneg ennreal_less_iff reorient: ennreal_plus)
       moreover
       have "?K \<subseteq> (\<Union>i. D i)" using K by auto
       moreover
@@ -332,9 +332,9 @@
         from INF_approx_ennreal[OF \<open>0 < e/(2 powr Suc i)\<close> this]
         show "\<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
           using \<open>0<e\<close>
-          by (auto simp: emeasure_eq_measure measure_nonneg sum_nonneg ennreal_less_iff ennreal_plus[symmetric] ennreal_minus
+          by (auto simp: emeasure_eq_measure sum_nonneg ennreal_less_iff ennreal_minus
                          finite_measure_mono sb
-                   simp del: ennreal_plus)
+                   reorient: ennreal_plus)
       qed
       then obtain U where U: "\<And>i. D i \<subseteq> U i" "\<And>i. open (U i)"
         "\<And>i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)"
@@ -367,7 +367,7 @@
       also have "\<dots> = ennreal e"
         by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
       finally have "emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e"
-        using \<open>0<e\<close> by (simp add: emeasure_eq_measure ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
+        using \<open>0<e\<close> by (simp add: emeasure_eq_measure reorient: ennreal_plus)
       moreover
       have "(\<Union>i. D i) \<subseteq> ?U" using U by auto
       moreover
--- a/src/HOL/Analysis/Set_Integral.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Set_Integral.thy	Thu May 03 15:07:14 2018 +0200
@@ -869,7 +869,7 @@
     2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)" by simp
 
   have le: "ennreal (norm (F n x - f x)) \<le> ennreal (norm (f x)) + ennreal (norm (F n x))" for n x
-    by (simp add: norm_minus_commute norm_triangle_ineq4 ennreal_plus[symmetric] ennreal_minus del: ennreal_plus)
+    by (simp add: norm_minus_commute norm_triangle_ineq4 ennreal_minus reorient: ennreal_plus)
   then have le2: "(\<integral>\<^sup>+ x. ennreal (norm (F n x - f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) + ennreal (norm (F n x)) \<partial>M)" for n
     by (rule nn_integral_mono)
 
@@ -882,7 +882,7 @@
   proof (intro arg_cong[where f=liminf] ext)
     fix n
     have "\<And>x. ennreal(G n x) = ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x))"
-      unfolding G_def by (simp add: ennreal_plus[symmetric] ennreal_minus del: ennreal_plus)
+      unfolding G_def by (simp add: ennreal_minus reorient: ennreal_plus)
     moreover have "(\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x)) \<partial>M)
             = (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)"
     proof (rule nn_integral_diff)
--- a/src/HOL/Analysis/Sigma_Algebra.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Sigma_Algebra.thy	Thu May 03 15:07:14 2018 +0200
@@ -2084,7 +2084,7 @@
   have "(\<lambda>x. g (f x)) \<in> X \<rightarrow> space M" "\<And>A. A \<inter> f -` Y \<inter> X = A \<inter> X"
     using * by auto
   with * show "sets ?VV = sets ?V"
-    by (simp add: sets_vimage_algebra2 ex_simps[symmetric] vimage_comp comp_def del: ex_simps)
+    by (simp add: sets_vimage_algebra2 vimage_comp comp_def reorient: ex_simps)
 qed (simp add: vimage_algebra_def emeasure_sigma)
 
 subsubsection \<open>Restricted Space Sigma Algebra\<close>
--- a/src/HOL/Analysis/Starlike.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Starlike.thy	Thu May 03 15:07:14 2018 +0200
@@ -18,10 +18,7 @@
   where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
 
 lemma midpoint_idem [simp]: "midpoint x x = x"
-  unfolding midpoint_def
-  unfolding scaleR_right_distrib
-  unfolding scaleR_left_distrib[symmetric]
-  by auto
+  unfolding midpoint_def  by simp
 
 lemma midpoint_sym: "midpoint a b = midpoint b a"
   unfolding midpoint_def by (auto simp add: scaleR_right_distrib)
@@ -927,61 +924,49 @@
 proof (cases "a = b")
   case True
   then show ?thesis
-    unfolding between_def split_conv
-    by (auto simp add: dist_commute)
+    by (auto simp add: between_def dist_commute)
 next
   case False
   then have Fal: "norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0"
     by auto
   have *: "\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)"
     by (auto simp add: algebra_simps)
-  show ?thesis
-    unfolding between_def split_conv closed_segment_def mem_Collect_eq
-    apply rule
-    apply (elim exE conjE)
-    apply (subst dist_triangle_eq)
+  have "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" if "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" for u
   proof -
-    fix u
-    assume as: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1"
-    then have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
-      unfolding as(1) by (auto simp add:algebra_simps)
+    have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
+      unfolding that(1) by (auto simp add:algebra_simps)
     show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
-      unfolding norm_minus_commute[of x a] * using as(2,3)
+      unfolding norm_minus_commute[of x a] * using \<open>0 \<le> u\<close> \<open>u \<le> 1\<close>
       by (auto simp add: field_simps)
-  next
-    assume as: "dist a b = dist a x + dist x b"
-    have "norm (a - x) / norm (a - b) \<le> 1"
-      using Fal2 unfolding as[unfolded dist_norm] norm_ge_zero by auto
-    then show "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1"
-      apply (rule_tac x="dist a x / dist a b" in exI)
-      unfolding dist_norm
-      apply (subst euclidean_eq_iff)
-      apply rule
-      defer
-      apply rule
-      prefer 3
-      apply rule
-    proof -
-      fix i :: 'a
-      assume i: "i \<in> Basis"
-      have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i =
-        ((norm (a - b) - norm (a - x)) * (a \<bullet> i) + norm (a - x) * (b \<bullet> i)) / norm (a - b)"
-        using Fal by (auto simp add: field_simps inner_simps)
-      also have "\<dots> = x\<bullet>i"
-        apply (rule divide_eq_imp[OF Fal])
-        unfolding as[unfolded dist_norm]
-        using as[unfolded dist_triangle_eq]
-        apply -
-        apply (subst (asm) euclidean_eq_iff)
-        using i
-        apply (erule_tac x=i in ballE)
-        apply (auto simp add: field_simps inner_simps)
-        done
-      finally show "x \<bullet> i =
-        ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i"
-        by auto
-    qed (insert Fal2, auto)
   qed
+  moreover have "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" if "dist a b = dist a x + dist x b" 
+  proof -
+    let ?\<beta> = "norm (a - x) / norm (a - b)"
+    show "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1"
+    proof (intro exI conjI)
+      show "?\<beta> \<le> 1"
+        using Fal2 unfolding that[unfolded dist_norm] norm_ge_zero by auto
+      show "x = (1 - ?\<beta>) *\<^sub>R a + (?\<beta>) *\<^sub>R b"
+      proof (subst euclidean_eq_iff; intro ballI)
+        fix i :: 'a
+        assume i: "i \<in> Basis"
+        have "((1 - ?\<beta>) *\<^sub>R a + (?\<beta>) *\<^sub>R b) \<bullet> i 
+              = ((norm (a - b) - norm (a - x)) * (a \<bullet> i) + norm (a - x) * (b \<bullet> i)) / norm (a - b)"
+          using Fal by (auto simp add: field_simps inner_simps)
+        also have "\<dots> = x\<bullet>i"
+          apply (rule divide_eq_imp[OF Fal])
+          unfolding that[unfolded dist_norm]
+          using that[unfolded dist_triangle_eq] i
+          apply (subst (asm) euclidean_eq_iff)
+           apply (auto simp add: field_simps inner_simps)
+          done
+        finally show "x \<bullet> i = ((1 - ?\<beta>) *\<^sub>R a + (?\<beta>) *\<^sub>R b) \<bullet> i"
+          by auto
+      qed
+    qed (use Fal2 in auto)
+  qed
+  ultimately show ?thesis
+    by (force simp add: between_def closed_segment_def dist_triangle_eq)
 qed
 
 lemma between_midpoint:
@@ -993,10 +978,7 @@
     by auto
   show ?t1 ?t2
     unfolding between midpoint_def dist_norm
-    apply(rule_tac[!] *)
-    unfolding euclidean_eq_iff[where 'a='a]
-    apply (auto simp add: field_simps inner_simps)
-    done
+    by (auto simp add: field_simps inner_simps euclidean_eq_iff[where 'a='a] intro!: *)
 qed
 
 lemma between_mem_convex_hull:
@@ -1061,25 +1043,23 @@
 subsection%unimportant \<open>Shrinking towards the interior of a convex set\<close>
 
 lemma mem_interior_convex_shrink:
-  fixes s :: "'a::euclidean_space set"
-  assumes "convex s"
-    and "c \<in> interior s"
-    and "x \<in> s"
+  fixes S :: "'a::euclidean_space set"
+  assumes "convex S"
+    and "c \<in> interior S"
+    and "x \<in> S"
     and "0 < e"
     and "e \<le> 1"
-  shows "x - e *\<^sub>R (x - c) \<in> interior s"
-proof -
-  obtain d where "d > 0" and d: "ball c d \<subseteq> s"
+  shows "x - e *\<^sub>R (x - c) \<in> interior S"
+proof -
+  obtain d where "d > 0" and d: "ball c d \<subseteq> S"
     using assms(2) unfolding mem_interior by auto
   show ?thesis
     unfolding mem_interior
-    apply (rule_tac x="e*d" in exI)
-    apply rule
-    defer
-    unfolding subset_eq Ball_def mem_ball
-  proof (rule, rule)
+  proof (intro exI subsetI conjI)
     fix y
-    assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d"
+    assume "y \<in> ball (x - e *\<^sub>R (x - c)) (e*d)"
+    then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d"
+      by simp
     have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
       using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
     have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \<bar>1/e\<bar> * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
@@ -1093,7 +1073,7 @@
     also have "\<dots> < d"
       using as[unfolded dist_norm] and \<open>e > 0\<close>
       by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
-    finally show "y \<in> s"
+    finally show "y \<in> S"
       apply (subst *)
       apply (rule assms(1)[unfolded convex_alt,rule_format])
       apply (rule d[unfolded subset_eq,rule_format])
@@ -1105,18 +1085,18 @@
 qed
 
 lemma mem_interior_closure_convex_shrink:
-  fixes s :: "'a::euclidean_space set"
-  assumes "convex s"
-    and "c \<in> interior s"
-    and "x \<in> closure s"
+  fixes S :: "'a::euclidean_space set"
+  assumes "convex S"
+    and "c \<in> interior S"
+    and "x \<in> closure S"
     and "0 < e"
     and "e \<le> 1"
-  shows "x - e *\<^sub>R (x - c) \<in> interior s"
-proof -
-  obtain d where "d > 0" and d: "ball c d \<subseteq> s"
+  shows "x - e *\<^sub>R (x - c) \<in> interior S"
+proof -
+  obtain d where "d > 0" and d: "ball c d \<subseteq> S"
     using assms(2) unfolding mem_interior by auto
-  have "\<exists>y\<in>s. norm (y - x) * (1 - e) < e * d"
-  proof (cases "x \<in> s")
+  have "\<exists>y\<in>S. norm (y - x) * (1 - e) < e * d"
+  proof (cases "x \<in> S")
     case True
     then show ?thesis
       using \<open>e > 0\<close> \<open>d > 0\<close>
@@ -1125,12 +1105,12 @@
       done
   next
     case False
-    then have x: "x islimpt s"
+    then have x: "x islimpt S"
       using assms(3)[unfolded closure_def] by auto
     show ?thesis
     proof (cases "e = 1")
       case True
-      obtain y where "y \<in> s" "y \<noteq> x" "dist y x < 1"
+      obtain y where "y \<in> S" "y \<noteq> x" "dist y x < 1"
         using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
       then show ?thesis
         apply (rule_tac x=y in bexI)
@@ -1142,7 +1122,7 @@
       case False
       then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
         using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto
-      then obtain y where "y \<in> s" "y \<noteq> x" "dist y x < e * d / (1 - e)"
+      then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)"
         using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
       then show ?thesis
         apply (rule_tac x=y in bexI)
@@ -1152,24 +1132,20 @@
         done
     qed
   qed
-  then obtain y where "y \<in> s" and y: "norm (y - x) * (1 - e) < e * d"
+  then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d"
     by auto
   define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)"
   have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
     unfolding z_def using \<open>e > 0\<close>
     by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
-  have "z \<in> interior s"
+  have "z \<in> interior S"
     apply (rule interior_mono[OF d,unfolded subset_eq,rule_format])
     unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5)
     apply (auto simp add:field_simps norm_minus_commute)
     done
   then show ?thesis
     unfolding *
-    apply -
-    apply (rule mem_interior_convex_shrink)
-    using assms(1,4-5) \<open>y\<in>s\<close>
-    apply auto
-    done
+    using mem_interior_convex_shrink \<open>y \<in> S\<close> assms by blast
 qed
 
 lemma in_interior_closure_convex_segment:
@@ -1255,23 +1231,20 @@
 subsection%unimportant \<open>Some obvious but surprisingly hard simplex lemmas\<close>
 
 lemma simplex:
-  assumes "finite s"
-    and "0 \<notin> s"
-  shows "convex hull (insert 0 s) =
-    {y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s \<le> 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y)}"
-  unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]]
-  apply (rule set_eqI, rule)
-  unfolding mem_Collect_eq
-  apply (erule_tac[!] exE)
-  apply (erule_tac[!] conjE)+
-  unfolding sum_clauses(2)[OF \<open>finite s\<close>]
-  apply (rule_tac x=u in exI)
-  defer
-  apply (rule_tac x="\<lambda>x. if x = 0 then 1 - sum u s else u x" in exI)
-  using assms(2)
-  unfolding if_smult and sum_delta_notmem[OF assms(2)]
-  apply auto
-  done
+  assumes "finite S"
+    and "0 \<notin> S"
+  shows "convex hull (insert 0 S) = {y. \<exists>u. (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S \<le> 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) S = y}"
+proof (simp add: convex_hull_finite set_eq_iff assms, safe)
+  fix x and u :: "'a \<Rightarrow> real"
+  assume "0 \<le> u 0" "\<forall>x\<in>S. 0 \<le> u x" "u 0 + sum u S = 1"
+  then show "\<exists>v. (\<forall>x\<in>S. 0 \<le> v x) \<and> sum v S \<le> 1 \<and> (\<Sum>x\<in>S. v x *\<^sub>R x) = (\<Sum>x\<in>S. u x *\<^sub>R x)"
+    by force
+next
+  fix x and u :: "'a \<Rightarrow> real"
+  assume "\<forall>x\<in>S. 0 \<le> u x" "sum u S \<le> 1"
+  then show "\<exists>v. 0 \<le> v 0 \<and> (\<forall>x\<in>S. 0 \<le> v x) \<and> v 0 + sum v S = 1 \<and> (\<Sum>x\<in>S. v x *\<^sub>R x) = (\<Sum>x\<in>S. u x *\<^sub>R x)"
+    by (rule_tac x="\<lambda>x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult)
+qed
 
 lemma substd_simplex:
   assumes d: "d \<subseteq> Basis"
@@ -1286,50 +1259,27 @@
     by (blast intro: finite_subset finite_Basis)
   show ?thesis
     unfolding simplex[OF \<open>finite d\<close> \<open>0 \<notin> ?p\<close>]
-    apply (rule set_eqI)
-    unfolding mem_Collect_eq
-    apply rule
-    apply (elim exE conjE)
-    apply (erule_tac[2] conjE)+
-  proof -
-    fix x :: "'a::euclidean_space"
-    fix u
-    assume as: "\<forall>x\<in>?D. 0 \<le> u x" "sum u ?D \<le> 1" "(\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
-    have *: "\<forall>i\<in>Basis. i \<in> d \<longrightarrow> u i = x\<bullet>i"
-      and "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
-      using as(3)
-      unfolding substdbasis_expansion_unique[OF assms]
-      by auto
-    then have **: "sum u ?D = sum ((\<bullet>) x) ?D"
-      apply -
-      apply (rule sum.cong)
-      using assms
-      apply auto
-      done
-    have "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum ((\<bullet>) x) ?D \<le> 1"
-    proof (rule,rule)
-      fix i :: 'a
-      assume i: "i \<in> Basis"
-      have "i \<in> d \<Longrightarrow> 0 \<le> x\<bullet>i"
-        unfolding *[rule_format,OF i,symmetric]
-         apply (rule_tac as(1)[rule_format])
-         apply auto
-         done
-      moreover have "i \<notin> d \<Longrightarrow> 0 \<le> x\<bullet>i"
-        using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close>[rule_format, OF i] by auto
-      ultimately show "0 \<le> x\<bullet>i" by auto
-    qed (insert as(2)[unfolded **], auto)
-    then show "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum ((\<bullet>) x) ?D \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
-      using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close> by auto
+  proof (intro set_eqI; safe)
+    fix u :: "'a \<Rightarrow> real"
+    assume as: "\<forall>x\<in>?D. 0 \<le> u x" "sum u ?D \<le> 1" 
+    let ?x = "(\<Sum>x\<in>?D. u x *\<^sub>R x)"
+    have ind: "\<forall>i\<in>Basis. i \<in> d \<longrightarrow> u i = ?x \<bullet> i"
+      and notind: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> ?x \<bullet> i = 0)"
+      using substdbasis_expansion_unique[OF assms] by blast+
+    then have **: "sum u ?D = sum ((\<bullet>) ?x) ?D"
+      using assms by (auto intro!: sum.cong)
+    show "0 \<le> ?x \<bullet> i" if "i \<in> Basis" for i
+      using as(1) ind notind that by fastforce
+    show "sum ((\<bullet>) ?x) ?D \<le> 1"
+      using "**" as(2) by linarith
+    show "?x \<bullet> i = 0" if "i \<in> Basis" "i \<notin> d" for i
+      using notind that by blast
   next
-    fix x :: "'a::euclidean_space"
-    assume as: "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "sum ((\<bullet>) x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
-    show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> sum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
-      using as d
-      unfolding substdbasis_expansion_unique[OF assms]
-      apply (rule_tac x="inner x" in exI)
-      apply auto
-      done
+    fix x 
+    assume "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "sum ((\<bullet>) x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
+    with d show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> sum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
+      unfolding substdbasis_expansion_unique[OF assms] 
+      by (rule_tac x="inner x" in exI) auto
   qed
 qed
 
@@ -1341,27 +1291,18 @@
 lemma interior_std_simplex:
   "interior (convex hull (insert 0 Basis)) =
     {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 < x\<bullet>i) \<and> sum (\<lambda>i. x\<bullet>i) Basis < 1}"
-  apply (rule set_eqI)
-  unfolding mem_interior std_simplex
-  unfolding subset_eq mem_Collect_eq Ball_def mem_ball
-  unfolding Ball_def[symmetric]
-  apply rule
-  apply (elim exE conjE)
-  defer
-  apply (erule conjE)
-proof -
+  unfolding set_eq_iff mem_interior std_simplex
+proof (intro allI iffI CollectI; clarify)
   fix x :: 'a
   fix e
-  assume "e > 0" and as: "\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x\<in>Basis. 0 \<le> xa \<bullet> x) \<and> sum ((\<bullet>) xa) Basis \<le> 1"
-  show "(\<forall>xa\<in>Basis. 0 < x \<bullet> xa) \<and> sum ((\<bullet>) x) Basis < 1"
-    apply safe
-  proof -
+  assume "e > 0" and as: "ball x e \<subseteq> {x. (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i) \<and> sum ((\<bullet>) x) Basis \<le> 1}"
+  show "(\<forall>i\<in>Basis. 0 < x \<bullet> i) \<and> sum ((\<bullet>) x) Basis < 1"
+  proof safe
     fix i :: 'a
     assume i: "i \<in> Basis"
     then show "0 < x \<bullet> i"
-      using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and \<open>e > 0\<close>
-      unfolding dist_norm
-      by (auto elim!: ballE[where x=i] simp: inner_simps)
+      using as[THEN subsetD[where c="x - (e / 2) *\<^sub>R i"]] and \<open>e > 0\<close> 
+      by (force simp add: inner_simps)
   next
     have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) < e" using \<open>e > 0\<close>
       unfolding dist_norm
@@ -1371,42 +1312,31 @@
       by (auto simp: SOME_Basis inner_Basis inner_simps)
     then have *: "sum ((\<bullet>) (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis =
       sum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis"
-      apply (rule_tac sum.cong)
-      apply auto
-      done
+      by (auto simp: intro!: sum.cong)
     have "sum ((\<bullet>) x) Basis < sum ((\<bullet>) (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis"
-      unfolding * sum.distrib
-      using \<open>e > 0\<close> DIM_positive[where 'a='a]
-      apply (subst sum.delta')
-      apply (auto simp: SOME_Basis)
-      done
+      using \<open>e > 0\<close> DIM_positive by (auto simp: SOME_Basis sum.distrib *)
     also have "\<dots> \<le> 1"
-      using **
-      apply (drule_tac as[rule_format])
-      apply auto
-      done
+      using ** as by force
     finally show "sum ((\<bullet>) x) Basis < 1" by auto
-  qed
+  qed 
 next
   fix x :: 'a
   assume as: "\<forall>i\<in>Basis. 0 < x \<bullet> i" "sum ((\<bullet>) x) Basis < 1"
   obtain a :: 'b where "a \<in> UNIV" using UNIV_witness ..
   let ?d = "(1 - sum ((\<bullet>) x) Basis) / real (DIM('a))"
-  show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> sum ((\<bullet>) y) Basis \<le> 1"
-  proof (rule_tac x="min (Min (((\<bullet>) x) ` Basis)) D" for D in exI, intro conjI impI allI)
+  show "\<exists>e>0. ball x e \<subseteq> {x. (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i) \<and> sum ((\<bullet>) x) Basis \<le> 1}"
+  proof (rule_tac x="min (Min (((\<bullet>) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI)
     fix y
-    assume y: "dist x y < min (Min ((\<bullet>) x ` Basis)) ?d"
+    assume y: "y \<in> ball x (min (Min ((\<bullet>) x ` Basis)) ?d)"
     have "sum ((\<bullet>) y) Basis \<le> sum (\<lambda>i. x\<bullet>i + ?d) Basis"
     proof (rule sum_mono)
       fix i :: 'a
       assume i: "i \<in> Basis"
-      then have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d"
-        apply -
-        apply (rule le_less_trans)
-        using Basis_le_norm[OF i, of "y - x"]
-        using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2]
-        apply (auto simp add: norm_minus_commute inner_diff_left)
-        done
+      have "\<bar>y\<bullet>i - x\<bullet>i\<bar> \<le> norm (y - x)"
+        by (metis Basis_le_norm i inner_commute inner_diff_right)
+      also have "... < ?d"
+        using y by (simp add: dist_norm norm_minus_commute)
+      finally have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d" .
       then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
     qed
     also have "\<dots> \<le> 1"
@@ -1417,12 +1347,11 @@
     proof safe
       fix i :: 'a
       assume i: "i \<in> Basis"
-      have "norm (x - y) < x\<bullet>i"
-        apply (rule less_le_trans)
-        apply (rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1])
-        using i
-        apply auto
-        done
+      have "norm (x - y) < MINIMUM Basis ((\<bullet>) x)"
+        using y by (auto simp: dist_norm less_eq_real_def)
+      also have "... \<le> x\<bullet>i"
+        using i by auto
+      finally have "norm (x - y) < x\<bullet>i" .
       then show "0 \<le> y\<bullet>i"
         using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i]
         by (auto simp: inner_simps)
@@ -1472,82 +1401,70 @@
 qed
 
 lemma rel_interior_substd_simplex:
-  assumes d: "d \<subseteq> Basis"
-  shows "rel_interior (convex hull (insert 0 d)) =
-    {x::'a::euclidean_space. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
+  assumes D: "D \<subseteq> Basis"
+  shows "rel_interior (convex hull (insert 0 D)) =
+    {x::'a::euclidean_space. (\<forall>i\<in>D. 0 < x\<bullet>i) \<and> (\<Sum>i\<in>D. x\<bullet>i) < 1 \<and> (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> x\<bullet>i = 0)}"
   (is "rel_interior (convex hull (insert 0 ?p)) = ?s")
 proof -
-  have "finite d"
-    apply (rule finite_subset)
-    using assms
-    apply auto
-    done
+  have "finite D"
+    using D finite_Basis finite_subset by blast
   show ?thesis
-  proof (cases "d = {}")
+  proof (cases "D = {}")
     case True
     then show ?thesis
       using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto
   next
     case False
     have h0: "affine hull (convex hull (insert 0 ?p)) =
-      {x::'a::euclidean_space. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
+      {x::'a::euclidean_space. (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> x\<bullet>i = 0)}"
       using affine_hull_convex_hull affine_hull_substd_basis assms by auto
-    have aux: "\<And>x::'a. \<forall>i\<in>Basis. (\<forall>i\<in>d. 0 \<le> x\<bullet>i) \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i"
+    have aux: "\<And>x::'a. \<forall>i\<in>Basis. (\<forall>i\<in>D. 0 \<le> x\<bullet>i) \<and> (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i"
       by auto
     {
       fix x :: "'a::euclidean_space"
       assume x: "x \<in> rel_interior (convex hull (insert 0 ?p))"
-      then obtain e where e0: "e > 0" and
-        "ball x e \<inter> {xa. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0)} \<subseteq> convex hull (insert 0 ?p)"
+      then obtain e where "e > 0" and
+        "ball x e \<inter> {xa. (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> xa\<bullet>i = 0)} \<subseteq> convex hull (insert 0 ?p)"
         using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto
-      then have as: "\<forall>xa. dist x xa < e \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0) \<longrightarrow>
-        (\<forall>i\<in>d. 0 \<le> xa \<bullet> i) \<and> sum ((\<bullet>) xa) d \<le> 1"
+      then have as [rule_format]: "\<And>y. dist x y < e \<and> (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> y\<bullet>i = 0) \<longrightarrow>
+        (\<forall>i\<in>D. 0 \<le> y \<bullet> i) \<and> sum ((\<bullet>) y) D \<le> 1"
         unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto
-      have x0: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
+      have x0: "(\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> x\<bullet>i = 0)"
         using x rel_interior_subset  substd_simplex[OF assms] by auto
-      have "(\<forall>i\<in>d. 0 < x \<bullet> i) \<and> sum ((\<bullet>) x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
-        apply rule
-        apply rule
-      proof -
+      have "(\<forall>i\<in>D. 0 < x \<bullet> i) \<and> sum ((\<bullet>) x) D < 1 \<and> (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> x\<bullet>i = 0)"
+      proof (intro conjI ballI)
         fix i :: 'a
-        assume "i \<in> d"
-        then have "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R i) \<bullet> ia"
+        assume "i \<in> D"
+        then have "\<forall>j\<in>D. 0 \<le> (x - (e / 2) *\<^sub>R i) \<bullet> j"
           apply -
-          apply (rule as[rule_format,THEN conjunct1])
-          unfolding dist_norm
-          using d \<open>e > 0\<close> x0
-          apply (auto simp: inner_simps inner_Basis)
+          apply (rule as[THEN conjunct1])
+          using D \<open>e > 0\<close> x0
+          apply (auto simp: dist_norm inner_simps inner_Basis)
           done
         then show "0 < x \<bullet> i"
-          apply (erule_tac x=i in ballE)
-          using \<open>e > 0\<close> \<open>i \<in> d\<close> d
-          apply (auto simp: inner_simps inner_Basis)
-          done
+          using \<open>e > 0\<close> \<open>i \<in> D\<close> D  by (force simp: inner_simps inner_Basis)
       next
-        obtain a where a: "a \<in> d"
-          using \<open>d \<noteq> {}\<close> by auto
+        obtain a where a: "a \<in> D"
+          using \<open>D \<noteq> {}\<close> by auto
         then have **: "dist x (x + (e / 2) *\<^sub>R a) < e"
-          using \<open>e > 0\<close> norm_Basis[of a] d
+          using \<open>e > 0\<close> norm_Basis[of a] D
           unfolding dist_norm
           by auto
         have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = x\<bullet>i + (if i = a then e/2 else 0)"
-          using a d by (auto simp: inner_simps inner_Basis)
-        then have *: "sum ((\<bullet>) (x + (e / 2) *\<^sub>R a)) d =
-          sum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d"
-          using d by (intro sum.cong) auto
+          using a D by (auto simp: inner_simps inner_Basis)
+        then have *: "sum ((\<bullet>) (x + (e / 2) *\<^sub>R a)) D =
+          sum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) D"
+          using D by (intro sum.cong) auto
         have "a \<in> Basis"
-          using \<open>a \<in> d\<close> d by auto
-        then have h1: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)"
-          using x0 d \<open>a\<in>d\<close> by (auto simp add: inner_add_left inner_Basis)
-        have "sum ((\<bullet>) x) d < sum ((\<bullet>) (x + (e / 2) *\<^sub>R a)) d"
-          unfolding * sum.distrib
-          using \<open>e > 0\<close> \<open>a \<in> d\<close>
-          using \<open>finite d\<close>
-          by (auto simp add: sum.delta')
+          using \<open>a \<in> D\<close> D by auto
+        then have h1: "(\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)"
+          using x0 D \<open>a\<in>D\<close> by (auto simp add: inner_add_left inner_Basis)
+        have "sum ((\<bullet>) x) D < sum ((\<bullet>) (x + (e / 2) *\<^sub>R a)) D"
+          using \<open>e > 0\<close> \<open>a \<in> D\<close> \<open>finite D\<close> by (auto simp add: * sum.distrib)
         also have "\<dots> \<le> 1"
           using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"]
           by auto
-        finally show "sum ((\<bullet>) x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
+        finally show "sum ((\<bullet>) x) D < 1" "\<And>i. i\<in>Basis \<Longrightarrow> i \<notin> D \<longrightarrow> x\<bullet>i = 0"
           using x0 by auto
       qed
     }
@@ -1557,63 +1474,62 @@
       assume as: "x \<in> ?s"
       have "\<forall>i. 0 < x\<bullet>i \<or> 0 = x\<bullet>i \<longrightarrow> 0 \<le> x\<bullet>i"
         by auto
-      moreover have "\<forall>i. i \<in> d \<or> i \<notin> d" by auto
+      moreover have "\<forall>i. i \<in> D \<or> i \<notin> D" by auto
       ultimately
-      have "\<forall>i. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<forall>i. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i"
+      have "\<forall>i. (\<forall>i\<in>D. 0 < x\<bullet>i) \<and> (\<forall>i. i \<notin> D \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i"
         by metis
       then have h2: "x \<in> convex hull (insert 0 ?p)"
         using as assms
         unfolding substd_simplex[OF assms] by fastforce
-      obtain a where a: "a \<in> d"
-        using \<open>d \<noteq> {}\<close> by auto
-      let ?d = "(1 - sum ((\<bullet>) x) d) / real (card d)"
-      have "0 < card d" using \<open>d \<noteq> {}\<close> \<open>finite d\<close>
+      obtain a where a: "a \<in> D"
+        using \<open>D \<noteq> {}\<close> by auto
+      let ?d = "(1 - sum ((\<bullet>) x) D) / real (card D)"
+      have "0 < card D" using \<open>D \<noteq> {}\<close> \<open>finite D\<close>
         by (simp add: card_gt_0_iff)
-      have "Min (((\<bullet>) x) ` d) > 0"
-        using as \<open>d \<noteq> {}\<close> \<open>finite d\<close> by (simp add: Min_gr_iff)
-      moreover have "?d > 0" using as using \<open>0 < card d\<close> by auto
-      ultimately have h3: "min (Min (((\<bullet>) x) ` d)) ?d > 0"
+      have "Min (((\<bullet>) x) ` D) > 0"
+        using as \<open>D \<noteq> {}\<close> \<open>finite D\<close> by (simp add: Min_gr_iff)
+      moreover have "?d > 0" using as using \<open>0 < card D\<close> by auto
+      ultimately have h3: "min (Min (((\<bullet>) x) ` D)) ?d > 0"
         by auto
 
       have "x \<in> rel_interior (convex hull (insert 0 ?p))"
         unfolding rel_interior_ball mem_Collect_eq h0
         apply (rule,rule h2)
         unfolding substd_simplex[OF assms]
-        apply (rule_tac x="min (Min (((\<bullet>) x) ` d)) ?d" in exI)
+        apply (rule_tac x="min (Min (((\<bullet>) x) ` D)) ?d" in exI)
         apply (rule, rule h3)
         apply safe
         unfolding mem_ball
       proof -
         fix y :: 'a
-        assume y: "dist x y < min (Min ((\<bullet>) x ` d)) ?d"
-        assume y2: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> y\<bullet>i = 0"
-        have "sum ((\<bullet>) y) d \<le> sum (\<lambda>i. x\<bullet>i + ?d) d"
+        assume y: "dist x y < min (Min ((\<bullet>) x ` D)) ?d"
+        assume y2: "\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> y\<bullet>i = 0"
+        have "sum ((\<bullet>) y) D \<le> sum (\<lambda>i. x\<bullet>i + ?d) D"
         proof (rule sum_mono)
           fix i
-          assume "i \<in> d"
-          with d have i: "i \<in> Basis"
+          assume "i \<in> D"
+          with D have i: "i \<in> Basis"
             by auto
-          have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d"
-            apply (rule le_less_trans)
-            using Basis_le_norm[OF i, of "y - x"]
-            using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2]
-            apply (auto simp add: norm_minus_commute inner_simps)
-            done
+          have "\<bar>y\<bullet>i - x\<bullet>i\<bar> \<le> norm (y - x)"
+            by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl)
+          also have "... < ?d"
+            by (metis dist_norm min_less_iff_conj norm_minus_commute y)
+          finally have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d" .
           then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
         qed
         also have "\<dots> \<le> 1"
-          unfolding sum.distrib sum_constant  using \<open>0 < card d\<close>
+          unfolding sum.distrib sum_constant  using \<open>0 < card D\<close>
           by auto
-        finally show "sum ((\<bullet>) y) d \<le> 1" .
+        finally show "sum ((\<bullet>) y) D \<le> 1" .
 
         fix i :: 'a
         assume i: "i \<in> Basis"
         then show "0 \<le> y\<bullet>i"
-        proof (cases "i\<in>d")
+        proof (cases "i\<in>D")
           case True
           have "norm (x - y) < x\<bullet>i"
             using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
-            using Min_gr_iff[of "(\<bullet>) x ` d" "norm (x - y)"] \<open>0 < card d\<close> \<open>i \<in> d\<close>
+            using Min_gr_iff[of "(\<bullet>) x ` D" "norm (x - y)"] \<open>0 < card D\<close> \<open>i \<in> D\<close>
             by (simp add: card_gt_0_iff)
           then show "0 \<le> y\<bullet>i"
             using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format]
@@ -1622,36 +1538,36 @@
       qed
     }
     ultimately have
-      "\<And>x. x \<in> rel_interior (convex hull insert 0 d) \<longleftrightarrow>
-        x \<in> {x. (\<forall>i\<in>d. 0 < x \<bullet> i) \<and> sum ((\<bullet>) x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)}"
+      "\<And>x. x \<in> rel_interior (convex hull insert 0 D) \<longleftrightarrow>
+        x \<in> {x. (\<forall>i\<in>D. 0 < x \<bullet> i) \<and> sum ((\<bullet>) x) D < 1 \<and> (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> x \<bullet> i = 0)}"
       by blast
     then show ?thesis by (rule set_eqI)
   qed
 qed
 
 lemma rel_interior_substd_simplex_nonempty:
-  assumes "d \<noteq> {}"
-    and "d \<subseteq> Basis"
+  assumes "D \<noteq> {}"
+    and "D \<subseteq> Basis"
   obtains a :: "'a::euclidean_space"
-    where "a \<in> rel_interior (convex hull (insert 0 d))"
-proof -
-  let ?D = d
-  let ?a = "sum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D"
-  have "finite d"
+    where "a \<in> rel_interior (convex hull (insert 0 D))"
+proof -
+  let ?D = D
+  let ?a = "sum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) ?D"
+  have "finite D"
     apply (rule finite_subset)
     using assms(2)
     apply auto
     done
-  then have d1: "0 < real (card d)"
-    using \<open>d \<noteq> {}\<close> by auto
+  then have d1: "0 < real (card D)"
+    using \<open>D \<noteq> {}\<close> by auto
   {
     fix i
-    assume "i \<in> d"
-    have "?a \<bullet> i = inverse (2 * real (card d))"
-      apply (rule trans[of _ "sum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
+    assume "i \<in> D"
+    have "?a \<bullet> i = inverse (2 * real (card D))"
+      apply (rule trans[of _ "sum (\<lambda>j. if i = j then inverse (2 * real (card D)) else 0) ?D"])
       unfolding inner_sum_left
       apply (rule sum.cong)
-      using \<open>i \<in> d\<close> \<open>finite d\<close> sum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
+      using \<open>i \<in> D\<close> \<open>finite D\<close> sum.delta'[of D i "(\<lambda>k. inverse (2 * real (card D)))"]
         d1 assms(2)
       by (auto simp: inner_Basis set_rev_mp[OF _ assms(2)])
   }
@@ -1661,14 +1577,14 @@
     unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq
   proof safe
     fix i
-    assume "i \<in> d"
-    have "0 < inverse (2 * real (card d))"
+    assume "i \<in> D"
+    have "0 < inverse (2 * real (card D))"
       using d1 by auto
-    also have "\<dots> = ?a \<bullet> i" using **[of i] \<open>i \<in> d\<close>
+    also have "\<dots> = ?a \<bullet> i" using **[of i] \<open>i \<in> D\<close>
       by auto
     finally show "0 < ?a \<bullet> i" by auto
   next
-    have "sum ((\<bullet>) ?a) ?D = sum (\<lambda>i. inverse (2 * real (card d))) ?D"
+    have "sum ((\<bullet>) ?a) ?D = sum (\<lambda>i. inverse (2 * real (card D))) ?D"
       by (rule sum.cong) (rule refl, rule **)
     also have "\<dots> < 1"
       unfolding sum_constant divide_real_def[symmetric]
@@ -1676,22 +1592,22 @@
     finally show "sum ((\<bullet>) ?a) ?D < 1" by auto
   next
     fix i
-    assume "i \<in> Basis" and "i \<notin> d"
-    have "?a \<in> span d"
-    proof (rule span_sum[of d "(\<lambda>b. b /\<^sub>R (2 * real (card d)))" d])
+    assume "i \<in> Basis" and "i \<notin> D"
+    have "?a \<in> span D"
+    proof (rule span_sum[of D "(\<lambda>b. b /\<^sub>R (2 * real (card D)))" D])
       {
         fix x :: "'a::euclidean_space"
-        assume "x \<in> d"
-        then have "x \<in> span d"
-          using span_base[of _ "d"] by auto
-        then have "x /\<^sub>R (2 * real (card d)) \<in> span d"
-          using span_scale[of x "d" "(inverse (real (card d)) / 2)"] by auto
+        assume "x \<in> D"
+        then have "x \<in> span D"
+          using span_superset[of _ "D"] by auto
+        then have "x /\<^sub>R (2 * real (card D)) \<in> span D"
+          using span_mul[of x "D" "(inverse (real (card D)) / 2)"] by auto
       }
-      then show "\<And>x. x\<in>d \<Longrightarrow> x /\<^sub>R (2 * real (card d)) \<in> span d"
+      then show "\<And>x. x\<in>D \<Longrightarrow> x /\<^sub>R (2 * real (card D)) \<in> span D"
         by auto
     qed
     then show "?a \<bullet> i = 0 "
-      using \<open>i \<notin> d\<close> unfolding span_substd_basis[OF assms(2)] using \<open>i \<in> Basis\<close> by auto
+      using \<open>i \<notin> D\<close> unfolding span_substd_basis[OF assms(2)] using \<open>i \<in> Basis\<close> by auto
   qed
 qed
 
@@ -1708,9 +1624,8 @@
   then show ?thesis using rel_interior_sing by auto
 next
   case False
-  obtain B where B: "independent B \<and> B \<le> S \<and> S \<le> span B \<and>
-  card B = dim S"
-    using basis_exists[of S] by auto
+  obtain B where B: "independent B \<and> B \<le> S \<and> S \<le> span B \<and> card B = dim S"
+    using basis_exists[of S] by metis
   then have "B \<noteq> {}"
     using B assms \<open>S \<noteq> {0}\<close> span_empty by auto
   have "insert 0 B \<le> span B"
@@ -1923,10 +1838,7 @@
 proof -
   define e where "e = a / (a + b)"
   have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)"
-    apply auto
-    using assms
-    apply simp
-    done
+    using assms  by (simp add: eq_vector_fraction_iff)
   also have "\<dots> = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)"
     using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"]
     by auto
@@ -2096,9 +2008,8 @@
 lemma rel_frontier_eq_empty:
     fixes S :: "'n::euclidean_space set"
     shows "rel_frontier S = {} \<longleftrightarrow> affine S"
-  apply (simp add: rel_frontier_def)
-  apply (simp add: rel_interior_eq_closure [symmetric])
-  using rel_interior_subset_closure by blast
+  unfolding rel_frontier_def
+  using rel_interior_subset_closure  by (auto simp add: rel_interior_eq_closure [symmetric])
 
 lemma rel_frontier_sing [simp]:
     fixes a :: "'n::euclidean_space"
@@ -2372,16 +2283,12 @@
   shows "z \<in> interior S \<longleftrightarrow> (\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S)"
 proof (cases "aff_dim S = int DIM('n)")
   case False
-  {
-    assume "z \<in> interior S"
+  { assume "z \<in> interior S"
     then have False
-      using False interior_rel_interior_gen[of S] by auto
-  }
+      using False interior_rel_interior_gen[of S] by auto }
   moreover
-  {
-    assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S"
-    {
-      fix x
+  { assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S"
+    { fix x
       obtain e1 where e1: "e1 > 0 \<and> z + e1 *\<^sub>R (x - z) \<in> S"
         using r by auto
       obtain e2 where e2: "e2 > 0 \<and> z + e2 *\<^sub>R (z - x) \<in> S"
@@ -3688,13 +3595,7 @@
     then have "x = (\<Sum>i\<in>I. c i *\<^sub>R s i) \<and> sum c I = 1"
       using k by (simp add: sum_prod)
     then have "x \<in> ?rhs"
-      using k
-      apply auto
-      apply (rule_tac x = c in exI)
-      apply (rule_tac x = s in exI)
-      using cs
-      apply auto
-      done
+      using k cs by auto
   }
   moreover
   {
@@ -3718,7 +3619,7 @@
       done
     then have "x \<in> ?lhs"
       using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x]
-      by (auto simp add: convex_convex_hull)
+      by auto
   }
   ultimately show ?thesis by blast
 qed
@@ -5112,10 +5013,7 @@
                       "S \<subseteq> S'"  "T \<subseteq> T'"  "S' \<inter> T' = {}"
 proof (cases "S = {} \<or> T = {}")
   case True with that show ?thesis
-    apply safe
-    using UT closedin_subset apply blast
-    using US closedin_subset apply blast
-    done
+    using UT US by (blast dest: closedin_subset)
 next
   case False
   define f where "f \<equiv> \<lambda>x. setdist {x} T - setdist {x} S"
@@ -5818,11 +5716,11 @@
 lemma dim_special_hyperplane:
   fixes k :: "'n::euclidean_space"
   shows "k \<in> Basis \<Longrightarrow> dim {x. k \<bullet> x = 0} = DIM('n) - 1"
-  apply (simp add: special_hyperplane_span)
-  apply (rule dim_unique [OF order_refl])
-    apply (auto simp: Diff_subset independent_substdbasis)
-  apply (metis member_remove remove_def span_clauses(1))
-  done
+apply (simp add: special_hyperplane_span)
+apply (rule Linear_Algebra.dim_unique [OF subset_refl])
+apply (auto simp: Diff_subset independent_substdbasis)
+apply (metis member_remove remove_def span_superset)
+done
 
 proposition dim_hyperplane:
   fixes a :: "'a::euclidean_space"
@@ -6852,15 +6750,12 @@
     using spanU by simp
   also have "... = span (insert a (S \<union> T))"
     apply (rule eq_span_insert_eq)
-    apply (simp add: a'_def span_neg span_sum span_clauses(1)
-        span_scale)
+    apply (simp add: a'_def span_neg span_sum span_superset span_mul)
     done
   also have "... = span (S \<union> insert a T)"
     by simp
   finally show ?case
-    apply (rule_tac x="insert a' U" in exI)
-    using orthU apply auto
-    done
+    by (rule_tac x="insert a' U" in exI) (use orthU in auto)
 qed
 
 
@@ -6870,7 +6765,7 @@
   obtains U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)"
 proof%unimportant -
   obtain B where "finite B" "span B = span T"
-    using basis_subspace_exists [of "span T"] subspace_span by auto
+    using basis_subspace_exists [of "span T"] subspace_span by metis
   with orthogonal_extension_aux [of B S]
   obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> B)"
     using assms pairwise_orthogonal_imp_finite by auto
@@ -6936,7 +6831,7 @@
              and "independent B" "card B = dim S" "span B = S"
     by (blast intro: orthogonal_basis_subspace [OF assms])
   have 1: "(\<lambda>x. x /\<^sub>R norm x) ` B \<subseteq> S"
-    using \<open>span B = S\<close> span_clauses(1) span_scale by fastforce
+    using \<open>span B = S\<close> span_superset span_mul by fastforce
   have 2: "pairwise orthogonal ((\<lambda>x. x /\<^sub>R norm x) ` B)"
     using orth by (force simp: pairwise_def orthogonal_clauses)
   have 3: "\<And>x. x \<in> (\<lambda>x. x /\<^sub>R norm x) ` B \<Longrightarrow> norm x = 1"
@@ -7027,7 +6922,7 @@
   have "dim {x} < DIM('a)"
     using assms by auto
   then show thesis
-    by (rule orthogonal_to_subspace_exists) (simp add: orthogonal_commute span_clauses(1) that)
+    by (rule orthogonal_to_subspace_exists) (simp add: orthogonal_commute span_superset that)
 qed
 
 proposition%important orthogonal_subspace_decomp_exists:
@@ -7297,7 +7192,7 @@
               have "e/2 / norm a \<noteq> 0"
                 using \<open>0 < e\<close> \<open>a \<noteq> 0\<close> by auto
               then show ?thesis
-                by (metis True \<open>a \<noteq> 0\<close> a orthogonal_scaleR orthogonal_self scale_eq_0_iff span_add_eq span_clauses(1))
+                by (metis True \<open>a \<noteq> 0\<close> a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_superset)
             qed
             ultimately show "?y \<in> S - U" by blast
           qed
@@ -8113,8 +8008,7 @@
     have "v = ?u + (v - ?u)"
       by simp
     moreover have "?u \<in> U"
-      by (metis (no_types, lifting) \<open>span B = U\<close> assms subspace_sum
-          span_clauses(1) span_scale)
+      by (metis (no_types, lifting) \<open>span B = U\<close> assms real_vector_class.subspace_sum span_superset span_mul)
     moreover have "(v - ?u) \<in> U\<^sup>\<bottom>"
     proof (clarsimp simp: orthogonal_comp_def orthogonal_def)
       fix y
--- a/src/HOL/Analysis/Vitali_Covering_Theorem.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Analysis/Vitali_Covering_Theorem.thy	Thu May 03 15:07:14 2018 +0200
@@ -1,3 +1,9 @@
+(*  Title:      HOL/Analysis/Vitali_Covering_Theorem.thy
+    Authors:    LC Paulson, based on material from HOL Light
+*)
+
+section\<open>Vitali Covering Theorem and an Application to Negligibility\<close>
+
 theory Vitali_Covering_Theorem
   imports Ball_Volume "HOL-Library.Permutations"
 
--- a/src/HOL/Code_Evaluation.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Code_Evaluation.thy	Thu May 03 15:07:14 2018 +0200
@@ -101,7 +101,7 @@
 ML_file "~~/src/HOL/Tools/value_command.ML"
 
 
-subsection \<open>\<open>term_of\<close> instances\<close>
+subsection \<open>Dedicated \<open>term_of\<close> instances\<close>
 
 instantiation "fun" :: (typerep, typerep) term_of
 begin
@@ -119,21 +119,6 @@
   "term_of :: typerep \<Rightarrow> _" "term_of :: term \<Rightarrow> _" "term_of :: String.literal \<Rightarrow> _"
   "term_of :: _ Predicate.pred \<Rightarrow> term" "term_of :: _ Predicate.seq \<Rightarrow> term"]]
 
-definition case_char :: "'a \<Rightarrow> (num \<Rightarrow> 'a) \<Rightarrow> char \<Rightarrow> 'a"
-  where "case_char f g c = (if c = 0 then f else g (num_of_nat (nat_of_char c)))"
-
-lemma term_of_char [unfolded typerep_fun_def typerep_char_def typerep_num_def, code]:
-  "term_of =
-    case_char (Const (STR ''Groups.zero_class.zero'') (TYPEREP(char)))
-    (\<lambda>k. App (Const (STR ''String.Char'') (TYPEREP(num \<Rightarrow> char))) (term_of k))"
-  by (rule ext) (rule term_of_anything [THEN meta_eq_to_obj_eq])
-
-lemma term_of_string [code]:
-   "term_of s = App (Const (STR ''STR'')
-     (Typerep.Typerep (STR ''fun'') [Typerep.Typerep (STR ''list'') [Typerep.Typerep (STR ''char'') []],
-       Typerep.Typerep (STR ''String.literal'') []])) (term_of (String.explode s))"
-  by (subst term_of_anything) rule 
-  
 code_printing
   constant "term_of :: integer \<Rightarrow> term" \<rightharpoonup> (Eval) "HOLogic.mk'_number/ HOLogic.code'_integerT"
 | constant "term_of :: String.literal \<Rightarrow> term" \<rightharpoonup> (Eval) "HOLogic.mk'_literal"
--- a/src/HOL/Code_Numeral.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Code_Numeral.thy	Thu May 03 15:07:14 2018 +0200
@@ -82,11 +82,15 @@
   unfolding dvd_def by transfer_prover
 
 lemma [transfer_rule]:
-  "rel_fun HOL.eq pcr_integer (of_nat :: nat \<Rightarrow> int) (of_nat :: nat \<Rightarrow> integer)"
+  "rel_fun (=) pcr_integer (of_bool :: bool \<Rightarrow> int) (of_bool :: bool \<Rightarrow> integer)"
+  by (unfold of_bool_def [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "rel_fun (=) pcr_integer (of_nat :: nat \<Rightarrow> int) (of_nat :: nat \<Rightarrow> integer)"
   by (rule transfer_rule_of_nat) transfer_prover+
 
 lemma [transfer_rule]:
-  "rel_fun HOL.eq pcr_integer (\<lambda>k :: int. k :: int) (of_int :: int \<Rightarrow> integer)"
+  "rel_fun (=) pcr_integer (\<lambda>k :: int. k :: int) (of_int :: int \<Rightarrow> integer)"
 proof -
   have "rel_fun HOL.eq pcr_integer (of_int :: int \<Rightarrow> int) (of_int :: int \<Rightarrow> integer)"
     by (rule transfer_rule_of_int) transfer_prover+
@@ -101,6 +105,10 @@
   "rel_fun HOL.eq (rel_fun HOL.eq pcr_integer) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> int) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> integer)"
   by (unfold Num.sub_def [abs_def]) transfer_prover
 
+lemma [transfer_rule]:
+  "rel_fun pcr_integer (rel_fun (=) pcr_integer) (power :: _ \<Rightarrow> _ \<Rightarrow> int) (power :: _ \<Rightarrow> _ \<Rightarrow> integer)"
+  by (unfold power_def [abs_def]) transfer_prover
+
 lemma int_of_integer_of_nat [simp]:
   "int_of_integer (of_nat n) = of_nat n"
   by transfer rule
@@ -265,6 +273,18 @@
 instance integer :: ring_parity
   by (standard; transfer) (simp_all add: of_nat_div division_segment_int_def)
 
+lemma [transfer_rule]:
+  "rel_fun (=) (rel_fun pcr_integer pcr_integer) (push_bit :: _ \<Rightarrow> _ \<Rightarrow> int) (push_bit :: _ \<Rightarrow> _ \<Rightarrow> integer)"
+  by (unfold push_bit_eq_mult [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "rel_fun (=) (rel_fun pcr_integer pcr_integer) (take_bit :: _ \<Rightarrow> _ \<Rightarrow> int) (take_bit :: _ \<Rightarrow> _ \<Rightarrow> integer)"
+  by (unfold take_bit_eq_mod [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "rel_fun (=) (rel_fun pcr_integer pcr_integer) (drop_bit :: _ \<Rightarrow> _ \<Rightarrow> int) (drop_bit :: _ \<Rightarrow> _ \<Rightarrow> integer)"
+  by (unfold drop_bit_eq_div [abs_def]) transfer_prover
+
 instantiation integer :: unique_euclidean_semiring_numeral
 begin
 
@@ -313,6 +333,7 @@
   "integer_of_nat (numeral n) = numeral n"
 by transfer simp
 
+
 subsection \<open>Code theorems for target language integers\<close>
 
 text \<open>Constructors\<close>
@@ -495,6 +516,20 @@
   "k mod l = snd (divmod_integer k l)"
   by simp
 
+definition bit_cut_integer :: "integer \<Rightarrow> integer \<times> bool"
+  where "bit_cut_integer k = (k div 2, odd k)"
+
+lemma bit_cut_integer_code [code]:
+  "bit_cut_integer k = (if k = 0 then (0, False)
+     else let (r, s) = Code_Numeral.divmod_abs k 2
+       in (if k > 0 then r else - r - s, s = 1))"
+proof -
+  have "bit_cut_integer k = (let (r, s) = divmod_integer k 2 in (r, s = 1))"
+    by (simp add: divmod_integer_def bit_cut_integer_def odd_iff_mod_2_eq_one)
+  then show ?thesis
+    by (simp add: divmod_integer_code) (auto simp add: split_def)
+qed
+
 lemma equal_integer_code [code]:
   "HOL.equal 0 (0::integer) \<longleftrightarrow> True"
   "HOL.equal 0 (Pos l) \<longleftrightarrow> False"
@@ -777,6 +812,10 @@
   unfolding dvd_def by transfer_prover
 
 lemma [transfer_rule]:
+  "rel_fun (=) pcr_natural (of_bool :: bool \<Rightarrow> nat) (of_bool :: bool \<Rightarrow> natural)"
+  by (unfold of_bool_def [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
   "rel_fun HOL.eq pcr_natural (\<lambda>n::nat. n) (of_nat :: nat \<Rightarrow> natural)"
 proof -
   have "rel_fun HOL.eq pcr_natural (of_nat :: nat \<Rightarrow> nat) (of_nat :: nat \<Rightarrow> natural)"
@@ -792,6 +831,10 @@
   then show ?thesis by simp
 qed
 
+lemma [transfer_rule]:
+  "rel_fun pcr_natural (rel_fun (=) pcr_natural) (power :: _ \<Rightarrow> _ \<Rightarrow> nat) (power :: _ \<Rightarrow> _ \<Rightarrow> natural)"
+  by (unfold power_def [abs_def]) transfer_prover
+
 lemma nat_of_natural_of_nat [simp]:
   "nat_of_natural (of_nat n) = n"
   by transfer rule
@@ -895,6 +938,18 @@
 instance natural :: semiring_parity
   by (standard; transfer) simp_all
 
+lemma [transfer_rule]:
+  "rel_fun (=) (rel_fun pcr_natural pcr_natural) (push_bit :: _ \<Rightarrow> _ \<Rightarrow> nat) (push_bit :: _ \<Rightarrow> _ \<Rightarrow> natural)"
+  by (unfold push_bit_eq_mult [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "rel_fun (=) (rel_fun pcr_natural pcr_natural) (take_bit :: _ \<Rightarrow> _ \<Rightarrow> nat) (take_bit :: _ \<Rightarrow> _ \<Rightarrow> natural)"
+  by (unfold take_bit_eq_mod [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "rel_fun (=) (rel_fun pcr_natural pcr_natural) (drop_bit :: _ \<Rightarrow> _ \<Rightarrow> nat) (drop_bit :: _ \<Rightarrow> _ \<Rightarrow> natural)"
+  by (unfold drop_bit_eq_div [abs_def]) transfer_prover
+
 lift_definition natural_of_integer :: "integer \<Rightarrow> natural"
   is "nat :: int \<Rightarrow> nat"
   .
--- a/src/HOL/Codegenerator_Test/Generate_Pretty_Char.thy	Wed May 02 13:49:38 2018 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,21 +0,0 @@
-
-(* Author: Florian Haftmann, TU Muenchen *)
-
-section \<open>Pervasive test of code generator\<close>
-
-theory Generate_Pretty_Char
-imports
-  Candidates
-  "HOL-Library.AList_Mapping"
-  "HOL-Library.Finite_Lattice"
-  "HOL-Library.Code_Char"
-begin
-
-text \<open>
-  If any of the checks fails, inspect the code generated
-  by a corresponding \<open>export_code\<close> command.
-\<close>
-
-export_code _ checking SML OCaml? Haskell? Scala
-
-end
--- a/src/HOL/Data_Structures/AA_Map.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/AA_Map.thy	Thu May 03 15:07:14 2018 +0200
@@ -23,7 +23,7 @@
      LT \<Rightarrow> adjust (Node lv (delete x l) (a,b) r) |
      GT \<Rightarrow> adjust (Node lv l (a,b) (delete x r)) |
      EQ \<Rightarrow> (if l = Leaf then r
-            else let (l',ab') = del_max l in adjust (Node lv l' ab' r)))"
+            else let (l',ab') = split_max l in adjust (Node lv l' ab' r)))"
 
 
 subsection "Invariance"
@@ -187,7 +187,7 @@
         by(auto simp: post_del_def invar.simps(2))
     next
       assume "l \<noteq> Leaf" thus ?thesis using equal Node.prems
-        by simp (metis inv_l post_del_adjL post_del_max pre_adj_if_postL)
+        by simp (metis inv_l post_del_adjL post_split_max pre_adj_if_postL)
     qed
   qed
 qed (simp add: post_del_def)
@@ -204,7 +204,7 @@
   inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR 
-              post_del_max post_delete del_maxD split: prod.splits)
+              post_split_max post_delete split_maxD split: prod.splits)
 
 interpretation I: Map_by_Ordered
 where empty = Leaf and lookup = lookup and update = update and delete = delete
--- a/src/HOL/Data_Structures/AA_Set.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/AA_Set.thy	Thu May 03 15:07:14 2018 +0200
@@ -72,14 +72,14 @@
 text\<open>In the paper, the last case of @{const adjust} is expressed with the help of an
 incorrect auxiliary function \texttt{nlvl}.
 
-Function @{text del_max} below is called \texttt{dellrg} in the paper.
+Function @{text split_max} below is called \texttt{dellrg} in the paper.
 The latter is incorrect for two reasons: \texttt{dellrg} is meant to delete the largest
 element but recurses on the left instead of the right subtree; the invariant
 is not restored.\<close>
 
-fun del_max :: "'a aa_tree \<Rightarrow> 'a aa_tree * 'a" where
-"del_max (Node lv l a Leaf) = (l,a)" |
-"del_max (Node lv l a r) = (let (r',b) = del_max r in (adjust(Node lv l a r'), b))"
+fun split_max :: "'a aa_tree \<Rightarrow> 'a aa_tree * 'a" where
+"split_max (Node lv l a Leaf) = (l,a)" |
+"split_max (Node lv l a r) = (let (r',b) = split_max r in (adjust(Node lv l a r'), b))"
 
 fun delete :: "'a::linorder \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where
 "delete _ Leaf = Leaf" |
@@ -88,7 +88,7 @@
      LT \<Rightarrow> adjust (Node lv (delete x l) a r) |
      GT \<Rightarrow> adjust (Node lv l a (delete x r)) |
      EQ \<Rightarrow> (if l = Leaf then r
-            else let (l',b) = del_max l in adjust (Node lv l' b r)))"
+            else let (l',b) = split_max l in adjust (Node lv l' b r)))"
 
 fun pre_adjust where
 "pre_adjust (Node lv l a r) = (invar l \<and> invar r \<and>
@@ -397,13 +397,13 @@
 
 declare prod.splits[split]
 
-theorem post_del_max:
- "\<lbrakk> invar t; (t', x) = del_max t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> post_del t t'"
-proof (induction t arbitrary: t' rule: del_max.induct)
+theorem post_split_max:
+ "\<lbrakk> invar t; (t', x) = split_max t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> post_del t t'"
+proof (induction t arbitrary: t' rule: split_max.induct)
   case (2 lv l a lvr rl ra rr)
   let ?r =  "\<langle>lvr, rl, ra, rr\<rangle>"
   let ?t = "\<langle>lv, l, a, ?r\<rangle>"
-  from "2.prems"(2) obtain r' where r': "(r', x) = del_max ?r"
+  from "2.prems"(2) obtain r' where r': "(r', x) = split_max ?r"
     and [simp]: "t' = adjust \<langle>lv, l, a, r'\<rangle>" by auto
   from  "2.IH"[OF _ r'] \<open>invar ?t\<close> have post: "post_del ?r r'" by simp
   note preR = pre_adj_if_postR[OF \<open>invar ?t\<close> post]
@@ -440,7 +440,7 @@
         by(auto simp: post_del_def invar.simps(2))
     next
       assume "l \<noteq> Leaf" thus ?thesis using equal
-        by simp (metis Node.prems inv_l post_del_adjL post_del_max pre_adj_if_postL)
+        by simp (metis Node.prems inv_l post_del_adjL post_split_max pre_adj_if_postL)
     qed
   qed
 qed (simp add: post_del_def)
@@ -471,16 +471,16 @@
   (auto simp: adjust_def inorder_skew inorder_split invar.simps(2) pre_adjust.simps
      split: tree.splits)
 
-lemma del_maxD:
-  "\<lbrakk> del_max t = (t',x); t \<noteq> Leaf; invar t \<rbrakk> \<Longrightarrow> inorder t' @ [x] = inorder t"
-by(induction t arbitrary: t' rule: del_max.induct)
-  (auto simp: sorted_lems inorder_adjust pre_adj_if_postR post_del_max split: prod.splits)
+lemma split_maxD:
+  "\<lbrakk> split_max t = (t',x); t \<noteq> Leaf; invar t \<rbrakk> \<Longrightarrow> inorder t' @ [x] = inorder t"
+by(induction t arbitrary: t' rule: split_max.induct)
+  (auto simp: sorted_lems inorder_adjust pre_adj_if_postR post_split_max split: prod.splits)
 
 lemma inorder_delete:
   "invar t \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR 
-              post_del_max post_delete del_maxD split: prod.splits)
+              post_split_max post_delete split_maxD split: prod.splits)
 
 interpretation I: Set_by_Ordered
 where empty = Leaf and isin = isin and insert = insert and delete = delete
--- a/src/HOL/Data_Structures/AVL_Map.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/AVL_Map.thy	Thu May 03 15:07:14 2018 +0200
@@ -34,7 +34,7 @@
   "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_balL inorder_balR
-     inorder_del_root inorder_del_maxD split: prod.splits)
+     inorder_del_root inorder_split_maxD split: prod.splits)
 
 interpretation Map_by_Ordered
 where empty = Leaf and lookup = lookup and update = update and delete = delete
--- a/src/HOL/Data_Structures/AVL_Set.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/AVL_Set.thy	Thu May 03 15:07:14 2018 +0200
@@ -58,16 +58,16 @@
    LT \<Rightarrow> balL (insert x l) a r |
    GT \<Rightarrow> balR l a (insert x r))"
 
-fun del_max :: "'a avl_tree \<Rightarrow> 'a avl_tree * 'a" where
-"del_max (Node _ l a r) =
-  (if r = Leaf then (l,a) else let (r',a') = del_max r in (balL l a r', a'))"
+fun split_max :: "'a avl_tree \<Rightarrow> 'a avl_tree * 'a" where
+"split_max (Node _ l a r) =
+  (if r = Leaf then (l,a) else let (r',a') = split_max r in (balL l a r', a'))"
 
-lemmas del_max_induct = del_max.induct[case_names Node Leaf]
+lemmas split_max_induct = split_max.induct[case_names Node Leaf]
 
 fun del_root :: "'a avl_tree \<Rightarrow> 'a avl_tree" where
 "del_root (Node h Leaf a r) = r" |
 "del_root (Node h l a Leaf) = l" |
-"del_root (Node h l a r) = (let (l', a') = del_max l in balR l' a' r)"
+"del_root (Node h l a r) = (let (l', a') = split_max l in balR l' a' r)"
 
 lemmas del_root_cases = del_root.cases[case_names Leaf_t Node_Leaf Node_Node]
 
@@ -103,22 +103,22 @@
 
 subsubsection "Proofs for delete"
 
-lemma inorder_del_maxD:
-  "\<lbrakk> del_max t = (t',a); t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
+lemma inorder_split_maxD:
+  "\<lbrakk> split_max t = (t',a); t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
    inorder t' @ [a] = inorder t"
-by(induction t arbitrary: t' rule: del_max.induct)
+by(induction t arbitrary: t' rule: split_max.induct)
   (auto simp: inorder_balL split: if_splits prod.splits tree.split)
 
 lemma inorder_del_root:
   "inorder (del_root (Node h l a r)) = inorder l @ inorder r"
 by(cases "Node h l a r" rule: del_root.cases)
-  (auto simp: inorder_balL inorder_balR inorder_del_maxD split: if_splits prod.splits)
+  (auto simp: inorder_balL inorder_balR inorder_split_maxD split: if_splits prod.splits)
 
 theorem inorder_delete:
   "sorted(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_balL inorder_balR
-    inorder_del_root inorder_del_maxD split: prod.splits)
+    inorder_del_root inorder_split_maxD split: prod.splits)
 
 
 subsubsection "Overall functional correctness"
@@ -301,12 +301,12 @@
 
 subsubsection \<open>Deletion maintains AVL balance\<close>
 
-lemma avl_del_max:
+lemma avl_split_max:
   assumes "avl x" and "x \<noteq> Leaf"
-  shows "avl (fst (del_max x))" "height x = height(fst (del_max x)) \<or>
-         height x = height(fst (del_max x)) + 1"
+  shows "avl (fst (split_max x))" "height x = height(fst (split_max x)) \<or>
+         height x = height(fst (split_max x)) + 1"
 using assms
-proof (induct x rule: del_max_induct)
+proof (induct x rule: split_max_induct)
   case (Node h l a r)
   case 1
   thus ?case using Node
@@ -316,7 +316,7 @@
 next
   case (Node h l a r)
   case 2
-  let ?r' = "fst (del_max r)"
+  let ?r' = "fst (split_max r)"
   from \<open>avl x\<close> Node 2 have "avl l" and "avl r" by simp_all
   thus ?case using Node 2 height_balL[of l ?r' a] height_balL2[of l ?r' a]
     apply (auto split:prod.splits simp del:avl.simps) by arith+
@@ -330,14 +330,14 @@
   case (Node_Node h lh ll ln lr n rh rl rn rr)
   let ?l = "Node lh ll ln lr"
   let ?r = "Node rh rl rn rr"
-  let ?l' = "fst (del_max ?l)"
+  let ?l' = "fst (split_max ?l)"
   from \<open>avl t\<close> and Node_Node have "avl ?r" by simp
   from \<open>avl t\<close> and Node_Node have "avl ?l" by simp
   hence "avl(?l')" "height ?l = height(?l') \<or>
-         height ?l = height(?l') + 1" by (rule avl_del_max,simp)+
+         height ?l = height(?l') + 1" by (rule avl_split_max,simp)+
   with \<open>avl t\<close> Node_Node have "height ?l' = height ?r \<or> height ?l' = height ?r + 1
             \<or> height ?r = height ?l' + 1 \<or> height ?r = height ?l' + 2" by fastforce
-  with \<open>avl ?l'\<close> \<open>avl ?r\<close> have "avl(balR ?l' (snd(del_max ?l)) ?r)"
+  with \<open>avl ?l'\<close> \<open>avl ?r\<close> have "avl(balR ?l' (snd(split_max ?l)) ?r)"
     by (rule avl_balR)
   with Node_Node show ?thesis by (auto split:prod.splits)
 qed simp_all
@@ -350,12 +350,12 @@
   case (Node_Node h lh ll ln lr n rh rl rn rr)
   let ?l = "Node lh ll ln lr"
   let ?r = "Node rh rl rn rr"
-  let ?l' = "fst (del_max ?l)"
-  let ?t' = "balR ?l' (snd(del_max ?l)) ?r"
+  let ?l' = "fst (split_max ?l)"
+  let ?t' = "balR ?l' (snd(split_max ?l)) ?r"
   from \<open>avl t\<close> and Node_Node have "avl ?r" by simp
   from \<open>avl t\<close> and Node_Node have "avl ?l" by simp
-  hence "avl(?l')"  by (rule avl_del_max,simp)
-  have l'_height: "height ?l = height ?l' \<or> height ?l = height ?l' + 1" using \<open>avl ?l\<close> by (intro avl_del_max) auto
+  hence "avl(?l')"  by (rule avl_split_max,simp)
+  have l'_height: "height ?l = height ?l' \<or> height ?l = height ?l' + 1" using \<open>avl ?l\<close> by (intro avl_split_max) auto
   have t_height: "height t = 1 + max (height ?l) (height ?r)" using \<open>avl t\<close> Node_Node by simp
   have "height t = height ?t' \<or> height t = height ?t' + 1" using  \<open>avl t\<close> Node_Node
   proof(cases "height ?r = height ?l' + 2")
@@ -364,7 +364,7 @@
   next
     case True
     show ?thesis
-    proof(cases rule: disjE[OF height_balR[OF True \<open>avl ?l'\<close> \<open>avl ?r\<close>, of "snd (del_max ?l)"]])
+    proof(cases rule: disjE[OF height_balR[OF True \<open>avl ?l'\<close> \<open>avl ?r\<close>, of "snd (split_max ?l)"]])
       case 1
       thus ?thesis using l'_height t_height True by arith
     next
--- a/src/HOL/Data_Structures/Brother12_Map.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/Brother12_Map.thy	Thu May 03 15:07:14 2018 +0200
@@ -44,7 +44,7 @@
   (case cmp x a of
      LT \<Rightarrow> n2 (del x l) (a,b) r |
      GT \<Rightarrow> n2 l (a,b) (del x r) |
-     EQ \<Rightarrow> (case del_min r of
+     EQ \<Rightarrow> (case split_min r of
               None \<Rightarrow> N1 l |
               Some (ab, r') \<Rightarrow> n2 l ab r'))"
 
@@ -84,7 +84,7 @@
 lemma inorder_del:
   "t \<in> T h \<Longrightarrow> sorted1(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)"
 by(induction h arbitrary: t) (auto simp: del_list_simps inorder_n2
-     inorder_del_min[OF UnI1] inorder_del_min[OF UnI2] split: option.splits)
+     inorder_split_min[OF UnI1] inorder_split_min[OF UnI2] split: option.splits)
 
 lemma inorder_delete:
   "t \<in> T h \<Longrightarrow> sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
@@ -151,16 +151,16 @@
     qed
     moreover
     have ?case if [simp]: "x=a"
-    proof (cases "del_min r")
+    proof (cases "split_min r")
       case None
       show ?thesis
       proof cases
         assume "r \<in> B h"
-        with del_minNoneN0[OF this None] lr show ?thesis by(simp)
+        with split_minNoneN0[OF this None] lr show ?thesis by(simp)
       next
         assume "r \<notin> B h"
         hence "r \<in> U h" using lr by auto
-        with del_minNoneN1[OF this None] lr(3) show ?thesis by (simp)
+        with split_minNoneN1[OF this None] lr(3) show ?thesis by (simp)
       qed
     next
       case [simp]: (Some br')
@@ -168,12 +168,12 @@
       show ?thesis
       proof cases
         assume "r \<in> B h"
-        from del_min_type(1)[OF this] n2_type3[OF lr(1)]
+        from split_min_type(1)[OF this] n2_type3[OF lr(1)]
         show ?thesis by simp
       next
         assume "r \<notin> B h"
         hence "l \<in> B h" and "r \<in> U h" using lr by auto
-        from del_min_type(2)[OF this(2)] n2_type2[OF this(1)]
+        from split_min_type(2)[OF this(2)] n2_type2[OF this(1)]
         show ?thesis by simp
       qed
     qed
--- a/src/HOL/Data_Structures/Brother12_Set.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/Brother12_Set.thy	Thu May 03 15:07:14 2018 +0200
@@ -92,14 +92,14 @@
   N2 (N1 (N2 t1 a1 t2)) a2 (N2 (N2 t3 a3 t4) a5 (N1 t5))" |
 "n2 t1 a1 t2 = N2 t1 a1 t2"
 
-fun del_min :: "'a bro \<Rightarrow> ('a \<times> 'a bro) option" where
-"del_min N0 = None" |
-"del_min (N1 t) =
-  (case del_min t of
+fun split_min :: "'a bro \<Rightarrow> ('a \<times> 'a bro) option" where
+"split_min N0 = None" |
+"split_min (N1 t) =
+  (case split_min t of
      None \<Rightarrow> None |
      Some (a, t') \<Rightarrow> Some (a, N1 t'))" |
-"del_min (N2 t1 a t2) =
-  (case del_min t1 of
+"split_min (N2 t1 a t2) =
+  (case split_min t1 of
      None \<Rightarrow> Some (a, N1 t2) |
      Some (b, t1') \<Rightarrow> Some (b, n2 t1' a t2))"
 
@@ -110,7 +110,7 @@
   (case cmp x a of
      LT \<Rightarrow> n2 (del x l) a r |
      GT \<Rightarrow> n2 l a (del x r) |
-     EQ \<Rightarrow> (case del_min r of
+     EQ \<Rightarrow> (case split_min r of
               None \<Rightarrow> N1 l |
               Some (b, r') \<Rightarrow> n2 l b r'))"
 
@@ -189,15 +189,15 @@
 lemma inorder_n2: "inorder(n2 l a r) = inorder l @ a # inorder r"
 by(cases "(l,a,r)" rule: n2.cases) (auto)
 
-lemma inorder_del_min:
-  "t \<in> T h \<Longrightarrow> (del_min t = None \<longleftrightarrow> inorder t = []) \<and>
-  (del_min t = Some(a,t') \<longrightarrow> inorder t = a # inorder t')"
+lemma inorder_split_min:
+  "t \<in> T h \<Longrightarrow> (split_min t = None \<longleftrightarrow> inorder t = []) \<and>
+  (split_min t = Some(a,t') \<longrightarrow> inorder t = a # inorder t')"
 by(induction h arbitrary: t a t') (auto simp: inorder_n2 split: option.splits)
 
 lemma inorder_del:
   "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)"
 by(induction h arbitrary: t) (auto simp: del_list_simps inorder_n2
-     inorder_del_min[OF UnI1] inorder_del_min[OF UnI2] split: option.splits)
+     inorder_split_min[OF UnI1] inorder_split_min[OF UnI2] split: option.splits)
 
 lemma inorder_delete:
   "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
@@ -331,15 +331,15 @@
 apply(erule exE bexE conjE imageE | simp | erule disjE)+
 done
 
-lemma del_minNoneN0: "\<lbrakk>t \<in> B h; del_min t = None\<rbrakk> \<Longrightarrow>  t = N0"
+lemma split_minNoneN0: "\<lbrakk>t \<in> B h; split_min t = None\<rbrakk> \<Longrightarrow>  t = N0"
 by (cases t) (auto split: option.splits)
 
-lemma del_minNoneN1 : "\<lbrakk>t \<in> U h; del_min t = None\<rbrakk> \<Longrightarrow> t = N1 N0"
-by (cases h) (auto simp: del_minNoneN0  split: option.splits)
+lemma split_minNoneN1 : "\<lbrakk>t \<in> U h; split_min t = None\<rbrakk> \<Longrightarrow> t = N1 N0"
+by (cases h) (auto simp: split_minNoneN0  split: option.splits)
 
-lemma del_min_type:
-  "t \<in> B h \<Longrightarrow> del_min t = Some (a, t') \<Longrightarrow> t' \<in> T h"
-  "t \<in> U h \<Longrightarrow> del_min t = Some (a, t') \<Longrightarrow> t' \<in> Um h"
+lemma split_min_type:
+  "t \<in> B h \<Longrightarrow> split_min t = Some (a, t') \<Longrightarrow> t' \<in> T h"
+  "t \<in> U h \<Longrightarrow> split_min t = Some (a, t') \<Longrightarrow> t' \<in> Um h"
 proof (induction h arbitrary: t a t')
   case (Suc h)
   { case 1
@@ -347,12 +347,12 @@
       t12: "t1 \<in> T h" "t2 \<in> T h" "t1 \<in> B h \<or> t2 \<in> B h"
       by auto
     show ?case
-    proof (cases "del_min t1")
+    proof (cases "split_min t1")
       case None
       show ?thesis
       proof cases
         assume "t1 \<in> B h"
-        with del_minNoneN0[OF this None] 1 show ?thesis by(auto)
+        with split_minNoneN0[OF this None] 1 show ?thesis by(auto)
       next
         assume "t1 \<notin> B h"
         thus ?thesis using 1 None by (auto)
@@ -376,9 +376,9 @@
   { case 2
     then obtain t1 where [simp]: "t = N1 t1" and t1: "t1 \<in> B h" by auto
     show ?case
-    proof (cases "del_min t1")
+    proof (cases "split_min t1")
       case None
-      with del_minNoneN0[OF t1 None] 2 show ?thesis by(auto)
+      with split_minNoneN0[OF t1 None] 2 show ?thesis by(auto)
     next
       case [simp]: (Some bt')
       obtain b t1' where [simp]: "bt' = (b,t1')" by fastforce
@@ -421,16 +421,16 @@
     qed
     moreover
     have ?case if [simp]: "x=a"
-    proof (cases "del_min r")
+    proof (cases "split_min r")
       case None
       show ?thesis
       proof cases
         assume "r \<in> B h"
-        with del_minNoneN0[OF this None] lr show ?thesis by(simp)
+        with split_minNoneN0[OF this None] lr show ?thesis by(simp)
       next
         assume "r \<notin> B h"
         hence "r \<in> U h" using lr by auto
-        with del_minNoneN1[OF this None] lr(3) show ?thesis by (simp)
+        with split_minNoneN1[OF this None] lr(3) show ?thesis by (simp)
       qed
     next
       case [simp]: (Some br')
@@ -438,12 +438,12 @@
       show ?thesis
       proof cases
         assume "r \<in> B h"
-        from del_min_type(1)[OF this] n2_type3[OF lr(1)]
+        from split_min_type(1)[OF this] n2_type3[OF lr(1)]
         show ?thesis by simp
       next
         assume "r \<notin> B h"
         hence "l \<in> B h" and "r \<in> U h" using lr by auto
-        from del_min_type(2)[OF this(2)] n2_type2[OF this(1)]
+        from split_min_type(2)[OF this(2)] n2_type2[OF this(1)]
         show ?thesis by simp
       qed
     qed
--- a/src/HOL/Data_Structures/Tree234_Map.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/Tree234_Map.thy	Thu May 03 15:07:14 2018 +0200
@@ -88,23 +88,23 @@
 "del x (Node2 l ab1 r) = (case cmp x (fst ab1) of
   LT \<Rightarrow> node21 (del x l) ab1 r |
   GT \<Rightarrow> node22 l ab1 (del x r) |
-  EQ \<Rightarrow> let (ab1',t) = del_min r in node22 l ab1' t)" |
+  EQ \<Rightarrow> let (ab1',t) = split_min r in node22 l ab1' t)" |
 "del x (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of
   LT \<Rightarrow> node31 (del x l) ab1 m ab2 r |
-  EQ \<Rightarrow> let (ab1',m') = del_min m in node32 l ab1' m' ab2 r |
+  EQ \<Rightarrow> let (ab1',m') = split_min m in node32 l ab1' m' ab2 r |
   GT \<Rightarrow> (case cmp x (fst ab2) of
            LT \<Rightarrow> node32 l ab1 (del x m) ab2 r |
-           EQ \<Rightarrow> let (ab2',r') = del_min r in node33 l ab1 m ab2' r' |
+           EQ \<Rightarrow> let (ab2',r') = split_min r in node33 l ab1 m ab2' r' |
            GT \<Rightarrow> node33 l ab1 m ab2 (del x r)))" |
 "del x (Node4 t1 ab1 t2 ab2 t3 ab3 t4) = (case cmp x (fst ab2) of
   LT \<Rightarrow> (case cmp x (fst ab1) of
            LT \<Rightarrow> node41 (del x t1) ab1 t2 ab2 t3 ab3 t4 |
-           EQ \<Rightarrow> let (ab',t2') = del_min t2 in node42 t1 ab' t2' ab2 t3 ab3 t4 |
+           EQ \<Rightarrow> let (ab',t2') = split_min t2 in node42 t1 ab' t2' ab2 t3 ab3 t4 |
            GT \<Rightarrow> node42 t1 ab1 (del x t2) ab2 t3 ab3 t4) |
-  EQ \<Rightarrow> let (ab',t3') = del_min t3 in node43 t1 ab1 t2 ab' t3' ab3 t4 |
+  EQ \<Rightarrow> let (ab',t3') = split_min t3 in node43 t1 ab1 t2 ab' t3' ab3 t4 |
   GT \<Rightarrow> (case cmp x (fst ab3) of
           LT \<Rightarrow> node43 t1 ab1 t2 ab2 (del x t3) ab3 t4 |
-          EQ \<Rightarrow> let (ab',t4') = del_min t4 in node44 t1 ab1 t2 ab2 t3 ab' t4' |
+          EQ \<Rightarrow> let (ab',t4') = split_min t4 in node44 t1 ab1 t2 ab2 t3 ab' t4' |
           GT \<Rightarrow> node44 t1 ab1 t2 ab2 t3 ab3 (del x t4)))"
 
 definition delete :: "'a::linorder \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) tree234" where
@@ -130,7 +130,7 @@
 lemma inorder_del: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)"
 by(induction t rule: del.induct)
-  (auto simp: del_list_simps inorder_nodes del_minD split!: if_splits prod.splits)
+  (auto simp: del_list_simps inorder_nodes split_minD split!: if_splits prod.splits)
 (* 30 secs (2016) *)
 
 lemma inorder_delete: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
@@ -148,11 +148,11 @@
 
 lemma height_del: "bal t \<Longrightarrow> height(del x t) = height t"
 by(induction x t rule: del.induct)
-  (auto simp add: heights height_del_min split!: if_split prod.split)
+  (auto simp add: heights height_split_min split!: if_split prod.split)
 
 lemma bal_tree\<^sub>d_del: "bal t \<Longrightarrow> bal(tree\<^sub>d(del x t))"
 by(induction x t rule: del.induct)
-  (auto simp: bals bal_del_min height_del height_del_min split!: if_split prod.split)
+  (auto simp: bals bal_split_min height_del height_split_min split!: if_split prod.split)
 
 corollary bal_delete: "bal t \<Longrightarrow> bal(delete x t)"
 by(simp add: delete_def bal_tree\<^sub>d_del)
--- a/src/HOL/Data_Structures/Tree234_Set.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/Tree234_Set.thy	Thu May 03 15:07:14 2018 +0200
@@ -154,13 +154,13 @@
 "node44 t1 a t2 b (Node3 t3 c t4 d t5) e (Up\<^sub>d t6) = T\<^sub>d(Node4 t1 a t2 b (Node2 t3 c t4) d (Node2 t5 e t6))" |
 "node44 t1 a t2 b (Node4 t3 c t4 d t5 e t6) f (Up\<^sub>d t7) = T\<^sub>d(Node4 t1 a t2 b (Node2 t3 c t4) d (Node3 t5 e t6 f t7))"
 
-fun del_min :: "'a tree234 \<Rightarrow> 'a * 'a up\<^sub>d" where
-"del_min (Node2 Leaf a Leaf) = (a, Up\<^sub>d Leaf)" |
-"del_min (Node3 Leaf a Leaf b Leaf) = (a, T\<^sub>d(Node2 Leaf b Leaf))" |
-"del_min (Node4 Leaf a Leaf b Leaf c Leaf) = (a, T\<^sub>d(Node3 Leaf b Leaf c Leaf))" |
-"del_min (Node2 l a r) = (let (x,l') = del_min l in (x, node21 l' a r))" |
-"del_min (Node3 l a m b r) = (let (x,l') = del_min l in (x, node31 l' a m b r))" |
-"del_min (Node4 l a m b n c r) = (let (x,l') = del_min l in (x, node41 l' a m b n c r))"
+fun split_min :: "'a tree234 \<Rightarrow> 'a * 'a up\<^sub>d" where
+"split_min (Node2 Leaf a Leaf) = (a, Up\<^sub>d Leaf)" |
+"split_min (Node3 Leaf a Leaf b Leaf) = (a, T\<^sub>d(Node2 Leaf b Leaf))" |
+"split_min (Node4 Leaf a Leaf b Leaf c Leaf) = (a, T\<^sub>d(Node3 Leaf b Leaf c Leaf))" |
+"split_min (Node2 l a r) = (let (x,l') = split_min l in (x, node21 l' a r))" |
+"split_min (Node3 l a m b r) = (let (x,l') = split_min l in (x, node31 l' a m b r))" |
+"split_min (Node4 l a m b n c r) = (let (x,l') = split_min l in (x, node41 l' a m b n c r))"
 
 fun del :: "'a::linorder \<Rightarrow> 'a tree234 \<Rightarrow> 'a up\<^sub>d" where
 "del k Leaf = T\<^sub>d Leaf" |
@@ -175,23 +175,23 @@
 "del k (Node2 l a r) = (case cmp k a of
   LT \<Rightarrow> node21 (del k l) a r |
   GT \<Rightarrow> node22 l a (del k r) |
-  EQ \<Rightarrow> let (a',t) = del_min r in node22 l a' t)" |
+  EQ \<Rightarrow> let (a',t) = split_min r in node22 l a' t)" |
 "del k (Node3 l a m b r) = (case cmp k a of
   LT \<Rightarrow> node31 (del k l) a m b r |
-  EQ \<Rightarrow> let (a',m') = del_min m in node32 l a' m' b r |
+  EQ \<Rightarrow> let (a',m') = split_min m in node32 l a' m' b r |
   GT \<Rightarrow> (case cmp k b of
            LT \<Rightarrow> node32 l a (del k m) b r |
-           EQ \<Rightarrow> let (b',r') = del_min r in node33 l a m b' r' |
+           EQ \<Rightarrow> let (b',r') = split_min r in node33 l a m b' r' |
            GT \<Rightarrow> node33 l a m b (del k r)))" |
 "del k (Node4 l a m b n c r) = (case cmp k b of
   LT \<Rightarrow> (case cmp k a of
           LT \<Rightarrow> node41 (del k l) a m b n c r |
-          EQ \<Rightarrow> let (a',m') = del_min m in node42 l a' m' b n c r |
+          EQ \<Rightarrow> let (a',m') = split_min m in node42 l a' m' b n c r |
           GT \<Rightarrow> node42 l a (del k m) b n c r) |
-  EQ \<Rightarrow> let (b',n') = del_min n in node43 l a m b' n' c r |
+  EQ \<Rightarrow> let (b',n') = split_min n in node43 l a m b' n' c r |
   GT \<Rightarrow> (case cmp k c of
            LT \<Rightarrow> node43 l a m b (del k n) c r |
-           EQ \<Rightarrow> let (c',r') = del_min r in node44 l a m b n c' r' |
+           EQ \<Rightarrow> let (c',r') = split_min r in node44 l a m b n c' r' |
            GT \<Rightarrow> node44 l a m b n c (del k r)))"
 
 definition delete :: "'a::linorder \<Rightarrow> 'a tree234 \<Rightarrow> 'a tree234" where
@@ -259,16 +259,16 @@
   inorder_node31 inorder_node32 inorder_node33
   inorder_node41 inorder_node42 inorder_node43 inorder_node44
 
-lemma del_minD:
-  "del_min t = (x,t') \<Longrightarrow> bal t \<Longrightarrow> height t > 0 \<Longrightarrow>
+lemma split_minD:
+  "split_min t = (x,t') \<Longrightarrow> bal t \<Longrightarrow> height t > 0 \<Longrightarrow>
   x # inorder(tree\<^sub>d t') = inorder t"
-by(induction t arbitrary: t' rule: del_min.induct)
+by(induction t arbitrary: t' rule: split_min.induct)
   (auto simp: inorder_nodes split: prod.splits)
 
 lemma inorder_del: "\<lbrakk> bal t ; sorted(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)"
 by(induction t rule: del.induct)
-  (auto simp: inorder_nodes del_list_simps del_minD split!: if_split prod.splits)
+  (auto simp: inorder_nodes del_list_simps split_minD split!: if_split prod.splits)
   (* 30 secs (2016) *)
 
 lemma inorder_delete: "\<lbrakk> bal t ; sorted(inorder t) \<rbrakk> \<Longrightarrow>
@@ -476,23 +476,23 @@
   height_node31 height_node32 height_node33
   height_node41 height_node42 height_node43 height_node44
 
-lemma height_del_min:
-  "del_min t = (x, t') \<Longrightarrow> height t > 0 \<Longrightarrow> bal t \<Longrightarrow> height t' = height t"
-by(induct t arbitrary: x t' rule: del_min.induct)
+lemma height_split_min:
+  "split_min t = (x, t') \<Longrightarrow> height t > 0 \<Longrightarrow> bal t \<Longrightarrow> height t' = height t"
+by(induct t arbitrary: x t' rule: split_min.induct)
   (auto simp: heights split: prod.splits)
 
 lemma height_del: "bal t \<Longrightarrow> height(del x t) = height t"
 by(induction x t rule: del.induct)
-  (auto simp add: heights height_del_min split!: if_split prod.split)
+  (auto simp add: heights height_split_min split!: if_split prod.split)
 
-lemma bal_del_min:
-  "\<lbrakk> del_min t = (x, t'); bal t; height t > 0 \<rbrakk> \<Longrightarrow> bal (tree\<^sub>d t')"
-by(induct t arbitrary: x t' rule: del_min.induct)
-  (auto simp: heights height_del_min bals split: prod.splits)
+lemma bal_split_min:
+  "\<lbrakk> split_min t = (x, t'); bal t; height t > 0 \<rbrakk> \<Longrightarrow> bal (tree\<^sub>d t')"
+by(induct t arbitrary: x t' rule: split_min.induct)
+  (auto simp: heights height_split_min bals split: prod.splits)
 
 lemma bal_tree\<^sub>d_del: "bal t \<Longrightarrow> bal(tree\<^sub>d(del x t))"
 by(induction x t rule: del.induct)
-  (auto simp: bals bal_del_min height_del height_del_min split!: if_split prod.split)
+  (auto simp: bals bal_split_min height_del height_split_min split!: if_split prod.split)
 
 corollary bal_delete: "bal t \<Longrightarrow> bal(delete x t)"
 by(simp add: delete_def bal_tree\<^sub>d_del)
--- a/src/HOL/Data_Structures/Tree23_Map.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/Tree23_Map.thy	Thu May 03 15:07:14 2018 +0200
@@ -57,13 +57,13 @@
 "del x (Node2 l ab1 r) = (case cmp x (fst ab1) of
   LT \<Rightarrow> node21 (del x l) ab1 r |
   GT \<Rightarrow> node22 l ab1 (del x r) |
-  EQ \<Rightarrow> let (ab1',t) = del_min r in node22 l ab1' t)" |
+  EQ \<Rightarrow> let (ab1',t) = split_min r in node22 l ab1' t)" |
 "del x (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of
   LT \<Rightarrow> node31 (del x l) ab1 m ab2 r |
-  EQ \<Rightarrow> let (ab1',m') = del_min m in node32 l ab1' m' ab2 r |
+  EQ \<Rightarrow> let (ab1',m') = split_min m in node32 l ab1' m' ab2 r |
   GT \<Rightarrow> (case cmp x (fst ab2) of
            LT \<Rightarrow> node32 l ab1 (del x m) ab2 r |
-           EQ \<Rightarrow> let (ab2',r') = del_min r in node33 l ab1 m ab2' r' |
+           EQ \<Rightarrow> let (ab2',r') = split_min r in node33 l ab1 m ab2' r' |
            GT \<Rightarrow> node33 l ab1 m ab2 (del x r)))"
 
 definition delete :: "'a::linorder \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
@@ -89,7 +89,7 @@
 lemma inorder_del: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)"
 by(induction t rule: del.induct)
-  (auto simp: del_list_simps inorder_nodes del_minD split!: if_split prod.splits)
+  (auto simp: del_list_simps inorder_nodes split_minD split!: if_split prod.splits)
 
 corollary inorder_delete: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(delete x t) = del_list x (inorder t)"
@@ -107,11 +107,11 @@
 
 lemma height_del: "bal t \<Longrightarrow> height(del x t) = height t"
 by(induction x t rule: del.induct)
-  (auto simp add: heights max_def height_del_min split: prod.split)
+  (auto simp add: heights max_def height_split_min split: prod.split)
 
 lemma bal_tree\<^sub>d_del: "bal t \<Longrightarrow> bal(tree\<^sub>d(del x t))"
 by(induction x t rule: del.induct)
-  (auto simp: bals bal_del_min height_del height_del_min split: prod.split)
+  (auto simp: bals bal_split_min height_del height_split_min split: prod.split)
 
 corollary bal_delete: "bal t \<Longrightarrow> bal(delete x t)"
 by(simp add: delete_def bal_tree\<^sub>d_del)
--- a/src/HOL/Data_Structures/Tree23_Set.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/Tree23_Set.thy	Thu May 03 15:07:14 2018 +0200
@@ -102,13 +102,13 @@
 "node33 t1 a (Node2 t2 b t3) c (Up\<^sub>d t4) = T\<^sub>d(Node2 t1 a (Node3 t2 b t3 c t4))" |
 "node33 t1 a (Node3 t2 b t3 c t4) d (Up\<^sub>d t5) = T\<^sub>d(Node3 t1 a (Node2 t2 b t3) c (Node2 t4 d t5))"
 
-fun del_min :: "'a tree23 \<Rightarrow> 'a * 'a up\<^sub>d" where
-"del_min (Node2 Leaf a Leaf) = (a, Up\<^sub>d Leaf)" |
-"del_min (Node3 Leaf a Leaf b Leaf) = (a, T\<^sub>d(Node2 Leaf b Leaf))" |
-"del_min (Node2 l a r) = (let (x,l') = del_min l in (x, node21 l' a r))" |
-"del_min (Node3 l a m b r) = (let (x,l') = del_min l in (x, node31 l' a m b r))"
+fun split_min :: "'a tree23 \<Rightarrow> 'a * 'a up\<^sub>d" where
+"split_min (Node2 Leaf a Leaf) = (a, Up\<^sub>d Leaf)" |
+"split_min (Node3 Leaf a Leaf b Leaf) = (a, T\<^sub>d(Node2 Leaf b Leaf))" |
+"split_min (Node2 l a r) = (let (x,l') = split_min l in (x, node21 l' a r))" |
+"split_min (Node3 l a m b r) = (let (x,l') = split_min l in (x, node31 l' a m b r))"
 
-text \<open>In the base cases of \<open>del_min\<close> and \<open>del\<close> it is enough to check if one subtree is a \<open>Leaf\<close>,
+text \<open>In the base cases of \<open>split_min\<close> and \<open>del\<close> it is enough to check if one subtree is a \<open>Leaf\<close>,
 in which case balancedness implies that so are the others. Exercise.\<close>
 
 fun del :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d" where
@@ -123,15 +123,15 @@
   (case cmp x a of
      LT \<Rightarrow> node21 (del x l) a r |
      GT \<Rightarrow> node22 l a (del x r) |
-     EQ \<Rightarrow> let (a',t) = del_min r in node22 l a' t)" |
+     EQ \<Rightarrow> let (a',t) = split_min r in node22 l a' t)" |
 "del x (Node3 l a m b r) =
   (case cmp x a of
      LT \<Rightarrow> node31 (del x l) a m b r |
-     EQ \<Rightarrow> let (a',m') = del_min m in node32 l a' m' b r |
+     EQ \<Rightarrow> let (a',m') = split_min m in node32 l a' m' b r |
      GT \<Rightarrow>
        (case cmp x b of
           LT \<Rightarrow> node32 l a (del x m) b r |
-          EQ \<Rightarrow> let (b',r') = del_min r in node33 l a m b' r' |
+          EQ \<Rightarrow> let (b',r') = split_min r in node33 l a m b' r' |
           GT \<Rightarrow> node33 l a m b (del x r)))"
 
 definition delete :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
@@ -182,16 +182,16 @@
 lemmas inorder_nodes = inorder_node21 inorder_node22
   inorder_node31 inorder_node32 inorder_node33
 
-lemma del_minD:
-  "del_min t = (x,t') \<Longrightarrow> bal t \<Longrightarrow> height t > 0 \<Longrightarrow>
+lemma split_minD:
+  "split_min t = (x,t') \<Longrightarrow> bal t \<Longrightarrow> height t > 0 \<Longrightarrow>
   x # inorder(tree\<^sub>d t') = inorder t"
-by(induction t arbitrary: t' rule: del_min.induct)
+by(induction t arbitrary: t' rule: split_min.induct)
   (auto simp: inorder_nodes split: prod.splits)
 
 lemma inorder_del: "\<lbrakk> bal t ; sorted(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)"
 by(induction t rule: del.induct)
-  (auto simp: del_list_simps inorder_nodes del_minD split!: if_split prod.splits)
+  (auto simp: del_list_simps inorder_nodes split_minD split!: if_split prod.splits)
 
 lemma inorder_delete: "\<lbrakk> bal t ; sorted(inorder t) \<rbrakk> \<Longrightarrow>
   inorder(delete x t) = del_list x (inorder t)"
@@ -350,23 +350,23 @@
 lemmas heights = height'_node21 height'_node22
   height'_node31 height'_node32 height'_node33
 
-lemma height_del_min:
-  "del_min t = (x, t') \<Longrightarrow> height t > 0 \<Longrightarrow> bal t \<Longrightarrow> height t' = height t"
-by(induct t arbitrary: x t' rule: del_min.induct)
+lemma height_split_min:
+  "split_min t = (x, t') \<Longrightarrow> height t > 0 \<Longrightarrow> bal t \<Longrightarrow> height t' = height t"
+by(induct t arbitrary: x t' rule: split_min.induct)
   (auto simp: heights split: prod.splits)
 
 lemma height_del: "bal t \<Longrightarrow> height(del x t) = height t"
 by(induction x t rule: del.induct)
-  (auto simp: heights max_def height_del_min split: prod.splits)
+  (auto simp: heights max_def height_split_min split: prod.splits)
 
-lemma bal_del_min:
-  "\<lbrakk> del_min t = (x, t'); bal t; height t > 0 \<rbrakk> \<Longrightarrow> bal (tree\<^sub>d t')"
-by(induct t arbitrary: x t' rule: del_min.induct)
-  (auto simp: heights height_del_min bals split: prod.splits)
+lemma bal_split_min:
+  "\<lbrakk> split_min t = (x, t'); bal t; height t > 0 \<rbrakk> \<Longrightarrow> bal (tree\<^sub>d t')"
+by(induct t arbitrary: x t' rule: split_min.induct)
+  (auto simp: heights height_split_min bals split: prod.splits)
 
 lemma bal_tree\<^sub>d_del: "bal t \<Longrightarrow> bal(tree\<^sub>d(del x t))"
 by(induction x t rule: del.induct)
-  (auto simp: bals bal_del_min height_del height_del_min split: prod.splits)
+  (auto simp: bals bal_split_min height_del height_split_min split: prod.splits)
 
 corollary bal_delete: "bal t \<Longrightarrow> bal(delete x t)"
 by(simp add: delete_def bal_tree\<^sub>d_del)
--- a/src/HOL/Data_Structures/Tree_Map.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/Tree_Map.thy	Thu May 03 15:07:14 2018 +0200
@@ -25,7 +25,7 @@
 "delete x (Node l (a,b) r) = (case cmp x a of
   LT \<Rightarrow> Node (delete x l) (a,b) r |
   GT \<Rightarrow> Node l (a,b) (delete x r) |
-  EQ \<Rightarrow> if r = Leaf then l else let (ab',r') = del_min r in Node l ab' r')"
+  EQ \<Rightarrow> if r = Leaf then l else let (ab',r') = split_min r in Node l ab' r')"
 
 
 subsection "Functional Correctness Proofs"
@@ -40,7 +40,7 @@
 
 lemma inorder_delete:
   "sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
-by(induction t) (auto simp: del_list_simps del_minD split: prod.splits)
+by(induction t) (auto simp: del_list_simps split_minD split: prod.splits)
 
 interpretation Map_by_Ordered
 where empty = Leaf and lookup = lookup and update = update and delete = delete
--- a/src/HOL/Data_Structures/Tree_Set.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Data_Structures/Tree_Set.thy	Thu May 03 15:07:14 2018 +0200
@@ -27,9 +27,9 @@
      EQ \<Rightarrow> Node l a r |
      GT \<Rightarrow> Node l a (insert x r))"
 
-fun del_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
-"del_min (Node l a r) =
-  (if l = Leaf then (a,r) else let (x,l') = del_min l in (x, Node l' a r))"
+fun split_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
+"split_min (Node l a r) =
+  (if l = Leaf then (a,r) else let (x,l') = split_min l in (x, Node l' a r))"
 
 fun delete :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
 "delete x Leaf = Leaf" |
@@ -37,7 +37,7 @@
   (case cmp x a of
      LT \<Rightarrow>  Node (delete x l) a r |
      GT \<Rightarrow>  Node l a (delete x r) |
-     EQ \<Rightarrow> if r = Leaf then l else let (a',r') = del_min r in Node l a' r')"
+     EQ \<Rightarrow> if r = Leaf then l else let (a',r') = split_min r in Node l a' r')"
 
 
 subsection "Functional Correctness Proofs"
@@ -50,14 +50,14 @@
 by(induction t) (auto simp: ins_list_simps)
 
 
-lemma del_minD:
-  "del_min t = (x,t') \<Longrightarrow> t \<noteq> Leaf \<Longrightarrow> x # inorder t' = inorder t"
-by(induction t arbitrary: t' rule: del_min.induct)
+lemma split_minD:
+  "split_min t = (x,t') \<Longrightarrow> t \<noteq> Leaf \<Longrightarrow> x # inorder t' = inorder t"
+by(induction t arbitrary: t' rule: split_min.induct)
   (auto simp: sorted_lems split: prod.splits if_splits)
 
 lemma inorder_delete:
   "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
-by(induction t) (auto simp: del_list_simps del_minD split: prod.splits)
+by(induction t) (auto simp: del_list_simps split_minD split: prod.splits)
 
 interpretation Set_by_Ordered
 where empty = Leaf and isin = isin and insert = insert and delete = delete
--- a/src/HOL/Hull.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Hull.thy	Thu May 03 15:07:14 2018 +0200
@@ -47,7 +47,7 @@
 lemma hull_unique: "s \<subseteq> t \<Longrightarrow> S t \<Longrightarrow> (\<And>t'. s \<subseteq> t' \<Longrightarrow> S t' \<Longrightarrow> t \<subseteq> t') \<Longrightarrow> (S hull s = t)"
   unfolding hull_def by auto
 
-lemma hull_induct: "(\<And>x. x\<in> S \<Longrightarrow> P x) \<Longrightarrow> Q {x. P x} \<Longrightarrow> \<forall>x\<in> Q hull S. P x"
+lemma hull_induct: "\<lbrakk>a \<in> Q hull S; \<And>x. x\<in> S \<Longrightarrow> P x; Q {x. P x}\<rbrakk> \<Longrightarrow> P a"
   using hull_minimal[of S "{x. P x}" Q]
   by (auto simp add: subset_eq)
 
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Thu May 03 15:07:14 2018 +0200
@@ -230,8 +230,10 @@
   obtains "r = x" "h' = h"
   using assms by (rule effectE) (simp add: execute_simps)
 
-definition raise :: "string \<Rightarrow> 'a Heap" where \<comment> \<open>the string is just decoration\<close>
-  [code del]: "raise s = Heap (\<lambda>_. None)"
+definition raise :: "String.literal \<Rightarrow> 'a Heap" \<comment> \<open>the literal is just decoration\<close>
+  where "raise s = Heap (\<lambda>_. None)"
+
+code_datatype raise \<comment> \<open>avoid @{const "Heap"} formally\<close>
 
 lemma execute_raise [execute_simps]:
   "execute (raise s) = (\<lambda>_. None)"
@@ -309,7 +311,7 @@
 subsubsection \<open>Assertions\<close>
 
 definition assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap" where
-  "assert P x = (if P x then return x else raise ''assert'')"
+  "assert P x = (if P x then return x else raise STR ''assert'')"
 
 lemma execute_assert [execute_simps]:
   "P x \<Longrightarrow> execute (assert P x) h = Some (x, h)"
@@ -516,32 +518,17 @@
 
 subsection \<open>Code generator setup\<close>
 
-subsubsection \<open>Logical intermediate layer\<close>
-
-definition raise' :: "String.literal \<Rightarrow> 'a Heap" where
-  [code del]: "raise' s = raise (String.explode s)"
-
-lemma [code_abbrev]: "raise' (STR s) = raise s"
-  unfolding raise'_def by (simp add: STR_inverse)
-
-lemma raise_raise': (* FIXME delete candidate *)
-  "raise s = raise' (STR s)"
-  unfolding raise'_def by (simp add: STR_inverse)
-
-code_datatype raise' \<comment> \<open>avoid @{const "Heap"} formally\<close>
-
-
 subsubsection \<open>SML and OCaml\<close>
 
 code_printing type_constructor Heap \<rightharpoonup> (SML) "(unit/ ->/ _)"
 code_printing constant bind \<rightharpoonup> (SML) "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())"
 code_printing constant return \<rightharpoonup> (SML) "!(fn/ ()/ =>/ _)"
-code_printing constant Heap_Monad.raise' \<rightharpoonup> (SML) "!(raise/ Fail/ _)"
+code_printing constant Heap_Monad.raise \<rightharpoonup> (SML) "!(raise/ Fail/ _)"
 
 code_printing type_constructor Heap \<rightharpoonup> (OCaml) "(unit/ ->/ _)"
 code_printing constant bind \<rightharpoonup> (OCaml) "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())"
 code_printing constant return \<rightharpoonup> (OCaml) "!(fun/ ()/ ->/ _)"
-code_printing constant Heap_Monad.raise' \<rightharpoonup> (OCaml) "failwith"
+code_printing constant Heap_Monad.raise \<rightharpoonup> (OCaml) "failwith"
 
 
 subsubsection \<open>Haskell\<close>
@@ -588,7 +575,7 @@
 code_printing type_constructor Heap \<rightharpoonup> (Haskell) "Heap.ST/ Heap.RealWorld/ _"
 code_monad bind Haskell
 code_printing constant return \<rightharpoonup> (Haskell) "return"
-code_printing constant Heap_Monad.raise' \<rightharpoonup> (Haskell) "error"
+code_printing constant Heap_Monad.raise \<rightharpoonup> (Haskell) "error"
 
 
 subsubsection \<open>Scala\<close>
@@ -633,7 +620,7 @@
 code_printing type_constructor Heap \<rightharpoonup> (Scala) "(Unit/ =>/ _)"
 code_printing constant bind \<rightharpoonup> (Scala) "Heap.bind"
 code_printing constant return \<rightharpoonup> (Scala) "('_: Unit)/ =>/ _"
-code_printing constant Heap_Monad.raise' \<rightharpoonup> (Scala) "!sys.error((_))"
+code_printing constant Heap_Monad.raise \<rightharpoonup> (Scala) "!sys.error((_))"
 
 
 subsubsection \<open>Target variants with less units\<close>
@@ -703,7 +690,7 @@
 
 \<close>
 
-hide_const (open) Heap heap guard raise' fold_map
+hide_const (open) Heap heap guard fold_map
 
 end
 
--- a/src/HOL/Imperative_HOL/ex/SatChecker.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy	Thu May 03 15:07:14 2018 +0200
@@ -173,7 +173,7 @@
 
 primrec res_mem :: "Lit \<Rightarrow> Clause \<Rightarrow> Clause Heap"
 where
-  "res_mem l [] = raise ''MiniSatChecked.res_thm: Cannot find literal''"
+  "res_mem l [] = raise STR ''MiniSatChecked.res_thm: Cannot find literal''"
 | "res_mem l (x#xs) = (if (x = l) then return xs else do { v \<leftarrow> res_mem l xs; return (x # v) })"
 
 fun resolve1 :: "Lit \<Rightarrow> Clause \<Rightarrow> Clause \<Rightarrow> Clause Heap"
@@ -183,7 +183,7 @@
   else (if (x < y) then do { v \<leftarrow> resolve1 l xs (y#ys); return (x # v) }
   else (if (x > y) then do { v \<leftarrow> resolve1 l (x#xs) ys; return (y # v) }
   else do { v \<leftarrow> resolve1 l xs ys; return (x # v) })))"
-| "resolve1 l [] ys = raise ''MiniSatChecked.res_thm: Cannot find literal''"
+| "resolve1 l [] ys = raise STR ''MiniSatChecked.res_thm: Cannot find literal''"
 | "resolve1 l xs [] = res_mem l xs"
 
 fun resolve2 :: "Lit \<Rightarrow> Clause \<Rightarrow> Clause \<Rightarrow> Clause Heap" 
@@ -193,7 +193,7 @@
   else (if (x < y) then do { v \<leftarrow> resolve2 l xs (y#ys); return (x # v) }
   else (if (x > y) then do { v \<leftarrow> resolve2 l (x#xs) ys; return (y # v) }
   else do { v \<leftarrow> resolve2 l xs ys; return (x # v) })))"
-  | "resolve2 l xs [] = raise ''MiniSatChecked.res_thm: Cannot find literal''"
+  | "resolve2 l xs [] = raise STR ''MiniSatChecked.res_thm: Cannot find literal''"
   | "resolve2 l [] ys = res_mem l ys"
 
 fun res_thm' :: "Lit \<Rightarrow> Clause \<Rightarrow> Clause \<Rightarrow> Clause Heap"
@@ -204,8 +204,8 @@
   else (if (x < y) then (res_thm' l xs (y#ys)) \<bind> (\<lambda>v. return (x # v))
   else (if (x > y) then (res_thm' l (x#xs) ys) \<bind> (\<lambda>v. return (y # v))
   else (res_thm' l xs ys) \<bind> (\<lambda>v. return (x # v))))))"
-| "res_thm' l [] ys = raise (''MiniSatChecked.res_thm: Cannot find literal'')"
-| "res_thm' l xs [] = raise (''MiniSatChecked.res_thm: Cannot find literal'')"
+| "res_thm' l [] ys = raise STR ''MiniSatChecked.res_thm: Cannot find literal''"
+| "res_thm' l xs [] = raise STR ''MiniSatChecked.res_thm: Cannot find literal''"
 
 declare res_mem.simps [simp del] resolve1.simps [simp del] resolve2.simps [simp del] res_thm'.simps [simp del]
 
@@ -414,7 +414,7 @@
 where
   "get_clause a i = 
        do { c \<leftarrow> Array.nth a i;
-           (case c of None \<Rightarrow> raise (''Clause not found'')
+           (case c of None \<Rightarrow> raise STR ''Clause not found''
                     | Some x \<Rightarrow> return x)
        }"
 
@@ -422,7 +422,7 @@
 primrec res_thm2 :: "Clause option array \<Rightarrow> (Lit * ClauseId) \<Rightarrow> Clause \<Rightarrow> Clause Heap"
 where
   "res_thm2 a (l, j) cli =
-  ( if l = 0 then raise(''Illegal literal'')
+  ( if l = 0 then raise STR ''Illegal literal''
     else
      do { clj \<leftarrow> get_clause a j;
          res_thm' l cli clj
@@ -445,8 +445,8 @@
   }"
 | "doProofStep2 a (Delete cid) rcs = do { Array.upd cid None a; return rcs }"
 | "doProofStep2 a (Root cid clause) rcs = do { Array.upd cid (Some (remdups (sort clause))) a; return (clause # rcs) }"
-| "doProofStep2 a (Xstep cid1 cid2) rcs = raise ''MiniSatChecked.doProofStep: Xstep constructor found.''"
-| "doProofStep2 a (ProofDone b) rcs = raise ''MiniSatChecked.doProofStep: ProofDone constructor found.''"
+| "doProofStep2 a (Xstep cid1 cid2) rcs = raise STR ''MiniSatChecked.doProofStep: Xstep constructor found.''"
+| "doProofStep2 a (ProofDone b) rcs = raise STR ''MiniSatChecked.doProofStep: ProofDone constructor found.''"
 
 definition checker :: "nat \<Rightarrow> ProofStep list \<Rightarrow> nat \<Rightarrow> Clause list Heap"
 where
@@ -456,7 +456,7 @@
      rcs \<leftarrow> foldM (doProofStep2 a) p [];
      ec \<leftarrow> Array.nth a i;
      (if ec = Some [] then return rcs 
-                else raise(''No empty clause''))
+                else raise STR ''No empty clause'')
   }"
 
 lemma effect_case_option:
@@ -641,24 +641,24 @@
 primrec lres_thm :: "Clause option list \<Rightarrow> (Lit * ClauseId) \<Rightarrow> Clause \<Rightarrow> Clause Heap" 
 where
   "lres_thm xs (l, j) cli = (if (j < List.length xs) then (case (xs ! j) of
-  None \<Rightarrow> raise (''MiniSatChecked.res_thm: No resolvant clause in thms array for Conflict step.'')
+  None \<Rightarrow> raise STR ''MiniSatChecked.res_thm: No resolvant clause in thms array for Conflict step.''
   | Some clj \<Rightarrow> res_thm' l cli clj
- ) else raise ''Error'')"
+ ) else raise STR ''Error'')"
 
 
 fun ldoProofStep :: " ProofStep \<Rightarrow> (Clause option list  * Clause list) \<Rightarrow> (Clause option list * Clause list) Heap"
 where
   "ldoProofStep (Conflict saveTo (i, rs)) (xs, rcl) =
      (case (xs ! i) of
-       None \<Rightarrow> raise (''MiniSatChecked.doProofStep: No starting clause in thms array for Conflict step.'')
+       None \<Rightarrow> raise STR ''MiniSatChecked.doProofStep: No starting clause in thms array for Conflict step.''
      | Some cli \<Rightarrow> do {
                       result \<leftarrow> foldM (lres_thm xs) rs cli ;
                       return ((xs[saveTo:=Some result]), rcl)
                    })"
 | "ldoProofStep (Delete cid) (xs, rcl) = return (xs[cid:=None], rcl)"
 | "ldoProofStep (Root cid clause) (xs, rcl) = return (xs[cid:=Some (sort clause)], (remdups(sort clause)) # rcl)"
-| "ldoProofStep (Xstep cid1 cid2) (xs, rcl) = raise ''MiniSatChecked.doProofStep: Xstep constructor found.''"
-| "ldoProofStep (ProofDone b) (xs, rcl) = raise ''MiniSatChecked.doProofStep: ProofDone constructor found.''"
+| "ldoProofStep (Xstep cid1 cid2) (xs, rcl) = raise STR ''MiniSatChecked.doProofStep: Xstep constructor found.''"
+| "ldoProofStep (ProofDone b) (xs, rcl) = raise STR ''MiniSatChecked.doProofStep: ProofDone constructor found.''"
 
 definition lchecker :: "nat \<Rightarrow> ProofStep list \<Rightarrow> nat \<Rightarrow> Clause list Heap"
 where
@@ -666,7 +666,7 @@
   do {
      rcs \<leftarrow> foldM (ldoProofStep) p ([], []);
      (if (fst rcs ! i) = Some [] then return (snd rcs) 
-                else raise(''No empty clause''))
+                else raise STR ''No empty clause'')
   }"
 
 
@@ -676,22 +676,22 @@
 where
   "tres_thm t (l, j) cli =
   (case (rbt_lookup t j) of 
-     None \<Rightarrow> raise (''MiniSatChecked.res_thm: No resolvant clause in thms array for Conflict step.'')
+     None \<Rightarrow> raise STR ''MiniSatChecked.res_thm: No resolvant clause in thms array for Conflict step.''
    | Some clj \<Rightarrow> res_thm' l cli clj)"
 
 fun tdoProofStep :: " ProofStep \<Rightarrow> ((ClauseId, Clause) RBT_Impl.rbt * Clause list) \<Rightarrow> ((ClauseId, Clause) RBT_Impl.rbt * Clause list) Heap"
 where
   "tdoProofStep (Conflict saveTo (i, rs)) (t, rcl) =
      (case (rbt_lookup t i) of
-       None \<Rightarrow> raise (''MiniSatChecked.doProofStep: No starting clause in thms array for Conflict step.'')
+       None \<Rightarrow> raise STR ''MiniSatChecked.doProofStep: No starting clause in thms array for Conflict step.''
      | Some cli \<Rightarrow> do {
                       result \<leftarrow> foldM (tres_thm t) rs cli;
                       return ((rbt_insert saveTo result t), rcl)
                    })"
 | "tdoProofStep (Delete cid) (t, rcl) = return ((rbt_delete cid t), rcl)"
 | "tdoProofStep (Root cid clause) (t, rcl) = return (rbt_insert cid (sort clause) t, (remdups(sort clause)) # rcl)"
-| "tdoProofStep (Xstep cid1 cid2) (t, rcl) = raise ''MiniSatChecked.doProofStep: Xstep constructor found.''"
-| "tdoProofStep (ProofDone b) (t, rcl) = raise ''MiniSatChecked.doProofStep: ProofDone constructor found.''"
+| "tdoProofStep (Xstep cid1 cid2) (t, rcl) = raise STR ''MiniSatChecked.doProofStep: Xstep constructor found.''"
+| "tdoProofStep (ProofDone b) (t, rcl) = raise STR ''MiniSatChecked.doProofStep: ProofDone constructor found.''"
 
 definition tchecker :: "nat \<Rightarrow> ProofStep list \<Rightarrow> nat \<Rightarrow> Clause list Heap"
 where
@@ -699,7 +699,7 @@
   do {
      rcs \<leftarrow> foldM (tdoProofStep) p (RBT_Impl.Empty, []);
      (if (rbt_lookup (fst rcs) i) = Some [] then return (snd rcs) 
-                else raise(''No empty clause''))
+                else raise STR ''No empty clause'')
   }"
 
 export_code checker tchecker lchecker checking SML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Adhoc_Overloading.thy	Thu May 03 15:07:14 2018 +0200
@@ -0,0 +1,15 @@
+(*  Title:      HOL/Library/Adhoc_Overloading.thy
+    Author:     Alexander Krauss, TU Muenchen
+    Author:     Christian Sternagel, University of Innsbruck
+*)
+
+section \<open>Adhoc overloading of constants based on their types\<close>
+
+theory Adhoc_Overloading
+  imports Main
+  keywords "adhoc_overloading" "no_adhoc_overloading" :: thy_decl
+begin
+
+ML_file "adhoc_overloading.ML"
+
+end
--- a/src/HOL/Library/Cardinality.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Cardinality.thy	Thu May 03 15:07:14 2018 +0200
@@ -27,15 +27,6 @@
 lemma (in type_definition) card: "card (UNIV :: 'b set) = card A"
   by (simp add: univ card_image inj_on_def Abs_inject)
 
-lemma finite_range_Some: "finite (range (Some :: 'a \<Rightarrow> 'a option)) = finite (UNIV :: 'a set)"
-by(auto dest: finite_imageD intro: inj_Some)
-
-lemma infinite_literal: "\<not> finite (UNIV :: String.literal set)"
-proof -
-  have "inj STR" by(auto intro: injI)
-  thus ?thesis
-    by(auto simp add: type_definition.univ[OF type_definition_literal] infinite_UNIV_listI dest: finite_imageD)
-qed
 
 subsection \<open>Cardinalities of types\<close>
 
--- a/src/HOL/Library/Char_ord.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Char_ord.thy	Thu May 03 15:07:14 2018 +0200
@@ -11,27 +11,29 @@
 instantiation char :: linorder
 begin
 
-definition "c1 \<le> c2 \<longleftrightarrow> nat_of_char c1 \<le> nat_of_char c2"
-definition "c1 < c2 \<longleftrightarrow> nat_of_char c1 < nat_of_char c2"
+definition less_eq_char :: "char \<Rightarrow> char \<Rightarrow> bool"
+  where "c1 \<le> c2 \<longleftrightarrow> of_char c1 \<le> (of_char c2 :: nat)"
+
+definition less_char :: "char \<Rightarrow> char \<Rightarrow> bool"
+  where "c1 < c2 \<longleftrightarrow> of_char c1 < (of_char c2 :: nat)"
+
 
 instance
   by standard (auto simp add: less_eq_char_def less_char_def)
 
 end
 
-lemma less_eq_char_simps:
-  "0 \<le> c"
-  "Char k \<le> 0 \<longleftrightarrow> numeral k mod 256 = (0 :: nat)"
-  "Char k \<le> Char l \<longleftrightarrow> numeral k mod 256 \<le> (numeral l mod 256 :: nat)"
-  for c :: char
-  by (simp_all add: Char_def less_eq_char_def)
+lemma less_eq_char_simp [simp]:
+  "Char b0 b1 b2 b3 b4 b5 b6 b7 \<le> Char c0 c1 c2 c3 c4 c5 c6 c7
+    \<longleftrightarrow> foldr (\<lambda>b k. of_bool b + k * 2) [b0, b1, b2, b3, b4, b5, b6, b7] 0
+      \<le> foldr (\<lambda>b k. of_bool b + k * 2) [c0, c1, c2, c3, c4, c5, c6, c7] (0::nat)"
+  by (simp add: less_eq_char_def)
 
-lemma less_char_simps:
-  "\<not> c < 0"
-  "0 < Char k \<longleftrightarrow> (0 :: nat) < numeral k mod 256"
-  "Char k < Char l \<longleftrightarrow> numeral k mod 256 < (numeral l mod 256 :: nat)"
-  for c :: char
-  by (simp_all add: Char_def less_char_def)
+lemma less_char_simp [simp]:
+  "Char b0 b1 b2 b3 b4 b5 b6 b7 < Char c0 c1 c2 c3 c4 c5 c6 c7
+    \<longleftrightarrow> foldr (\<lambda>b k. of_bool b + k * 2) [b0, b1, b2, b3, b4, b5, b6, b7] 0
+      < foldr (\<lambda>b k. of_bool b + k * 2) [c0, c1, c2, c3, c4, c5, c6, c7] (0::nat)"
+  by (simp add: less_char_def)
 
 instantiation char :: distrib_lattice
 begin
@@ -44,39 +46,4 @@
 
 end
 
-instantiation String.literal :: linorder
-begin
-
-context includes literal.lifting
-begin
-
-lift_definition less_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool"
-  is "ord.lexordp (<)" .
-
-lift_definition less_eq_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool"
-  is "ord.lexordp_eq (<)" .
-
-instance
-proof -
-  interpret linorder "ord.lexordp_eq (<)" "ord.lexordp (<) :: string \<Rightarrow> string \<Rightarrow> bool"
-    by (rule linorder.lexordp_linorder[where less_eq="(\<le>)"]) unfold_locales
-  show "PROP ?thesis"
-    by intro_classes (transfer, simp add: less_le_not_le linear)+
-qed
-
 end
-
-end
-
-lemma less_literal_code [code]:
-  "(<) = (\<lambda>xs ys. ord.lexordp (<) (String.explode xs) (String.explode ys))"
-  by (simp add: less_literal.rep_eq fun_eq_iff)
-
-lemma less_eq_literal_code [code]:
-  "(\<le>) = (\<lambda>xs ys. ord.lexordp_eq (<) (String.explode xs) (String.explode ys))"
-  by (simp add: less_eq_literal.rep_eq fun_eq_iff)
-
-lifting_update literal.lifting
-lifting_forget literal.lifting
-
-end
--- a/src/HOL/Library/Code_Char.thy	Wed May 02 13:49:38 2018 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,96 +0,0 @@
-(*  Title:      HOL/Library/Code_Char.thy
-    Author:     Florian Haftmann
-*)
-
-section \<open>Code generation of pretty characters (and strings)\<close>
-
-theory Code_Char
-imports Main Char_ord
-begin
-
-code_printing
-  type_constructor char \<rightharpoonup>
-    (SML) "char"
-    and (OCaml) "char"
-    and (Haskell) "Prelude.Char"
-    and (Scala) "Char"
-
-setup \<open>
-  fold String_Code.add_literal_char ["SML", "OCaml", "Haskell", "Scala"]
-  #> String_Code.add_literal_list_string "Haskell"
-\<close>
-
-code_printing
-  constant integer_of_char \<rightharpoonup>
-    (SML) "!(IntInf.fromInt o Char.ord)"
-    and (OCaml) "Big'_int.big'_int'_of'_int (Char.code _)"
-    and (Haskell) "Prelude.toInteger (Prelude.fromEnum (_ :: Prelude.Char))"
-    and (Scala) "BigInt(_.toInt)"
-| constant char_of_integer \<rightharpoonup>
-    (SML) "!(Char.chr o IntInf.toInt)"
-    and (OCaml) "Char.chr (Big'_int.int'_of'_big'_int _)"
-    and (Haskell) "!(let chr k | (0 <= k && k < 256) = Prelude.toEnum k :: Prelude.Char in chr . Prelude.fromInteger)"
-    and (Scala) "!((k: BigInt) => if (BigInt(0) <= k && k < BigInt(256)) k.charValue else sys.error(\"character value out of range\"))"
-| class_instance char :: equal \<rightharpoonup>
-    (Haskell) -
-| constant "HOL.equal :: char \<Rightarrow> char \<Rightarrow> bool" \<rightharpoonup>
-    (SML) "!((_ : char) = _)"
-    and (OCaml) "!((_ : char) = _)"
-    and (Haskell) infix 4 "=="
-    and (Scala) infixl 5 "=="
-| constant "Code_Evaluation.term_of :: char \<Rightarrow> term" \<rightharpoonup>
-    (Eval) "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))"
-
-code_reserved SML
-  char
-
-code_reserved OCaml
-  char
-
-code_reserved Scala
-  char
-
-code_reserved SML String
-
-code_printing
-  constant String.implode \<rightharpoonup>
-    (SML) "String.implode"
-    and (OCaml) "!(let l = _ in let res = String.create (List.length l) in let rec imp i = function | [] -> res | c :: l -> String.set res i c; imp (i + 1) l in imp 0 l)"
-    and (Haskell) "_"
-    and (Scala) "!(\"\" ++/ _)"
-| constant String.explode \<rightharpoonup>
-    (SML) "String.explode"
-    and (OCaml) "!(let s = _ in let rec exp i l = if i < 0 then l else exp (i - 1) (String.get s i :: l) in exp (String.length s - 1) [])"
-    and (Haskell) "_"
-    and (Scala) "!(_.toList)"
-
-code_printing
-  constant "Orderings.less_eq :: char \<Rightarrow> char \<Rightarrow> bool" \<rightharpoonup>
-    (SML) "!((_ : char) <= _)"
-    and (OCaml) "!((_ : char) <= _)"
-    and (Haskell) infix 4 "<="
-    and (Scala) infixl 4 "<="
-    and (Eval) infixl 6 "<="
-| constant "Orderings.less :: char \<Rightarrow> char \<Rightarrow> bool" \<rightharpoonup>
-    (SML) "!((_ : char) < _)"
-    and (OCaml) "!((_ : char) < _)"
-    and (Haskell) infix 4 "<"
-    and (Scala) infixl 4 "<"
-    and (Eval) infixl 6 "<"
-|  constant "Orderings.less_eq :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
-    (SML) "!((_ : string) <= _)"
-    and (OCaml) "!((_ : string) <= _)"
-    \<comment> \<open>Order operations for @{typ String.literal} work in Haskell only
-          if no type class instance needs to be generated, because String = [Char] in Haskell
-          and @{typ "char list"} need not have the same order as @{typ String.literal}.\<close>
-    and (Haskell) infix 4 "<="
-    and (Scala) infixl 4 "<="
-    and (Eval) infixl 6 "<="
-| constant "Orderings.less :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
-    (SML) "!((_ : string) < _)"
-    and (OCaml) "!((_ : string) < _)"
-    and (Haskell) infix 4 "<"
-    and (Scala) infixl 4 "<"
-    and (Eval) infixl 6 "<"
-
-end
--- a/src/HOL/Library/Code_Target_Int.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Code_Target_Int.thy	Thu May 03 15:07:14 2018 +0200
@@ -143,6 +143,22 @@
   "nat = nat_of_integer \<circ> of_int"
   including integer.lifting by transfer (simp add: fun_eq_iff)
 
+definition char_of_int :: "int \<Rightarrow> char"
+  where [code_abbrev]: "char_of_int = char_of"
+
+definition int_of_char :: "char \<Rightarrow> int"
+  where [code_abbrev]: "int_of_char = of_char"
+
+lemma [code]:
+  "char_of_int = char_of_integer \<circ> integer_of_int"
+  including integer.lifting unfolding char_of_integer_def char_of_int_def
+  by transfer (simp add: fun_eq_iff)
+
+lemma [code]:
+  "int_of_char = int_of_integer \<circ> integer_of_char"
+  including integer.lifting unfolding integer_of_char_def int_of_char_def
+  by transfer (simp add: fun_eq_iff)
+
 code_identifier
   code_module Code_Target_Int \<rightharpoonup>
     (SML) Arith and (OCaml) Arith and (Haskell) Arith
--- a/src/HOL/Library/Code_Target_Nat.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Code_Target_Nat.thy	Thu May 03 15:07:14 2018 +0200
@@ -147,6 +147,21 @@
   "integer_of_nat (nat k) = max 0 (integer_of_int k)"
   including integer.lifting by transfer auto
 
+definition char_of_nat :: "nat \<Rightarrow> char"
+  where [code_abbrev]: "char_of_nat = char_of"
+
+definition nat_of_char :: "char \<Rightarrow> nat"
+  where [code_abbrev]: "nat_of_char = of_char"
+
+lemma [code]:
+  "char_of_nat = char_of_integer \<circ> integer_of_nat"
+  including integer.lifting unfolding char_of_integer_def char_of_nat_def
+  by transfer (simp add: fun_eq_iff)
+
+lemma [code abstract]:
+  "integer_of_nat (nat_of_char c) = integer_of_char c"
+  by (cases c) (simp add: nat_of_char_def integer_of_char_def integer_of_nat_eq_of_nat)
+
 lemma term_of_nat_code [code]:
   \<comment> \<open>Use @{term Code_Numeral.nat_of_integer} in term reconstruction
         instead of @{term Code_Target_Nat.Nat} such that reconstructed
--- a/src/HOL/Library/Code_Test.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Code_Test.thy	Thu May 03 15:07:14 2018 +0200
@@ -84,8 +84,8 @@
 
 definition list where "list f xs = map (node \<circ> f) xs"
 
-definition X :: yxml_of_term where "X = yot_literal (STR [Char (num.Bit1 (num.Bit0 num.One))])"
-definition Y :: yxml_of_term where "Y = yot_literal (STR [Char (num.Bit0 (num.Bit1 num.One))])"
+definition X :: yxml_of_term where "X = yot_literal (STR 0x05)"
+definition Y :: yxml_of_term where "Y = yot_literal (STR 0x06)"
 definition XY :: yxml_of_term where "XY = yot_append X Y"
 definition XYX :: yxml_of_term where "XYX = yot_append XY X"
 
--- a/src/HOL/Library/Countable.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Countable.thy	Thu May 03 15:07:14 2018 +0200
@@ -256,7 +256,7 @@
 text \<open>String literals\<close>
 
 instance String.literal :: countable
-  by (rule countable_classI [of "to_nat \<circ> String.explode"]) (auto simp add: explode_inject)
+  by (rule countable_classI [of "to_nat \<circ> String.explode"]) (simp add: String.explode_inject)
 
 text \<open>Functions\<close>
 
--- a/src/HOL/Library/Library.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Library.thy	Thu May 03 15:07:14 2018 +0200
@@ -2,6 +2,7 @@
 theory Library
 imports
   AList
+  Adhoc_Overloading
   BigO
   Bit
   BNF_Axiomatization
--- a/src/HOL/Library/Monad_Syntax.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Monad_Syntax.thy	Thu May 03 15:07:14 2018 +0200
@@ -1,11 +1,12 @@
-(* Author: Alexander Krauss, TU Muenchen
-   Author: Christian Sternagel, University of Innsbruck
+(*  Title:      HOL/Library/Monad_Syntax.thy
+    Author:     Alexander Krauss, TU Muenchen
+    Author:     Christian Sternagel, University of Innsbruck
 *)
 
 section \<open>Monad notation for arbitrary types\<close>
 
 theory Monad_Syntax
-imports Main "~~/src/Tools/Adhoc_Overloading"
+  imports Adhoc_Overloading
 begin
 
 text \<open>
--- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Thu May 03 15:07:14 2018 +0200
@@ -244,7 +244,7 @@
 
 section \<open>Setup for String.literal\<close>
 
-setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name "STR"}]\<close>
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name String.Literal}]\<close>
 
 section \<open>Simplification rules for optimisation\<close>
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/adhoc_overloading.ML	Thu May 03 15:07:14 2018 +0200
@@ -0,0 +1,246 @@
+(*  Author:     Alexander Krauss, TU Muenchen
+    Author:     Christian Sternagel, University of Innsbruck
+
+Adhoc overloading of constants based on their types.
+*)
+
+signature ADHOC_OVERLOADING =
+sig
+  val is_overloaded: Proof.context -> string -> bool
+  val generic_add_overloaded: string -> Context.generic -> Context.generic
+  val generic_remove_overloaded: string -> Context.generic -> Context.generic
+  val generic_add_variant: string -> term -> Context.generic -> Context.generic
+  (*If the list of variants is empty at the end of "generic_remove_variant", then
+  "generic_remove_overloaded" is called implicitly.*)
+  val generic_remove_variant: string -> term -> Context.generic -> Context.generic
+  val show_variants: bool Config.T
+end
+
+structure Adhoc_Overloading: ADHOC_OVERLOADING =
+struct
+
+val show_variants = Attrib.setup_config_bool @{binding show_variants} (K false);
+
+
+(* errors *)
+
+fun err_duplicate_variant oconst =
+  error ("Duplicate variant of " ^ quote oconst);
+
+fun err_not_a_variant oconst =
+  error ("Not a variant of " ^  quote oconst);
+
+fun err_not_overloaded oconst =
+  error ("Constant " ^ quote oconst ^ " is not declared as overloaded");
+
+fun err_unresolved_overloading ctxt0 (c, T) t instances =
+  let
+    val ctxt = Config.put show_variants true ctxt0
+    val const_space = Proof_Context.const_space ctxt
+    val prt_const =
+      Pretty.block [Name_Space.pretty ctxt const_space c, Pretty.str " ::", Pretty.brk 1,
+        Pretty.quote (Syntax.pretty_typ ctxt T)]
+  in
+    error (Pretty.string_of (Pretty.chunks [
+      Pretty.block [
+        Pretty.str "Unresolved adhoc overloading of constant", Pretty.brk 1,
+        prt_const, Pretty.brk 1, Pretty.str "in term", Pretty.brk 1,
+        Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t)]],
+      Pretty.block (
+        (if null instances then [Pretty.str "no instances"]
+        else Pretty.fbreaks (
+          Pretty.str "multiple instances:" ::
+          map (Pretty.mark Markup.item o Syntax.pretty_term ctxt) instances)))]))
+  end;
+
+
+(* generic data *)
+
+fun variants_eq ((v1, T1), (v2, T2)) =
+  Term.aconv_untyped (v1, v2) andalso T1 = T2;
+
+structure Overload_Data = Generic_Data
+(
+  type T =
+    {variants : (term * typ) list Symtab.table,
+     oconsts : string Termtab.table};
+  val empty = {variants = Symtab.empty, oconsts = Termtab.empty};
+  val extend = I;
+
+  fun merge
+    ({variants = vtab1, oconsts = otab1},
+     {variants = vtab2, oconsts = otab2}) : T =
+    let
+      fun merge_oconsts _ (oconst1, oconst2) =
+        if oconst1 = oconst2 then oconst1
+        else err_duplicate_variant oconst1;
+    in
+      {variants = Symtab.merge_list variants_eq (vtab1, vtab2),
+       oconsts = Termtab.join merge_oconsts (otab1, otab2)}
+    end;
+);
+
+fun map_tables f g =
+  Overload_Data.map (fn {variants = vtab, oconsts = otab} =>
+    {variants = f vtab, oconsts = g otab});
+
+val is_overloaded = Symtab.defined o #variants o Overload_Data.get o Context.Proof;
+val get_variants = Symtab.lookup o #variants o Overload_Data.get o Context.Proof;
+val get_overloaded = Termtab.lookup o #oconsts o Overload_Data.get o Context.Proof;
+
+fun generic_add_overloaded oconst context =
+  if is_overloaded (Context.proof_of context) oconst then context
+  else map_tables (Symtab.update (oconst, [])) I context;
+
+fun generic_remove_overloaded oconst context =
+  let
+    fun remove_oconst_and_variants context oconst =
+      let
+        val remove_variants =
+          (case get_variants (Context.proof_of context) oconst of
+            NONE => I
+          | SOME vs => fold (Termtab.remove (op =) o rpair oconst o fst) vs);
+      in map_tables (Symtab.delete_safe oconst) remove_variants context end;
+  in
+    if is_overloaded (Context.proof_of context) oconst then remove_oconst_and_variants context oconst
+    else err_not_overloaded oconst
+  end;
+
+local
+  fun generic_variant add oconst t context =
+    let
+      val ctxt = Context.proof_of context;
+      val _ = if is_overloaded ctxt oconst then () else err_not_overloaded oconst;
+      val T = t |> fastype_of;
+      val t' = Term.map_types (K dummyT) t;
+    in
+      if add then
+        let
+          val _ =
+            (case get_overloaded ctxt t' of
+              NONE => ()
+            | SOME oconst' => err_duplicate_variant oconst');
+        in
+          map_tables (Symtab.cons_list (oconst, (t', T))) (Termtab.update (t', oconst)) context
+        end
+      else
+        let
+          val _ =
+            if member variants_eq (the (get_variants ctxt oconst)) (t', T) then ()
+            else err_not_a_variant oconst;
+        in
+          map_tables (Symtab.map_entry oconst (remove1 variants_eq (t', T)))
+            (Termtab.delete_safe t') context
+          |> (fn context =>
+            (case get_variants (Context.proof_of context) oconst of
+              SOME [] => generic_remove_overloaded oconst context
+            | _ => context))
+        end
+    end;
+in
+  val generic_add_variant = generic_variant true;
+  val generic_remove_variant = generic_variant false;
+end;
+
+
+(* check / uncheck *)
+
+fun unifiable_with thy T1 T2 =
+  let
+    val maxidx1 = Term.maxidx_of_typ T1;
+    val T2' = Logic.incr_tvar (maxidx1 + 1) T2;
+    val maxidx2 = Term.maxidx_typ T2' maxidx1;
+  in can (Sign.typ_unify thy (T1, T2')) (Vartab.empty, maxidx2) end;
+
+fun get_candidates ctxt (c, T) =
+  get_variants ctxt c
+  |> Option.map (map_filter (fn (t, T') =>
+    if unifiable_with (Proof_Context.theory_of ctxt) T T' then SOME t
+    else NONE));
+
+fun insert_variants ctxt t (oconst as Const (c, T)) =
+      (case get_candidates ctxt (c, T) of
+        SOME [] => err_unresolved_overloading ctxt (c, T) t []
+      | SOME [variant] => variant
+      | _ => oconst)
+  | insert_variants _ _ oconst = oconst;
+
+fun insert_overloaded ctxt =
+  let
+    fun proc t =
+      Term.map_types (K dummyT) t
+      |> get_overloaded ctxt
+      |> Option.map (Const o rpair (Term.type_of t));
+  in
+    Pattern.rewrite_term_top (Proof_Context.theory_of ctxt) [] [proc]
+  end;
+
+fun check ctxt =
+  map (fn t => Term.map_aterms (insert_variants ctxt t) t);
+
+fun uncheck ctxt ts =
+  if Config.get ctxt show_variants orelse exists (is_none o try Term.type_of) ts then ts
+  else map (insert_overloaded ctxt) ts;
+
+fun reject_unresolved ctxt =
+  let
+    val the_candidates = the o get_candidates ctxt;
+    fun check_unresolved t =
+      (case filter (is_overloaded ctxt o fst) (Term.add_consts t []) of
+        [] => t
+      | (cT :: _) => err_unresolved_overloading ctxt cT t (the_candidates cT));
+  in map check_unresolved end;
+
+
+(* setup *)
+
+val _ = Context.>>
+  (Syntax_Phases.term_check 0 "adhoc_overloading" check
+   #> Syntax_Phases.term_check 1 "adhoc_overloading_unresolved_check" reject_unresolved
+   #> Syntax_Phases.term_uncheck 0 "adhoc_overloading" uncheck);
+
+
+(* commands *)
+
+fun generic_adhoc_overloading_cmd add =
+  if add then
+    fold (fn (oconst, ts) =>
+      generic_add_overloaded oconst
+      #> fold (generic_add_variant oconst) ts)
+  else
+    fold (fn (oconst, ts) =>
+      fold (generic_remove_variant oconst) ts);
+
+fun adhoc_overloading_cmd' add args phi =
+  let val args' = args
+    |> map (apsnd (map_filter (fn t =>
+         let val t' = Morphism.term phi t;
+         in if Term.aconv_untyped (t, t') then SOME t' else NONE end)));
+  in generic_adhoc_overloading_cmd add args' end;
+
+fun adhoc_overloading_cmd add raw_args lthy =
+  let
+    fun const_name ctxt =
+      fst o dest_Const o Proof_Context.read_const {proper = false, strict = false} ctxt;  (* FIXME {proper = true, strict = true} (!?) *)
+    fun read_term ctxt = singleton (Variable.polymorphic ctxt) o Syntax.read_term ctxt;
+    val args =
+      raw_args
+      |> map (apfst (const_name lthy))
+      |> map (apsnd (map (read_term lthy)));
+  in
+    Local_Theory.declaration {syntax = true, pervasive = false}
+      (adhoc_overloading_cmd' add args) lthy
+  end;
+
+val _ =
+  Outer_Syntax.local_theory @{command_keyword adhoc_overloading}
+    "add adhoc overloading for constants / fixed variables"
+    (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd true);
+
+val _ =
+  Outer_Syntax.local_theory @{command_keyword no_adhoc_overloading}
+    "delete adhoc overloading for constants / fixed variables"
+    (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd false);
+
+end;
+
--- a/src/HOL/Limits.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Limits.thy	Thu May 03 15:07:14 2018 +0200
@@ -787,15 +787,32 @@
 lemmas tendsto_scaleR [tendsto_intros] =
   bounded_bilinear.tendsto [OF bounded_bilinear_scaleR]
 
-lemmas tendsto_mult [tendsto_intros] =
-  bounded_bilinear.tendsto [OF bounded_bilinear_mult]
+
+text\<open>Analogous type class for multiplication\<close>
+class topological_semigroup_mult = topological_space + semigroup_mult +
+  assumes tendsto_mult_Pair: "LIM x (nhds a \<times>\<^sub>F nhds b). fst x * snd x :> nhds (a * b)"
+
+instance real_normed_algebra < topological_semigroup_mult
+proof
+  fix a b :: 'a
+  show "((\<lambda>x. fst x * snd x) \<longlongrightarrow> a * b) (nhds a \<times>\<^sub>F nhds b)"
+    unfolding nhds_prod[symmetric]
+    using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"]
+    by (simp add: bounded_bilinear.tendsto [OF bounded_bilinear_mult])
+qed
+
+lemma tendsto_mult [tendsto_intros]:
+  fixes a b :: "'a::topological_semigroup_mult"
+  shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x * g x) \<longlongrightarrow> a * b) F"
+  using filterlim_compose[OF tendsto_mult_Pair, of "\<lambda>x. (f x, g x)" a b F]
+  by (simp add: nhds_prod[symmetric] tendsto_Pair)
 
 lemma tendsto_mult_left: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) \<longlongrightarrow> c * l) F"
-  for c :: "'a::real_normed_algebra"
+  for c :: "'a::topological_semigroup_mult"
   by (rule tendsto_mult [OF tendsto_const])
 
 lemma tendsto_mult_right: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) \<longlongrightarrow> l * c) F"
-  for c :: "'a::real_normed_algebra"
+  for c :: "'a::topological_semigroup_mult"
   by (rule tendsto_mult [OF _ tendsto_const])
 
 lemmas continuous_of_real [continuous_intros] =
@@ -2069,14 +2086,14 @@
 qed
 
 lemma convergent_add:
-  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
+  fixes X Y :: "nat \<Rightarrow> 'a::topological_monoid_add"
   assumes "convergent (\<lambda>n. X n)"
     and "convergent (\<lambda>n. Y n)"
   shows "convergent (\<lambda>n. X n + Y n)"
   using assms unfolding convergent_def by (blast intro: tendsto_add)
 
 lemma convergent_sum:
-  fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
+  fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::topological_comm_monoid_add"
   shows "(\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)) \<Longrightarrow> convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
   by (induct A rule: infinite_finite_induct) (simp_all add: convergent_const convergent_add)
 
@@ -2091,16 +2108,13 @@
   shows "convergent (\<lambda>n. X n ** Y n)"
   using assms unfolding convergent_def by (blast intro: tendsto)
 
-lemma convergent_minus_iff: "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
-  for X :: "nat \<Rightarrow> 'a::real_normed_vector"
-  apply (simp add: convergent_def)
-  apply (auto dest: tendsto_minus)
-  apply (drule tendsto_minus)
-  apply auto
-  done
+lemma convergent_minus_iff:
+  fixes X :: "nat \<Rightarrow> 'a::topological_group_add"
+  shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
+  unfolding convergent_def by (force dest: tendsto_minus)
 
 lemma convergent_diff:
-  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
+  fixes X Y :: "nat \<Rightarrow> 'a::topological_group_add"
   assumes "convergent (\<lambda>n. X n)"
   assumes "convergent (\<lambda>n. Y n)"
   shows "convergent (\<lambda>n. X n - Y n)"
@@ -2123,7 +2137,7 @@
   unfolding convergent_def by (blast intro!: tendsto_of_real)
 
 lemma convergent_add_const_iff:
-  "convergent (\<lambda>n. c + f n :: 'a::real_normed_vector) \<longleftrightarrow> convergent f"
+  "convergent (\<lambda>n. c + f n :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f"
 proof
   assume "convergent (\<lambda>n. c + f n)"
   from convergent_diff[OF this convergent_const[of c]] show "convergent f"
@@ -2135,15 +2149,15 @@
 qed
 
 lemma convergent_add_const_right_iff:
-  "convergent (\<lambda>n. f n + c :: 'a::real_normed_vector) \<longleftrightarrow> convergent f"
+  "convergent (\<lambda>n. f n + c :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f"
   using convergent_add_const_iff[of c f] by (simp add: add_ac)
 
 lemma convergent_diff_const_right_iff:
-  "convergent (\<lambda>n. f n - c :: 'a::real_normed_vector) \<longleftrightarrow> convergent f"
+  "convergent (\<lambda>n. f n - c :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f"
   using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac)
 
 lemma convergent_mult:
-  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
+  fixes X Y :: "nat \<Rightarrow> 'a::topological_semigroup_mult"
   assumes "convergent (\<lambda>n. X n)"
     and "convergent (\<lambda>n. Y n)"
   shows "convergent (\<lambda>n. X n * Y n)"
@@ -2151,7 +2165,7 @@
 
 lemma convergent_mult_const_iff:
   assumes "c \<noteq> 0"
-  shows "convergent (\<lambda>n. c * f n :: 'a::real_normed_field) \<longleftrightarrow> convergent f"
+  shows "convergent (\<lambda>n. c * f n :: 'a::{field,topological_semigroup_mult}) \<longleftrightarrow> convergent f"
 proof
   assume "convergent (\<lambda>n. c * f n)"
   from assms convergent_mult[OF this convergent_const[of "inverse c"]]
@@ -2163,7 +2177,7 @@
 qed
 
 lemma convergent_mult_const_right_iff:
-  fixes c :: "'a::real_normed_field"
+  fixes c :: "'a::{field,topological_semigroup_mult}"
   assumes "c \<noteq> 0"
   shows "convergent (\<lambda>n. f n * c) \<longleftrightarrow> convergent f"
   using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac)
--- a/src/HOL/List.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/List.thy	Thu May 03 15:07:14 2018 +0200
@@ -6980,10 +6980,6 @@
 
 signature LIST_CODE =
 sig
-  val implode_list: Code_Thingol.iterm -> Code_Thingol.iterm list option
-  val default_list: int * string
-    -> (Code_Printer.fixity -> Code_Thingol.iterm -> Pretty.T)
-    -> Code_Printer.fixity -> Code_Thingol.iterm -> Code_Thingol.iterm -> Pretty.T
   val add_literal_list: string -> theory -> theory
 end;
 
@@ -7002,7 +6998,7 @@
     | _ => NONE
   end;
 
-fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
+fun print_list (target_fxy, target_cons) pr fxy t1 t2 =
   Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy (
     pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1,
     Code_Printer.str target_cons,
@@ -7016,7 +7012,7 @@
        of SOME ts =>
             Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts)
         | NONE =>
-            default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
+            print_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
   in
     Code_Target.set_printings (Code_Symbol.Constant (@{const_name Cons},
       [(target, SOME (Code_Printer.complex_const_syntax (2, pretty)))]))
--- a/src/HOL/Modules.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Modules.thy	Thu May 03 15:07:14 2018 +0200
@@ -164,7 +164,7 @@
     by (intro exI[of _ u] exI[of _ S]) (auto intro: fS)
 qed
 
-lemma span_induct_alt[consumes 1]:
+lemma span_induct_alt [consumes 1, case_names base step, induct set: span]:
   assumes x: "x \<in> span S"
   assumes h0: "h 0" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c *s x + y)"
   shows "h x"
@@ -210,7 +210,7 @@
     by (intro exI[of _ t] exI[of _ "\<lambda>a. c * r a"]) (auto simp: * scale_sum_right)
 qed
 
-lemma subspace_span: "subspace (span S)"
+lemma subspace_span [iff]: "subspace (span S)"
   by (auto simp: subspace_def span_zero span_add span_scale)
 
 lemma span_neg: "x \<in> span S \<Longrightarrow> - x \<in> span S"
@@ -245,17 +245,13 @@
     by (metis subset_eq)
 qed
 
-lemma (in module) span_induct[consumes 1]:
+lemma (in module) span_induct[consumes 1, case_names base step, induct set: span]:
   assumes x: "x \<in> span S"
     and P: "subspace (Collect P)"
     and SP: "\<And>x. x \<in> S \<Longrightarrow> P x"
   shows "P x"
   using P SP span_subspace_induct x by fastforce
 
-lemma (in module) span_induct':
-  "\<forall>x \<in> S. P x \<Longrightarrow> subspace {x. P x} \<Longrightarrow> \<forall>x \<in> span S. P x"
-  unfolding span_def by (rule hull_induct) auto
-
 lemma span_empty[simp]: "span {} = {0}"
   by (rule span_unique) (auto simp add: subspace_def)
 
--- a/src/HOL/Option.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Option.thy	Thu May 03 15:07:14 2018 +0200
@@ -311,6 +311,9 @@
 
 end
 
+lemma finite_range_Some: "finite (range (Some :: 'a \<Rightarrow> 'a option)) = finite (UNIV :: 'a set)"
+  by (auto dest: finite_imageD intro: inj_Some)
+
 
 subsection \<open>Transfer rules for the Transfer package\<close>
 
--- a/src/HOL/Parity.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Parity.thy	Thu May 03 15:07:14 2018 +0200
@@ -407,6 +407,18 @@
     by (simp add: div_mult2_eq' mult_commute)
 qed
 
+lemma div_mult2_numeral_eq:
+  "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B")
+proof -
+  have "?A = a div of_nat (numeral k) div of_nat (numeral l)"
+    by simp
+  also have "\<dots> = a div (of_nat (numeral k) * of_nat (numeral l))"
+    by (fact div_mult2_eq' [symmetric])
+  also have "\<dots> = ?B"
+    by simp
+  finally show ?thesis .
+qed
+
 end
 
 class ring_parity = ring + semiring_parity
@@ -689,10 +701,10 @@
   where push_bit_eq_mult: "push_bit n a = a * 2 ^ n"
  
 definition take_bit :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
-  where take_bit_eq_mod: "take_bit n a = a mod of_nat (2 ^ n)"
+  where take_bit_eq_mod: "take_bit n a = a mod 2 ^ n"
 
 definition drop_bit :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
-  where drop_bit_eq_div: "drop_bit n a = a div of_nat (2 ^ n)"
+  where drop_bit_eq_div: "drop_bit n a = a div 2 ^ n"
 
 lemma bit_ident:
   "push_bit n (drop_bit n a) + take_bit n a = a"
@@ -807,6 +819,10 @@
   "push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))"
   by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric]) (simp add: ac_simps)
 
+lemma push_bit_of_nat:
+  "push_bit n (of_nat m) = of_nat (push_bit n m)"
+  by (simp add: push_bit_eq_mult Parity.push_bit_eq_mult)
+
 lemma take_bit_0 [simp]:
   "take_bit 0 a = 0"
   by (simp add: take_bit_eq_mod)
@@ -858,6 +874,10 @@
   by (simp only: numeral_eq_Suc power_Suc numeral_Bit1 [of k] mult_2 [symmetric] take_bit_Suc
     ac_simps even_add even_mult_iff div_mult_self1 [OF numeral_neq_zero]) (simp add: ac_simps)
 
+lemma take_bit_of_nat:
+  "take_bit n (of_nat m) = of_nat (take_bit n m)"
+  by (simp add: take_bit_eq_mod Parity.take_bit_eq_mod of_nat_mod [of m "2 ^ n"])
+
 lemma drop_bit_0 [simp]:
   "drop_bit 0 = id"
   by (simp add: fun_eq_iff drop_bit_eq_div)
@@ -907,6 +927,10 @@
   by (simp only: numeral_eq_Suc power_Suc numeral_Bit1 [of k] mult_2 [symmetric] drop_bit_Suc
     div_mult_self4 [OF numeral_neq_zero]) simp
 
+lemma drop_bit_of_nat:
+  "drop_bit n (of_nat m) = of_nat (drop_bit n m)"
+	by (simp add: drop_bit_eq_div Parity.drop_bit_eq_div of_nat_div [of m "2 ^ n"])
+
 end
 
 lemma push_bit_of_Suc_0 [simp]:
--- a/src/HOL/Quickcheck_Exhaustive.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Quickcheck_Exhaustive.thy	Thu May 03 15:07:14 2018 +0200
@@ -714,24 +714,6 @@
 
 ML_file "Tools/Quickcheck/abstract_generators.ML"
 
-instantiation char :: full_exhaustive
-begin
-
-definition full_exhaustive_char
-where
-  "full_exhaustive_char f i =
-     (if 0 < i then
-      case f (0, \<lambda>_ :: unit. Code_Evaluation.Const (STR ''Groups.zero_class.zero'') (TYPEREP(char))) of
-          Some x \<Rightarrow> Some x
-      | None \<Rightarrow> full_exhaustive_class.full_exhaustive
-          (\<lambda>(num, t). f (char_of_nat (nat_of_num num), \<lambda>_ :: unit. Code_Evaluation.App (Code_Evaluation.Const (STR ''String.Char'') TYPEREP(num \<Rightarrow> char)) (t ())))
-          (min (i - 1) 8) \<comment> \<open>generate at most 8 bits\<close>
-      else None)"
-
-instance ..
-
-end
-
 hide_fact (open) orelse_def
 no_notation orelse  (infixr "orelse" 55)
 
--- a/src/HOL/ROOT	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/ROOT	Thu May 03 15:07:14 2018 +0200
@@ -39,7 +39,6 @@
     (*data refinements and dependent applications*)
     AList_Mapping
     Code_Binary_Nat
-    Code_Char
     Code_Prolog
     Code_Real_Approx_By_Float
     Code_Target_Numeral
@@ -248,7 +247,6 @@
     Generate_Binary_Nat
     Generate_Target_Nat
     Generate_Efficient_Datastructures
-    Generate_Pretty_Char
     Code_Test_PolyML
     Code_Test_Scala
   theories [condition = ISABELLE_GHC]
@@ -583,6 +581,7 @@
     Records
     Reflection_Examples
     Refute_Examples
+    Residue_Ring
     Rewrite_Examples
     SOS
     SOS_Cert
--- a/src/HOL/Series.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Series.thy	Thu May 03 15:07:14 2018 +0200
@@ -170,6 +170,9 @@
 lemma summable_LIMSEQ: "summable f \<Longrightarrow> (\<lambda>n. \<Sum>i<n. f i) \<longlonglongrightarrow> suminf f"
   by (rule summable_sums [unfolded sums_def])
 
+lemma summable_LIMSEQ': "summable f \<Longrightarrow> (\<lambda>n. \<Sum>i\<le>n. f i) \<longlonglongrightarrow> suminf f"
+  using sums_def_le by blast
+
 lemma sums_unique: "f sums s \<Longrightarrow> s = suminf f"
   by (metis limI suminf_eq_lim sums_def)
 
--- a/src/HOL/Set_Interval.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Set_Interval.thy	Thu May 03 15:07:14 2018 +0200
@@ -2459,6 +2459,17 @@
   finally show ?thesis .
 qed
 
+lemma prod_nat_group: "(\<Prod>m<n::nat. prod f {m * k ..< m*k + k}) = prod f {..< n * k}"
+proof (cases "k = 0")
+  case True
+  then show ?thesis
+    by auto
+next
+  case False
+  then show ?thesis 
+    by (induct "n"; simp add: prod.atLeastLessThan_concat algebra_simps atLeast0_lessThan_Suc atLeast0LessThan[symmetric])
+qed
+
 
 subsection \<open>Efficient folding over intervals\<close>
 
--- a/src/HOL/Statespace/state_fun.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Statespace/state_fun.ML	Thu May 03 15:07:14 2018 +0200
@@ -76,7 +76,7 @@
 
 fun string_eq_simp_tac ctxt =
   simp_tac (put_simpset HOL_basic_ss ctxt
-    addsimps @{thms list.inject list.distinct Char_eq_Char_iff
+    addsimps @{thms list.inject list.distinct char.inject
       cong_exp_iff_simps simp_thms}
     addsimprocs [lazy_conj_simproc]
     |> Simplifier.add_cong @{thm block_conj_cong});
@@ -85,7 +85,7 @@
 
 val lookup_ss =
   simpset_of (put_simpset HOL_basic_ss @{context}
-    addsimps (@{thms list.inject} @ @{thms Char_eq_Char_iff}
+    addsimps (@{thms list.inject} @ @{thms char.inject}
       @ @{thms list.distinct} @ @{thms simp_thms}
       @ [@{thm StateFun.lookup_update_id_same}, @{thm StateFun.id_id_cancel},
         @{thm StateFun.lookup_update_same}, @{thm StateFun.lookup_update_other}])
@@ -161,7 +161,7 @@
 val meta_ext = @{thm StateFun.meta_ext};
 val ss' =
   simpset_of (put_simpset HOL_ss @{context} addsimps
-    (@{thm StateFun.update_apply} :: @{thm Fun.o_apply} :: @{thms list.inject} @ @{thms Char_eq_Char_iff}
+    (@{thm StateFun.update_apply} :: @{thm Fun.o_apply} :: @{thms list.inject} @ @{thms char.inject}
       @ @{thms list.distinct})
     addsimprocs [lazy_conj_simproc, StateSpace.distinct_simproc]
     |> fold Simplifier.add_cong @{thms block_conj_cong});
--- a/src/HOL/String.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/String.thy	Thu May 03 15:07:14 2018 +0200
@@ -6,191 +6,190 @@
 imports Enum
 begin
 
-subsection \<open>Characters and strings\<close>
+subsection \<open>Strings as list of bytes\<close>
+
+text \<open>
+  When modelling strings, we follow the approach given
+  in @{url "http://utf8everywhere.org/"}:
+
+  \<^item> Strings are a list of bytes (8 bit).
+
+  \<^item> Byte values from 0 to 127 are US-ASCII.
 
-subsubsection \<open>Characters as finite algebraic type\<close>
+  \<^item> Byte values from 128 to 255 are uninterpreted blobs.
+\<close>
+
+subsubsection \<open>Bytes as datatype\<close>
+
+datatype char =
+  Char (digit0: bool) (digit1: bool) (digit2: bool) (digit3: bool)
+       (digit4: bool) (digit5: bool) (digit6: bool) (digit7: bool)
+
+context comm_semiring_1
+begin
 
-typedef char = "{n::nat. n < 256}"
-  morphisms nat_of_char Abs_char
-proof
-  show "Suc 0 \<in> {n. n < 256}" by simp
+definition of_char :: "char \<Rightarrow> 'a"
+  where "of_char c = ((((((of_bool (digit7 c) * 2
+    + of_bool (digit6 c)) * 2
+    + of_bool (digit5 c)) * 2
+    + of_bool (digit4 c)) * 2
+    + of_bool (digit3 c)) * 2
+    + of_bool (digit2 c)) * 2
+    + of_bool (digit1 c)) * 2
+    + of_bool (digit0 c)"
+
+lemma of_char_Char [simp]:
+  "of_char (Char b0 b1 b2 b3 b4 b5 b6 b7) =
+    foldr (\<lambda>b k. of_bool b + k * 2) [b0, b1, b2, b3, b4, b5, b6, b7] 0"
+  by (simp add: of_char_def ac_simps)
+
+end
+
+context semiring_parity
+begin
+
+definition char_of :: "'a \<Rightarrow> char"
+  where "char_of n = Char (odd n) (odd (drop_bit 1 n))
+    (odd (drop_bit 2 n)) (odd (drop_bit 3 n))
+    (odd (drop_bit 4 n)) (odd (drop_bit 5 n))
+    (odd (drop_bit 6 n)) (odd (drop_bit 7 n))"
+
+lemma char_of_char [simp]:
+  "char_of (of_char c) = c"
+proof (cases c)
+  have **: "drop_bit n (q * 2 + of_bool d) = drop_bit (n - 1) q + drop_bit n (of_bool d)"
+    if "n > 0" for q :: 'a and n :: nat and d :: bool
+    using that by (cases n) simp_all
+  case (Char d0 d1 d2 d3 d4 d5 d6 d7)
+  then show ?thesis
+    by (simp only: of_char_def char_of_def char.simps char.sel drop_bit_of_bool **) simp
 qed
 
-setup_lifting type_definition_char
-
-definition char_of_nat :: "nat \<Rightarrow> char"
-where
-  "char_of_nat n = Abs_char (n mod 256)"
+lemma char_of_comp_of_char [simp]:
+  "char_of \<circ> of_char = id"
+  by (simp add: fun_eq_iff)
 
-lemma char_cases [case_names char_of_nat, cases type: char]:
-  "(\<And>n. c = char_of_nat n \<Longrightarrow> n < 256 \<Longrightarrow> P) \<Longrightarrow> P"
-  by (cases c) (simp add: char_of_nat_def)
-
-lemma char_of_nat_of_char [simp]:
-  "char_of_nat (nat_of_char c) = c"
-  by (cases c) (simp add: char_of_nat_def Abs_char_inject Abs_char_inverse)
-
-lemma inj_nat_of_char:
-  "inj nat_of_char"
+lemma inj_of_char:
+  "inj of_char"
 proof (rule injI)
   fix c d
-  assume "nat_of_char c = nat_of_char d"
-  then have "char_of_nat (nat_of_char c) = char_of_nat (nat_of_char d)"
+  assume "of_char c = of_char d"
+  then have "char_of (of_char c) = char_of (of_char d)"
     by simp
   then show "c = d"
     by simp
 qed
   
-lemma nat_of_char_eq_iff [simp]:
-  "nat_of_char c = nat_of_char d \<longleftrightarrow> c = d"
-  by (fact nat_of_char_inject)
+lemma of_char_eq_iff [simp]:
+  "of_char c = of_char d \<longleftrightarrow> c = d"
+  by (simp add: inj_eq inj_of_char)
 
-lemma nat_of_char_of_nat [simp]:
-  "nat_of_char (char_of_nat n) = n mod 256"
-  by (cases "char_of_nat n") (simp add: char_of_nat_def Abs_char_inject Abs_char_inverse)
+lemma of_char_of [simp]:
+  "of_char (char_of a) = a mod 256"
+proof -
+  have *: "{0::nat..<8} = {0, 1, 2, 3, 4, 5, 6, 7}"
+    by auto
+  have "of_char (char_of (take_bit 8 a)) =
+    (\<Sum>k\<in>{0, 1, 2, 3, 4, 5, 6, 7}. push_bit k (of_bool (odd (drop_bit k a))))"
+    by (simp add: of_char_def char_of_def push_bit_of_1 drop_bit_take_bit)
+  also have "\<dots> = take_bit 8 a"
+    using * take_bit_sum [of 8 a] by simp
+  also have "char_of(take_bit 8 a) = char_of a"
+    by (simp add: char_of_def drop_bit_take_bit)
+  finally show ?thesis
+    by (simp add: take_bit_eq_mod)
+qed
 
-lemma char_of_nat_mod_256 [simp]:
-  "char_of_nat (n mod 256) = char_of_nat n"
-  by (cases "char_of_nat (n mod 256)") (simp add: char_of_nat_def)
+lemma char_of_mod_256 [simp]:
+  "char_of (n mod 256) = char_of n"
+  by (metis char_of_char of_char_of)
+
+lemma of_char_mod_256 [simp]:
+  "of_char c mod 256 = of_char c"
+  by (metis char_of_char of_char_of)
 
-lemma char_of_nat_quasi_inj [simp]:
-  "char_of_nat m = char_of_nat n \<longleftrightarrow> m mod 256 = n mod 256"
-  by (simp add: char_of_nat_def Abs_char_inject)
+lemma char_of_quasi_inj [simp]:
+  "char_of m = char_of n \<longleftrightarrow> m mod 256 = n mod 256"
+  by (metis char_of_mod_256 of_char_of)
+
+lemma char_of_nat_eq_iff:
+  "char_of n = c \<longleftrightarrow> take_bit 8 n = of_char c"
+  by (simp add: take_bit_eq_mod) (use of_char_eq_iff in fastforce)
+
+lemma char_of_nat [simp]:
+  "char_of (of_nat n) = char_of n"
+  by (simp add: char_of_def String.char_of_def drop_bit_of_nat)
+
+end
 
 lemma inj_on_char_of_nat [simp]:
-  "inj_on char_of_nat {..<256}"
+  "inj_on char_of {0::nat..<256}"
   by (rule inj_onI) simp
 
-lemma nat_of_char_mod_256 [simp]:
-  "nat_of_char c mod 256 = nat_of_char c"
-  by (cases c) simp
-
 lemma nat_of_char_less_256 [simp]:
-  "nat_of_char c < 256"
+  "of_char c < (256 :: nat)"
 proof -
-  have "nat_of_char c mod 256 < 256"
+  have "of_char c mod (256 :: nat) < 256"
     by arith
   then show ?thesis by simp
 qed
 
+lemma range_nat_of_char:
+  "range of_char = {0::nat..<256}"
+proof (rule; rule)
+  fix n :: nat
+  assume "n \<in> range of_char"
+  then show "n \<in> {0..<256}"
+    by auto
+next
+  fix n :: nat
+  assume "n \<in> {0..<256}"
+  then have "n = of_char (char_of n)"
+    by simp
+  then show "n \<in> range of_char"
+    by (rule range_eqI)
+qed
+
 lemma UNIV_char_of_nat:
-  "UNIV = char_of_nat ` {..<256}"
+  "UNIV = char_of ` {0::nat..<256}"
 proof -
-  { fix c
-    have "c \<in> char_of_nat ` {..<256}"
-      by (cases c) auto
-  } then show ?thesis by auto
+  have "range (of_char :: char \<Rightarrow> nat) = of_char ` char_of ` {0::nat..<256}"
+    by (auto simp add: range_nat_of_char intro!: image_eqI)
+  with inj_of_char [where ?'a = nat] show ?thesis 
+    by (simp add: inj_image_eq_iff)
 qed
 
 lemma card_UNIV_char:
   "card (UNIV :: char set) = 256"
   by (auto simp add: UNIV_char_of_nat card_image)
 
-lemma range_nat_of_char:
-  "range nat_of_char = {..<256}"
-  by (auto simp add: UNIV_char_of_nat image_image image_def)
-
-
-subsubsection \<open>Character literals as variant of numerals\<close>
-
-instantiation char :: zero
+context
+  includes lifting_syntax integer.lifting natural.lifting
 begin
 
-definition zero_char :: char
-  where "0 = char_of_nat 0"
+lemma [transfer_rule]:
+  "(pcr_integer ===> (=)) (char_of :: int \<Rightarrow> char) (char_of :: integer \<Rightarrow> char)"
+  by (unfold char_of_def [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "((=) ===> pcr_integer) (of_char :: char \<Rightarrow> int) (of_char :: char \<Rightarrow> integer)"
+  by (unfold of_char_def [abs_def]) transfer_prover
 
-instance ..
+lemma [transfer_rule]:
+  "(pcr_natural ===> (=)) (char_of :: nat \<Rightarrow> char) (char_of :: natural \<Rightarrow> char)"
+  by (unfold char_of_def [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "((=) ===> pcr_natural) (of_char :: char \<Rightarrow> nat) (of_char :: char \<Rightarrow> natural)"
+  by (unfold of_char_def [abs_def]) transfer_prover
 
 end
 
-definition Char :: "num \<Rightarrow> char"
-  where "Char k = char_of_nat (numeral k)"
-
-code_datatype "0 :: char" Char
-
-lemma nat_of_char_zero [simp]:
-  "nat_of_char 0 = 0"
-  by (simp add: zero_char_def)
-
-lemma nat_of_char_Char [simp]:
-  "nat_of_char (Char k) = numeral k mod 256"
-  by (simp add: Char_def)
-
-lemma Char_eq_Char_iff:
-  "Char k = Char l \<longleftrightarrow> numeral k mod (256 :: nat) = numeral l mod 256" (is "?P \<longleftrightarrow> ?Q")
-proof -
-  have "?P \<longleftrightarrow> nat_of_char (Char k) = nat_of_char (Char l)"
-    by (rule sym, rule inj_eq) (fact inj_nat_of_char)
-  also have "nat_of_char (Char k) = nat_of_char (Char l) \<longleftrightarrow> ?Q"
-    by (simp only: Char_def nat_of_char_of_nat)
-  finally show ?thesis .
-qed
-
-lemma zero_eq_Char_iff:
-  "0 = Char k \<longleftrightarrow> numeral k mod (256 :: nat) = 0"
-  by (auto simp add: zero_char_def Char_def)
-
-lemma Char_eq_zero_iff:
-  "Char k = 0 \<longleftrightarrow> numeral k mod (256 :: nat) = 0"
-  by (auto simp add: zero_char_def Char_def) 
-
-simproc_setup char_eq
-  ("Char m = Char n" | "Char m = 0" | "0 = Char n") = \<open>
-  let
-    val ss = put_simpset HOL_ss @{context}
-      |> fold Simplifier.add_simp @{thms Char_eq_Char_iff zero_eq_Char_iff Char_eq_zero_iff cong_exp_iff_simps}
-      |> simpset_of 
-  in
-    fn _ => fn ctxt => SOME o Simplifier.rewrite (put_simpset ss ctxt)
-  end
-\<close>
+lifting_update integer.lifting
+lifting_forget integer.lifting
 
-definition integer_of_char :: "char \<Rightarrow> integer"
-where
-  "integer_of_char = integer_of_nat \<circ> nat_of_char"
-
-definition char_of_integer :: "integer \<Rightarrow> char"
-where
-  "char_of_integer = char_of_nat \<circ> nat_of_integer"
-
-lemma integer_of_char_zero [simp, code]:
-  "integer_of_char 0 = 0"
-  by (simp add: integer_of_char_def integer_of_nat_def)
-
-lemma integer_of_char_Char [simp]:
-  "integer_of_char (Char k) = numeral k mod 256"
-  by (simp only: integer_of_char_def integer_of_nat_def comp_apply nat_of_char_Char map_fun_def
-    id_apply zmod_int modulo_integer.abs_eq [symmetric]) simp
-
-lemma integer_of_char_Char_code [code]:
-  "integer_of_char (Char k) = integer_of_num k mod 256"
-  by simp
-  
-lemma nat_of_char_code [code]:
-  "nat_of_char = nat_of_integer \<circ> integer_of_char"
-  by (simp add: integer_of_char_def fun_eq_iff)
-
-lemma char_of_nat_code [code]:
-  "char_of_nat = char_of_integer \<circ> integer_of_nat"
-  by (simp add: char_of_integer_def fun_eq_iff)
-
-instantiation char :: equal
-begin
-
-definition equal_char
-  where "equal_char (c :: char) d \<longleftrightarrow> c = d"
-
-instance
-  by standard (simp add: equal_char_def)
-
-end
-
-lemma equal_char_simps [code]:
-  "HOL.equal (0::char) 0 \<longleftrightarrow> True"
-  "HOL.equal (Char k) (Char l) \<longleftrightarrow> HOL.equal (numeral k mod 256 :: nat) (numeral l mod 256)"
-  "HOL.equal 0 (Char k) \<longleftrightarrow> HOL.equal (numeral k mod 256 :: nat) 0"
-  "HOL.equal (Char k) 0 \<longleftrightarrow> HOL.equal (numeral k mod 256 :: nat) 0"
-  by (simp_all only: equal refl Char_eq_Char_iff zero_eq_Char_iff Char_eq_zero_iff)
+lifting_update natural.lifting
+lifting_forget natural.lifting
 
 syntax
   "_Char" :: "str_position \<Rightarrow> char"    ("CHR _")
@@ -199,7 +198,7 @@
 type_synonym string = "char list"
 
 syntax
-  "_String" :: "str_position => string"    ("_")
+  "_String" :: "str_position \<Rightarrow> string"    ("_")
 
 ML_file "Tools/string_syntax.ML"
 
@@ -207,7 +206,8 @@
 begin
 
 definition
-  "Enum.enum = [0, CHR 0x01, CHR 0x02, CHR 0x03,
+  "Enum.enum = [
+    CHR 0x00, CHR 0x01, CHR 0x02, CHR 0x03,
     CHR 0x04, CHR 0x05, CHR 0x06, CHR 0x07,
     CHR 0x08, CHR 0x09, CHR ''\<newline>'', CHR 0x0B,
     CHR 0x0C, CHR 0x0D, CHR 0x0E, CHR 0x0F,
@@ -279,14 +279,15 @@
   "Enum.enum_ex P \<longleftrightarrow> list_ex P (Enum.enum :: char list)"
 
 lemma enum_char_unfold:
-  "Enum.enum = map char_of_nat [0..<256]"
+  "Enum.enum = map char_of [0..<256]"
 proof -
-  have *: "Suc (Suc 0) = 2" by simp
-  have "map nat_of_char Enum.enum = [0..<256]"
-    by (simp add: enum_char_def upt_conv_Cons_Cons *)
-  then have "map char_of_nat (map nat_of_char Enum.enum) =
-    map char_of_nat [0..<256]" by simp
-  then show ?thesis by (simp add: comp_def)
+  have "map (of_char :: char \<Rightarrow> nat) Enum.enum = [0..<256]"
+    by (simp add: enum_char_def of_char_def upt_conv_Cons_Cons numeral_2_eq_2 [symmetric])
+  then have "map char_of (map (of_char :: char \<Rightarrow> nat) Enum.enum) =
+    map char_of [0..<256]"
+    by simp
+  then show ?thesis
+    by simp
 qed
 
 instance proof
@@ -302,157 +303,413 @@
 
 end
 
+lemma linorder_char:
+  "class.linorder (\<lambda>c d. of_char c \<le> (of_char d :: nat)) (\<lambda>c d. of_char c < (of_char d :: nat))"
+  by standard auto
+
+text \<open>Optimized version for execution\<close>
+
+definition char_of_integer :: "integer \<Rightarrow> char"
+  where [code_abbrev]: "char_of_integer = char_of"
+
+definition integer_of_char :: "char \<Rightarrow> integer"
+  where [code_abbrev]: "integer_of_char = of_char"
+
 lemma char_of_integer_code [code]:
-  "char_of_integer n = Enum.enum ! (nat_of_integer n mod 256)"
-  by (simp add: char_of_integer_def enum_char_unfold)
+  "char_of_integer k = (let
+     (q0, b0) = bit_cut_integer k;
+     (q1, b1) = bit_cut_integer q0;
+     (q2, b2) = bit_cut_integer q1;
+     (q3, b3) = bit_cut_integer q2;
+     (q4, b4) = bit_cut_integer q3;
+     (q5, b5) = bit_cut_integer q4;
+     (q6, b6) = bit_cut_integer q5;
+     (_, b7) = bit_cut_integer q6
+    in Char b0 b1 b2 b3 b4 b5 b6 b7)"
+  by (simp add: bit_cut_integer_def char_of_integer_def char_of_def div_mult2_numeral_eq odd_iff_mod_2_eq_one drop_bit_eq_div)
 
-lifting_update char.lifting
-lifting_forget char.lifting
+lemma integer_of_char_code [code]:
+  "integer_of_char (Char b0 b1 b2 b3 b4 b5 b6 b7) =
+    ((((((of_bool b7 * 2 + of_bool b6) * 2 +
+      of_bool b5) * 2 + of_bool b4) * 2 +
+        of_bool b3) * 2 + of_bool b2) * 2 +
+          of_bool b1) * 2 + of_bool b0"
+  by (simp only: integer_of_char_def of_char_def char.sel)
 
 
-subsection \<open>Strings as dedicated type\<close>
+subsection \<open>Strings as dedicated type for target language code generation\<close>
+
+subsubsection \<open>Logical specification\<close>
+
+context
+begin
+
+qualified definition ascii_of :: "char \<Rightarrow> char"
+  where "ascii_of c = Char (digit0 c) (digit1 c) (digit2 c) (digit3 c) (digit4 c) (digit5 c) (digit6 c) False"
 
-typedef literal = "UNIV :: string set"
-  morphisms explode STR ..
+qualified lemma ascii_of_Char [simp]:
+  "ascii_of (Char b0 b1 b2 b3 b4 b5 b6 b7) = Char b0 b1 b2 b3 b4 b5 b6 False"
+  by (simp add: ascii_of_def)
+
+qualified lemma not_digit7_ascii_of [simp]:
+  "\<not> digit7 (ascii_of c)"
+  by (simp add: ascii_of_def)
+
+qualified lemma ascii_of_idem:
+  "ascii_of c = c" if "\<not> digit7 c"
+  using that by (cases c) simp
 
-setup_lifting type_definition_literal
+qualified lemma char_of_ascii_of [simp]:
+  "of_char (ascii_of c) = take_bit 7 (of_char c :: nat)"
+  by (cases c)
+    (simp add: numeral_3_eq_3 [symmetric] numeral_2_eq_2 [symmetric])
+
+qualified typedef literal = "{cs. \<forall>c\<in>set cs. \<not> digit7 c}"
+  morphisms explode Abs_literal
+proof
+  show "[] \<in> {cs. \<forall>c\<in>set cs. \<not> digit7 c}"
+    by simp
+qed
+
+qualified setup_lifting type_definition_literal
 
-lemma STR_inject' [simp]:
-  "STR s = STR t \<longleftrightarrow> s = t"
+qualified lift_definition implode :: "string \<Rightarrow> literal"
+  is "map ascii_of"
+  by auto
+
+qualified lemma implode_explode_eq [simp]:
+  "String.implode (String.explode s) = s"
+proof transfer
+  fix cs
+  show "map ascii_of cs = cs" if "\<forall>c\<in>set cs. \<not> digit7 c"
+    using that
+      by (induction cs) (simp_all add: ascii_of_idem)
+qed
+
+qualified lemma explode_implode_eq [simp]:
+  "String.explode (String.implode cs) = map ascii_of cs"
   by transfer rule
 
-definition implode :: "string \<Rightarrow> String.literal"
-where
-  [code del]: "implode = STR"
+end
+
+
+subsubsection \<open>Syntactic representation\<close>
+
+text \<open>
+  Logical ground representations for literals are:
+
+  \<^enum> @{text 0} for the empty literal;
 
-instantiation literal :: size
+  \<^enum> @{text "Literal b0 \<dots> b6 s"} for a literal starting with one
+    character and continued by another literal.
+
+  Syntactic representations for literals are:
+
+  \<^enum> Printable text as string prefixed with @{text STR};
+
+  \<^enum> A single ascii value as numerical hexadecimal value prefixed with @{text STR}.
+\<close>
+
+instantiation String.literal :: zero
 begin
 
-definition size_literal :: "literal \<Rightarrow> nat"
-where
-  [code]: "size_literal (s::literal) = 0"
+context
+begin
+
+qualified lift_definition zero_literal :: String.literal
+  is Nil
+  by simp
 
 instance ..
 
 end
 
-instantiation literal :: equal
+end
+
+context
 begin
 
-lift_definition equal_literal :: "literal \<Rightarrow> literal \<Rightarrow> bool" is "(=)" .
+qualified abbreviation (output) empty_literal :: String.literal
+  where "empty_literal \<equiv> 0"
+
+qualified lift_definition Literal :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> String.literal \<Rightarrow> String.literal"
+  is "\<lambda>b0 b1 b2 b3 b4 b5 b6 cs. Char b0 b1 b2 b3 b4 b5 b6 False # cs"
+  by auto
 
-instance by intro_classes (transfer, simp)
+qualified lemma Literal_eq_iff [simp]:
+  "Literal b0 b1 b2 b3 b4 b5 b6 s = Literal c0 c1 c2 c3 c4 c5 c6 t
+     \<longleftrightarrow> (b0 \<longleftrightarrow> c0) \<and> (b1 \<longleftrightarrow> c1) \<and> (b2 \<longleftrightarrow> c2) \<and> (b3 \<longleftrightarrow> c3)
+         \<and> (b4 \<longleftrightarrow> c4) \<and> (b5 \<longleftrightarrow> c5) \<and> (b6 \<longleftrightarrow> c6) \<and> s = t"
+  by transfer simp
+
+qualified lemma empty_neq_Literal [simp]:
+  "empty_literal \<noteq> Literal b0 b1 b2 b3 b4 b5 b6 s"
+  by transfer simp
+
+qualified lemma Literal_neq_empty [simp]:
+  "Literal b0 b1 b2 b3 b4 b5 b6 s \<noteq> empty_literal"
+  by transfer simp
 
 end
 
-declare equal_literal.rep_eq[code]
+code_datatype "0 :: String.literal" String.Literal
+
+syntax
+  "_Literal" :: "str_position \<Rightarrow> String.literal"   ("STR _")
+  "_Ascii" :: "num_const \<Rightarrow> String.literal"        ("STR _")
 
-lemma [code nbe]:
-  fixes s :: "String.literal"
-  shows "HOL.equal s s \<longleftrightarrow> True"
-  by (simp add: equal)
+ML_file "Tools/literal.ML"
+
 
-instantiation literal :: zero
+subsubsection \<open>Operations\<close>
+
+instantiation String.literal :: plus
 begin
 
-lift_definition zero_literal :: "literal"
-  is "[]"
-  .
+context
+begin
+
+qualified lift_definition plus_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> String.literal"
+  is "(@)"
+  by auto
 
 instance ..
 
 end
 
-lemma [code]:
-  "0 = STR ''''"
-  by (fact zero_literal.abs_eq)
+end
 
-instantiation literal :: plus
+instance String.literal :: monoid_add
+  by (standard; transfer) simp_all
+
+instantiation String.literal :: size
 begin
 
-lift_definition plus_literal :: "literal \<Rightarrow> literal \<Rightarrow> literal"
-  is "(@)"
-  .
+context
+  includes literal.lifting
+begin
+
+lift_definition size_literal :: "String.literal \<Rightarrow> nat"
+  is length .
+
+end
 
 instance ..
 
 end
 
-lemma [code]:
-  "s + t = String.implode (String.explode s @ String.explode t)"
-  using plus_literal.abs_eq [of "String.explode s" "String.explode t"]
-  by (simp add: explode_inverse implode_def)
+instantiation String.literal :: equal
+begin
+
+context
+begin
+
+qualified lift_definition equal_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool"
+  is HOL.equal .
+
+instance
+  by (standard; transfer) (simp add: equal)
+
+end
+
+end
+
+instantiation String.literal :: linorder
+begin
 
-instance literal :: monoid_add
-  by standard (transfer; simp)+
+context
+begin
+
+qualified lift_definition less_eq_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool"
+  is "ord.lexordp_eq (\<lambda>c d. of_char c < (of_char d :: nat))"
+  .
+
+qualified lift_definition less_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool"
+  is "ord.lexordp (\<lambda>c d. of_char c < (of_char d :: nat))"
+  .
+
+instance proof -
+  from linorder_char interpret linorder "ord.lexordp_eq (\<lambda>c d. of_char c < (of_char d :: nat))"
+    "ord.lexordp (\<lambda>c d. of_char c < (of_char d :: nat)) :: string \<Rightarrow> string \<Rightarrow> bool"
+    by (rule linorder.lexordp_linorder)
+  show "PROP ?thesis"
+    by (standard; transfer) (simp_all add: less_le_not_le linear)
+qed
+
+end
+
+end
 
-lifting_update literal.lifting
-lifting_forget literal.lifting
+lemma infinite_literal:
+  "infinite (UNIV :: String.literal set)"
+proof -
+  define S where "S = range (\<lambda>n. replicate n CHR ''A'')"
+  have "inj_on String.implode S"
+  proof (rule inj_onI)
+    fix cs ds
+    assume "String.implode cs = String.implode ds"
+    then have "String.explode (String.implode cs) = String.explode (String.implode ds)"
+      by simp
+    moreover assume "cs \<in> S" and "ds \<in> S"
+    ultimately show "cs = ds"
+      by (auto simp add: S_def)
+  qed
+  moreover have "infinite S"
+    by (auto simp add: S_def dest: finite_range_imageI [of _ length])
+  ultimately have "infinite (String.implode ` S)"
+    by (simp add: finite_image_iff)
+  then show ?thesis
+    by (auto intro: finite_subset)
+qed
+
+
+subsubsection \<open>Executable conversions\<close>
+
+context
+begin
+
+qualified lift_definition asciis_of_literal :: "String.literal \<Rightarrow> integer list"
+  is "map of_char"
+  .
+
+qualified lift_definition literal_of_asciis :: "integer list \<Rightarrow> String.literal"
+  is "map (String.ascii_of \<circ> char_of)"
+  by auto
 
-  
-subsection \<open>Dedicated conversion for generated computations\<close>
+qualified lemma literal_of_asciis_of_literal [simp]:
+  "literal_of_asciis (asciis_of_literal s) = s"
+proof transfer
+  fix cs
+  assume "\<forall>c\<in>set cs. \<not> digit7 c"
+  then show "map (String.ascii_of \<circ> char_of) (map of_char cs) = cs"
+    by (induction cs) (simp_all add: String.ascii_of_idem) 
+qed
+
+qualified lemma explode_code [code]:
+  "String.explode s = map char_of (asciis_of_literal s)"
+  by transfer simp
+
+qualified lemma implode_code [code]:
+  "String.implode cs = literal_of_asciis (map of_char cs)"
+  by transfer simp
 
-definition char_of_num :: "num \<Rightarrow> char"
-  where "char_of_num = char_of_nat \<circ> nat_of_num"
+end
+
+declare [[code drop: String.literal_of_asciis String.asciis_of_literal]]
+
+
+subsubsection \<open>Technical code generation setup\<close>
+
+text \<open>Alternative constructor for generated computations\<close>
+
+context
+begin  
+
+qualified definition Literal' :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> String.literal \<Rightarrow> String.literal"
+  where [simp]: "Literal' = String.Literal"
+
+lemma [code]:
+  "Literal' b0 b1 b2 b3 b4 b5 b6 s = String.literal_of_asciis
+    [foldr (\<lambda>b k. of_bool b + k * 2) [b0, b1, b2, b3, b4, b5, b6] 0] + s" 
+  unfolding Literal'_def by transfer (simp add: char_of_def)
 
 lemma [code_computation_unfold]:
-  "Char = char_of_num"
-  by (simp add: fun_eq_iff char_of_num_def nat_of_num_numeral Char_def)
-
+  "String.Literal = Literal'"
+  by simp
 
-subsection \<open>Code generator\<close>
+end
 
-ML_file "Tools/string_code.ML"
-
-code_reserved SML string
-code_reserved OCaml string
+code_reserved SML string Char
+code_reserved OCaml string String Char List
+code_reserved Haskell Prelude
 code_reserved Scala string
 
 code_printing
-  type_constructor literal \<rightharpoonup>
+  type_constructor String.literal \<rightharpoonup>
     (SML) "string"
     and (OCaml) "string"
     and (Haskell) "String"
     and (Scala) "String"
+| constant "STR ''''" \<rightharpoonup>
+    (SML) "\"\""
+    and (OCaml) "\"\""
+    and (Haskell) "\"\""
+    and (Scala) "\"\""
 
 setup \<open>
-  fold String_Code.add_literal_string ["SML", "OCaml", "Haskell", "Scala"]
+  fold Literal.add_code ["SML", "OCaml", "Haskell", "Scala"]
 \<close>
 
 code_printing
-  class_instance literal :: equal \<rightharpoonup>
+  constant "(+) :: String.literal \<Rightarrow> String.literal \<Rightarrow> String.literal" \<rightharpoonup>
+    (SML) infixl 18 "^"
+    and (OCaml) infixr 6 "^"
+    and (Haskell) infixr 5 "++"
+    and (Scala) infixl 7 "+"
+| constant String.literal_of_asciis \<rightharpoonup>
+    (SML) "!(String.implode/ o map (fn k => if 0 <= k andalso k < 128 then (Char.chr o IntInf.toInt) k else raise Fail \"Non-ASCII character in literal\"))"
+    and (OCaml) "!(let ks = _ in let res = Bytes.create (List.length ks) in let rec imp i = function | [] -> res | k :: ks ->
+      let l = Big'_int.int'_of'_big'_int k in if 0 <= l && l < 128 then Bytes.set res i (Char.chr l) else failwith \"Non-ASCII character in literal\"; imp (i + 1) ks in imp 0 ks)"
+    and (Haskell) "map/ (let chr k | (0 <= k && k < 128) = Prelude.toEnum k :: Prelude.Char in chr . Prelude.fromInteger)"
+    and (Scala) "\"\"/ ++/ _.map((k: BigInt) => if (BigInt(0) <= k && k < BigInt(128)) k.charValue else sys.error(\"Non-ASCII character in literal\"))"
+| constant String.asciis_of_literal \<rightharpoonup>
+    (SML) "!(map (fn c => let val k = Char.ord c in if k < 128 then IntInf.fromInt k else raise Fail \"Non-ASCII character in literal\" end) /o String.explode)"
+    and (OCaml) "!(let s = _ in let rec exp i l = if i < 0 then l else exp (i - 1) (let k = Char.code (Bytes.get s i) in
+      if k < 128 then Big'_int.big'_int'_of'_int k :: l else failwith \"Non-ASCII character in literal\") in exp (Bytes.length s - 1) [])"
+    and (Haskell) "map/ (let ord k | (k < 128) = Prelude.toInteger k in ord . (Prelude.fromEnum :: Prelude.Char -> Prelude.Int))"
+    and (Scala) "!(_.toList.map(c => { val k: Int = c.toInt; if (k < 128) BigInt(k) else sys.error(\"Non-ASCII character in literal\") }))"
+| class_instance String.literal :: equal \<rightharpoonup>
     (Haskell) -
-| constant "HOL.equal :: literal \<Rightarrow> literal \<Rightarrow> bool" \<rightharpoonup>
+| constant "HOL.equal :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
     (SML) "!((_ : string) = _)"
     and (OCaml) "!((_ : string) = _)"
     and (Haskell) infix 4 "=="
     and (Scala) infixl 5 "=="
+| constant "(\<le>) :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
+    (SML) "!((_ : string) <= _)"
+    and (OCaml) "!((_ : string) <= _)"
+    \<comment> \<open>Order operations for @{typ String.literal} work in Haskell only
+          if no type class instance needs to be generated, because String = [Char] in Haskell
+          and @{typ "char list"} need not have the same order as @{typ String.literal}.\<close>
+    and (Haskell) infix 4 "<="
+    and (Scala) infixl 4 "<="
+    and (Eval) infixl 6 "<="
+| constant "(<) :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
+    (SML) "!((_ : string) < _)"
+    and (OCaml) "!((_ : string) < _)"
+    and (Haskell) infix 4 "<"
+    and (Scala) infixl 4 "<"
+    and (Eval) infixl 6 "<"
+
+
+subsubsection \<open>Code generation utility\<close>
 
 setup \<open>Sign.map_naming (Name_Space.mandatory_path "Code")\<close>
 
-definition abort :: "literal \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a"
-where [simp, code del]: "abort _ f = f ()"
+definition abort :: "String.literal \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a"
+  where [simp]: "abort _ f = f ()"
 
-lemma abort_cong: "msg = msg' ==> Code.abort msg f = Code.abort msg' f"
-by simp
+declare [[code drop: Code.abort]]
+
+lemma abort_cong:
+  "msg = msg' \<Longrightarrow> Code.abort msg f = Code.abort msg' f"
+  by simp
 
 setup \<open>Sign.map_naming Name_Space.parent_path\<close>
 
 setup \<open>Code_Simp.map_ss (Simplifier.add_cong @{thm Code.abort_cong})\<close>
 
-code_printing constant Code.abort \<rightharpoonup>
+code_printing
+  constant Code.abort \<rightharpoonup>
     (SML) "!(raise/ Fail/ _)"
     and (OCaml) "failwith"
     and (Haskell) "!(error/ ::/ forall a./ String -> (() -> a) -> a)"
-    and (Scala) "!{/ sys.error((_));/  ((_)).apply(())/ }"
-| constant "(+) :: literal \<Rightarrow> literal \<Rightarrow> literal" \<rightharpoonup>
-    (SML) infixl 18 "^"
-    and (OCaml) infixr 6 "@"
-    and (Haskell) infixr 5 "++"
-    and (Scala) infixl 7 "+"
+    and (Scala) "!{/ sys.error((_));/ ((_)).apply(())/ }"
+
 
-hide_type (open) literal
+subsubsection \<open>Finally\<close>
 
-hide_const (open) implode explode
+lifting_update literal.lifting
+lifting_forget literal.lifting
 
 end
--- a/src/HOL/Tools/SMT/z3_proof.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Tools/SMT/z3_proof.ML	Thu May 03 15:07:14 2018 +0200
@@ -216,9 +216,15 @@
   let
     val match = Sign.typ_match (Proof_Context.theory_of ctxt)
 
+    fun objT_of bound =
+      (case Symtab.lookup env bound of
+        SOME objT => objT
+      | NONE => raise Fail ("Replaying the proof trace produced by Z3 failed: " ^
+          "the bound " ^ quote bound ^ " is undeclared; this indicates a bug in Z3"))
+
     val t' = singleton (Variable.polymorphic ctxt) t
     val patTs = map snd (Term.strip_qnt_vars @{const_name Pure.all} t')
-    val objTs = map (the o Symtab.lookup env) bounds
+    val objTs = map objT_of bounds
     val subst = subst_of (fold match (patTs ~~ objTs) Vartab.empty)
   in Same.commit (Term_Subst.map_types_same (substTs_same subst)) t' end
 
--- a/src/HOL/Tools/hologic.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Tools/hologic.ML	Thu May 03 15:07:14 2018 +0200
@@ -98,8 +98,6 @@
   val one_const: term
   val bit0_const: term
   val bit1_const: term
-  val mk_bit: int -> term
-  val dest_bit: term -> int
   val mk_numeral: int -> term
   val dest_numeral: term -> int
   val numeral_const: typ -> term
@@ -519,18 +517,15 @@
 and bit0_const = Const ("Num.num.Bit0", numT --> numT)
 and bit1_const = Const ("Num.num.Bit1", numT --> numT);
 
-fun mk_bit 0 = bit0_const
-  | mk_bit 1 = bit1_const
-  | mk_bit _ = raise TERM ("mk_bit", []);
-
-fun dest_bit (Const ("Num.num.Bit0", _)) = 0
-  | dest_bit (Const ("Num.num.Bit1", _)) = 1
-  | dest_bit t = raise TERM ("dest_bit", [t]);
-
 fun mk_numeral i =
-  let fun mk 1 = one_const
-        | mk i = let val (q, r) = Integer.div_mod i 2 in mk_bit r $ mk q end
-  in if i > 0 then mk i else raise TERM ("mk_numeral: " ^ string_of_int i, [])
+  let
+    fun mk 1 = one_const
+      | mk i =
+          let
+            val (q, r) = Integer.div_mod i 2
+          in (if r = 0 then bit0_const else bit1_const) $ mk q end
+  in
+    if i > 0 then mk i else raise TERM ("mk_numeral: " ^ string_of_int i, [])
   end
 
 fun dest_numeral (Const ("Num.num.One", _)) = 1
@@ -555,7 +550,7 @@
   | dest_number (Const ("Groups.one_class.one", T)) = (T, 1)
   | dest_number (Const ("Num.numeral_class.numeral", Type ("fun", [_, T])) $ t) =
       (T, dest_numeral t)
-  | dest_number (Const ("Groups.uminus_class.uminus", Type ("fun", [_, T])) $ t) =
+  | dest_number (Const ("Groups.uminus_class.uminus", Type ("fun", _)) $ t) =
       apsnd (op ~) (dest_number t)
   | dest_number t = raise TERM ("dest_number", [t]);
 
@@ -594,27 +589,35 @@
   | dest_list t = raise TERM ("dest_list", [t]);
 
 
+(* booleans as bits *)
+
+fun mk_bit b = if b = 0 then @{term False}
+  else if b = 1 then @{term True}
+  else raise TERM ("mk_bit", []);
+
+fun mk_bits len = map mk_bit o Integer.radicify 2 len;
+
+fun dest_bit @{term False} = 0
+  | dest_bit @{term True} = 1
+  | dest_bit _ = raise TERM ("dest_bit", []);
+
+val dest_bits = Integer.eval_radix 2 o map dest_bit;
+
+
 (* char *)
 
 val charT = Type ("String.char", []);
 
-val Char_const = Const ("String.Char", numT --> charT);
-
-fun mk_char 0 = Const ("Groups.zero_class.zero", charT)
-  | mk_char i =
-      if 1 <= i andalso i <= 255 then Char_const $ mk_numeral i
-      else raise TERM ("mk_char", []);
+val Char_const = Const ("String.char.Char", replicate 8 boolT ---> charT);
 
-fun dest_char t =
-  let
-    val (T, n) = case t of
-      Const ("Groups.zero_class.zero", T) => (T, 0)
-    | (Const ("String.Char", Type ("fun", [_, T])) $ t) => (T, dest_numeral t)
-    | _ => raise TERM ("dest_char", [t]);
-  in
-    if T = charT then n
-    else raise TERM ("dest_char", [t])
-  end;
+fun mk_char i =
+  if 0 <= i andalso i <= 255
+  then list_comb (Char_const, mk_bits 8 i)
+  else raise TERM ("mk_char", [])
+
+fun dest_char (Const ("String.char.Char", _) $ b0 $ b1 $ b2 $ b3 $ b4 $ b5 $ b6 $ b7) =
+      dest_bits [b0, b1, b2, b3, b4, b5, b6, b7]
+  | dest_char t = raise TERM ("dest_char", [t]);
 
 
 (* string *)
@@ -629,11 +632,28 @@
 
 val literalT = Type ("String.literal", []);
 
-fun mk_literal s = Const ("String.literal.STR", stringT --> literalT)
-      $ mk_string s;
-fun dest_literal (Const ("String.literal.STR", _) $ t) =
-      dest_string t
-  | dest_literal t = raise TERM ("dest_literal", [t]);
+val Literal_const = Const ("String.Literal", replicate 7 boolT ---> literalT --> literalT);
+
+fun mk_literal s = 
+  let
+    fun mk [] =
+          Const ("Groups.zero_class.zero", literalT)
+      | mk (c :: cs) =
+          list_comb (Literal_const, mk_bits 7 c) $ mk cs;
+    val cs = map ord (raw_explode s);
+  in
+    if forall (fn c => c < 128) cs
+    then mk cs
+    else raise TERM ("mk_literal", [])
+  end;
+
+val dest_literal =
+  let
+    fun dest (Const ("Groups.zero_class.zero", Type ("String.literal", []))) = []
+      | dest (Const ("String.Literal", _) $ b0 $ b1 $ b2 $ b3 $ b4 $ b5 $ b6 $ t) =
+          chr (dest_bits [b0, b1, b2, b3, b4, b5, b6]) :: dest t
+      | dest t = raise TERM ("dest_literal", [t]);
+  in implode o dest end;
 
 
 (* typerep and term *)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/literal.ML	Thu May 03 15:07:14 2018 +0200
@@ -0,0 +1,123 @@
+(*  Title:      HOL/Tools/literal.ML
+    Author:     Florian Haftmann, TU Muenchen
+
+Logical and syntactic operations on literals (see also HOL/Tools/hologic.ML).
+*)
+
+signature LITERAL =
+sig
+  val add_code: string -> theory -> theory
+end
+
+structure Literal: LITERAL =
+struct
+
+datatype character = datatype String_Syntax.character;
+
+fun mk_literal_syntax [] = Syntax.const @{const_syntax String.empty_literal}
+  | mk_literal_syntax (c :: cs) =
+      list_comb (Syntax.const @{const_syntax String.Literal}, String_Syntax.mk_bits_syntax 7 c)
+        $ mk_literal_syntax cs;
+
+val dest_literal_syntax =
+  let
+    fun dest (Const (@{const_syntax String.empty_literal}, _)) = []
+      | dest (Const (@{const_syntax String.Literal}, _) $ b0 $ b1 $ b2 $ b3 $ b4 $ b5 $ b6 $ t) =
+          String_Syntax.classify_character (String_Syntax.dest_bits_syntax [b0, b1, b2, b3, b4, b5, b6]) :: dest t
+      | dest t = raise Match;
+  in dest end;
+
+fun ascii_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
+      c $ ascii_tr [t] $ u
+  | ascii_tr [Const (num, _)] =
+      (mk_literal_syntax o single o (fn n => n mod 128) o #value o Lexicon.read_num) num
+  | ascii_tr ts = raise TERM ("ascii_tr", ts);
+
+fun literal_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
+      c $ literal_tr [t] $ u
+  | literal_tr [Free (str, _)] =
+      (mk_literal_syntax o map ord o String_Syntax.plain_strings_of) str
+  | literal_tr ts = raise TERM ("literal_tr", ts);
+
+fun ascii k = Syntax.const @{syntax_const "_Ascii"}
+  $ Syntax.free (String_Syntax.hex k);
+
+fun literal cs = Syntax.const @{syntax_const "_Literal"}
+  $ Syntax.const (Lexicon.implode_str cs);
+
+fun empty_literal_tr' _ = literal [];
+
+fun literal_tr' [b0, b1, b2, b3, b4, b5, b6, t] =
+      let
+        val cs =
+          dest_literal_syntax (list_comb (Syntax.const @{const_syntax String.Literal}, [b0, b1, b2, b3, b4, b5, b6, t]))
+        fun is_printable (Char _) = true
+          | is_printable (Ord _) = false;
+        fun the_char (Char c) = c;
+        fun the_ascii [Ord i] = i;
+      in
+        if forall is_printable cs
+        then literal (map the_char cs)
+        else if length cs = 1
+        then ascii (the_ascii cs)
+        else raise Match
+      end
+  | literal_tr' _ = raise Match;
+
+open Basic_Code_Thingol;
+
+fun map_partial g f =
+  let
+    fun mapp [] = SOME []
+      | mapp (x :: xs) =
+          (case f x of
+            NONE => NONE
+          | SOME y => 
+            (case mapp xs of
+              NONE => NONE
+            | SOME ys => SOME (y :: ys)))
+  in Option.map g o mapp end;
+
+fun implode_bit (IConst { sym = Code_Symbol.Constant @{const_name False}, ... }) = SOME 0
+  | implode_bit (IConst { sym = Code_Symbol.Constant @{const_name True}, ... }) = SOME 1
+  | implode_bit _ = NONE
+
+fun implode_ascii (b0, b1, b2, b3, b4, b5, b6) =
+  map_partial (chr o Integer.eval_radix 2) implode_bit [b0, b1, b2, b3, b4, b5, b6];
+
+fun implode_literal b0 b1 b2 b3 b4 b5 b6 t =
+  let
+    fun dest_literal (IConst { sym = Code_Symbol.Constant @{const_name String.Literal}, ... }
+          `$ b0 `$ b1 `$ b2 `$ b3 `$ b4 `$ b5 `$ b6 `$ t) = SOME ((b0, b1, b2, b3, b4, b5, b6), t)
+      | dest_literal _ = NONE;
+    val (bss', t') = Code_Thingol.unfoldr dest_literal t;
+    val bss = (b0, b1, b2, b3, b4, b5, b6) :: bss';
+  in
+    case t' of
+      IConst { sym = Code_Symbol.Constant @{const_name String.zero_literal_inst.zero_literal}, ... }
+        => map_partial implode implode_ascii bss
+    | _ => NONE
+  end;
+
+fun add_code target thy =
+  let
+    fun pretty literals _ thm _ _ [(b0, _), (b1, _), (b2, _), (b3, _), (b4, _), (b5, _), (b6, _), (t, _)] =
+      case implode_literal b0 b1 b2 b3 b4 b5 b6 t of
+        SOME s => (Code_Printer.str o Code_Printer.literal_string literals) s
+      | NONE => Code_Printer.eqn_error thy thm "Illegal string literal expression";
+  in
+    thy
+    |> Code_Target.set_printings (Code_Symbol.Constant (@{const_name String.Literal},
+      [(target, SOME (Code_Printer.complex_const_syntax (8, pretty)))]))
+  end;
+
+val _ =
+  Theory.setup
+   (Sign.parse_translation
+     [(@{syntax_const "_Ascii"}, K ascii_tr),
+      (@{syntax_const "_Literal"}, K literal_tr)] #>
+    Sign.print_translation
+     [(@{const_syntax String.Literal}, K literal_tr'),
+      (@{const_syntax String.empty_literal}, K empty_literal_tr')]);
+
+end
--- a/src/HOL/Tools/numeral.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Tools/numeral.ML	Thu May 03 15:07:14 2018 +0200
@@ -1,17 +1,14 @@
 (*  Title:      HOL/Tools/numeral.ML
     Author:     Makarius
 
-Logical operations on numerals (see also HOL/Tools/hologic.ML).
+Logical and syntactic operations on numerals (see also HOL/Tools/hologic.ML).
 *)
 
 signature NUMERAL =
 sig
   val mk_cnumber: ctyp -> int -> cterm
   val mk_number_syntax: int -> term
-  val mk_cnumeral: int -> cterm
-  val mk_num_syntax: int -> term
   val dest_num_syntax: term -> int
-  val dest_num: Code_Thingol.iterm -> int option
   val add_code: string -> (int -> int) -> (Code_Printer.literals -> int -> string) -> string -> theory -> theory
 end;
 
@@ -92,25 +89,23 @@
 
 local open Basic_Code_Thingol in
 
-fun dest_num (IConst { sym = Code_Symbol.Constant @{const_name Num.One}, ... }) = SOME 1
-  | dest_num (IConst { sym = Code_Symbol.Constant @{const_name Num.Bit0}, ... } `$ t) =
-     (case dest_num t of
+fun dest_num_code (IConst { sym = Code_Symbol.Constant @{const_name Num.One}, ... }) = SOME 1
+  | dest_num_code (IConst { sym = Code_Symbol.Constant @{const_name Num.Bit0}, ... } `$ t) =
+     (case dest_num_code t of
         SOME n => SOME (2 * n)
       | _ => NONE)
-  | dest_num (IConst { sym = Code_Symbol.Constant @{const_name Num.Bit1}, ... } `$ t) =
-     (case dest_num t of
+  | dest_num_code (IConst { sym = Code_Symbol.Constant @{const_name Num.Bit1}, ... } `$ t) =
+     (case dest_num_code t of
         SOME n => SOME (2 * n + 1)
       | _ => NONE)
-  | dest_num _ = NONE;
+  | dest_num_code _ = NONE;
 
 fun add_code number_of preproc print target thy =
   let
     fun pretty literals _ thm _ _ [(t, _)] =
-      let
-        val n = case dest_num t of
-          SOME n => n
-        | NONE => Code_Printer.eqn_error thy thm "Illegal numeral expression: illegal term"
-      in (Code_Printer.str o print literals o preproc) n end;
+      case dest_num_code t of
+        SOME n => (Code_Printer.str o print literals o preproc) n
+      | NONE => Code_Printer.eqn_error thy thm "Illegal numeral expression: illegal term";
   in
     thy |> Code_Target.set_printings (Code_Symbol.Constant (number_of,
       [(target, SOME (Code_Printer.complex_const_syntax (1, pretty)))]))
--- a/src/HOL/Tools/string_code.ML	Wed May 02 13:49:38 2018 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,85 +0,0 @@
-(*  Title:      HOL/Tools/string_code.ML
-    Author:     Florian Haftmann, TU Muenchen
-
-Code generation for character and string literals.
-*)
-
-signature STRING_CODE =
-sig
-  val add_literal_list_string: string -> theory -> theory
-  val add_literal_char: string -> theory -> theory
-  val add_literal_string: string -> theory -> theory
-end;
-
-structure String_Code : STRING_CODE =
-struct
-
-open Basic_Code_Thingol;
-
-fun decode_char_nonzero t =
-  case Numeral.dest_num t of
-    SOME n => if 0 < n andalso n < 256 then SOME n else NONE
-  | _ => NONE;
-
-fun decode_char (IConst { sym = Code_Symbol.Constant @{const_name String.zero_char_inst.zero_char}, ... }) =
-     SOME 0
-  | decode_char (IConst { sym = Code_Symbol.Constant @{const_name Char}, ... } `$ t) =
-      decode_char_nonzero t
-  | decode_char _ = NONE
-
-fun implode_string literals ts =
-  let
-    val is = map_filter decode_char ts;
-  in if length ts = length is
-    then (SOME o Code_Printer.str o Code_Printer.literal_string literals o implode o map chr) is
-    else NONE
-  end;
-
-fun add_literal_list_string target =
-  let
-    fun pretty literals pr _ vars fxy [(t1, _), (t2, _)] =
-      case Option.map (cons t1) (List_Code.implode_list t2)
-       of SOME ts => (case implode_string literals ts
-             of SOME p => p
-              | NONE =>
-                  Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts))
-        | NONE =>
-            List_Code.default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
-  in
-    Code_Target.set_printings (Code_Symbol.Constant (@{const_name Cons},
-      [(target, SOME (Code_Printer.complex_const_syntax (2, pretty)))]))
-  end;
-
-fun add_literal_char target thy =
-  let
-    fun pr literals =
-      Code_Printer.str o Code_Printer.literal_char literals o chr;
-    fun pretty_zero literals _ _ _ _ [] =
-      pr literals 0
-    fun pretty_Char literals _ thm _ _ [(t, _)] =
-      case decode_char_nonzero t
-       of SOME i => pr literals i
-        | NONE => Code_Printer.eqn_error thy thm "Illegal character expression";
-  in
-    thy
-    |> Code_Target.set_printings (Code_Symbol.Constant (@{const_name String.zero_char_inst.zero_char},
-      [(target, SOME (Code_Printer.complex_const_syntax (0, pretty_zero)))]))
-    |> Code_Target.set_printings (Code_Symbol.Constant (@{const_name Char},
-      [(target, SOME (Code_Printer.complex_const_syntax (1, pretty_Char)))]))
-  end;
-
-fun add_literal_string target thy =
-  let
-    fun pretty literals _ thm _ _ [(t, _)] =
-      case List_Code.implode_list t
-       of SOME ts => (case implode_string literals ts
-             of SOME p => p
-              | NONE => Code_Printer.eqn_error thy thm "Illegal string literal expression")
-        | NONE => Code_Printer.eqn_error thy thm "Illegal string literal expression";
-  in
-    thy
-    |> Code_Target.set_printings (Code_Symbol.Constant (@{const_name STR},
-      [(target, SOME (Code_Printer.complex_const_syntax (1, pretty)))]))
-  end;
-
-end;
--- a/src/HOL/Tools/string_syntax.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Tools/string_syntax.ML	Thu May 03 15:07:14 2018 +0200
@@ -1,10 +1,19 @@
 (*  Title:      HOL/Tools/string_syntax.ML
     Author:     Makarius
 
-Concrete syntax for chars and strings.
+Concrete syntax for characters and strings.
 *)
 
-structure String_Syntax: sig end =
+signature STRING_SYNTAX = sig
+  val hex: int -> string
+  val mk_bits_syntax: int -> int -> term list
+  val dest_bits_syntax: term list -> int
+  val plain_strings_of: string -> string list
+  datatype character = Char of string | Ord of int
+  val classify_character: int -> character
+end
+
+structure String_Syntax: STRING_SYNTAX =
 struct
 
 (* numeral *)
@@ -22,47 +31,59 @@
 fun hex n = hex_prefix (map hex_digit (radixpand (16, n)));
 
 
+(* booleans as bits *)
+
+fun mk_bit_syntax b =
+  Syntax.const (if b = 1 then @{const_syntax True} else @{const_syntax False});
+
+fun mk_bits_syntax len = map mk_bit_syntax o Integer.radicify 2 len;
+
+fun dest_bit_syntax (Const (@{const_syntax True}, _)) = 1 
+  | dest_bit_syntax (Const (@{const_syntax False}, _)) = 0
+  | dest_bit_syntax _ = raise Match;
+
+val dest_bits_syntax = Integer.eval_radix 2 o map dest_bit_syntax;
+
+
 (* char *)
 
-fun mk_char_syntax n =
-  if n = 0 then Syntax.const @{const_name Groups.zero}
-  else Syntax.const @{const_syntax Char} $ Numeral.mk_num_syntax n;
+fun mk_char_syntax i =
+  list_comb (Syntax.const @{const_syntax Char}, mk_bits_syntax 8 i);
 
 fun mk_char_syntax' c =
   if Symbol.is_ascii c then mk_char_syntax (ord c)
   else if c = "\<newline>" then mk_char_syntax 10
   else error ("Bad character: " ^ quote c);
 
-fun plain_string_of str =
+fun plain_strings_of str =
   map fst (Lexicon.explode_str (str, Position.none));
 
-datatype character = Char of string | Ord of int | Unprintable;
+datatype character = Char of string | Ord of int;
 
 val specials = raw_explode "\\\"`'";
 
-fun dest_char_syntax c =
-  case try Numeral.dest_num_syntax c of
-    SOME n =>
-      if n < 256 then
-        let
-          val s = chr n
-        in
-          if not (member (op =) specials s) andalso Symbol.is_ascii s andalso Symbol.is_printable s
-          then Char s
-          else if s = "\n" then Char "\<newline>"
-          else Ord n
-        end
-      else Unprintable
-  | NONE => Unprintable;
+fun classify_character i =
+  let
+    val c = chr i
+  in
+    if not (member (op =) specials c) andalso Symbol.is_ascii c andalso Symbol.is_printable c
+    then Char c
+    else if c = "\n"
+    then Char "\<newline>"
+    else Ord i
+  end;
+
+fun dest_char_syntax b0 b1 b2 b3 b4 b5 b6 b7 =
+  classify_character (dest_bits_syntax [b0, b1, b2, b3, b4, b5, b6, b7])
 
 fun dest_char_ast (Ast.Appl [Ast.Constant @{syntax_const "_Char"}, Ast.Constant s]) =
-      plain_string_of s
+      plain_strings_of s
   | dest_char_ast _ = raise Match;
 
 fun char_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
       c $ char_tr [t] $ u
   | char_tr [Free (str, _)] =
-      (case plain_string_of str of
+      (case plain_strings_of str of
         [c] => mk_char_syntax' c
       | _ => error ("Single character expected: " ^ str))
   | char_tr ts = raise TERM ("char_tr", ts);
@@ -73,15 +94,12 @@
       (mk_char_syntax o #value o Lexicon.read_num) num
   | char_ord_tr ts = raise TERM ("char_ord_tr", ts);
 
-fun char_tr' [t] = (case dest_char_syntax t of
+fun char_tr' [b1, b2, b3, b4, b5, b6, b7, b8] =
+      (case dest_char_syntax b1 b2 b3 b4 b5 b6 b7 b8 of
         Char s => Syntax.const @{syntax_const "_Char"} $
           Syntax.const (Lexicon.implode_str [s])
-      | Ord n => 
-          if n = 0
-          then Syntax.const @{const_syntax Groups.zero}
-          else Syntax.const @{syntax_const "_Char_ord"} $
-            Syntax.free (hex n)
-      | _ => raise Match)
+      | Ord n => Syntax.const @{syntax_const "_Char_ord"} $
+          Syntax.free (hex n))
   | char_tr' _ = raise Match;
 
 
@@ -98,7 +116,7 @@
 fun string_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
       c $ string_tr [t] $ u
   | string_tr [Free (str, _)] =
-      mk_string_syntax (plain_string_of str)
+      mk_string_syntax (plain_strings_of str)
   | string_tr ts = raise TERM ("char_tr", ts);
 
 fun list_ast_tr' [args] =
@@ -120,4 +138,4 @@
     Sign.print_ast_translation
      [(@{syntax_const "_list"}, K list_ast_tr')]);
 
-end;
+end
--- a/src/HOL/Topological_Spaces.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Topological_Spaces.thy	Thu May 03 15:07:14 2018 +0200
@@ -3105,6 +3105,12 @@
 class open_uniformity = "open" + uniformity +
   assumes open_uniformity:
     "\<And>U. open U \<longleftrightarrow> (\<forall>x\<in>U. eventually (\<lambda>(x', y). x' = x \<longrightarrow> y \<in> U) uniformity)"
+begin
+
+subclass topological_space
+  by standard (force elim: eventually_mono eventually_elim2 simp: split_beta' open_uniformity)+
+
+end
 
 class uniform_space = open_uniformity +
   assumes uniformity_refl: "eventually E uniformity \<Longrightarrow> E (x, x)"
@@ -3114,9 +3120,6 @@
         \<exists>D. eventually D uniformity \<and> (\<forall>x y z. D (x, y) \<longrightarrow> D (y, z) \<longrightarrow> E (x, z))"
 begin
 
-subclass topological_space
-  by standard (force elim: eventually_mono eventually_elim2 simp: split_beta' open_uniformity)+
-
 lemma uniformity_bot: "uniformity \<noteq> bot"
   using uniformity_refl by auto
 
--- a/src/HOL/Vector_Spaces.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/Vector_Spaces.thy	Thu May 03 15:07:14 2018 +0200
@@ -213,7 +213,7 @@
 
 lemma maximal_independent_subset_extend:
   assumes "S \<subseteq> V" "independent S"
-  shows "\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
+  obtains B where "S \<subseteq> B" "B \<subseteq> V" "independent B" "V \<subseteq> span B"
 proof -
   let ?C = "{B. S \<subseteq> B \<and> independent B \<and> B \<subseteq> V}"
   have "\<exists>M\<in>?C. \<forall>X\<in>?C. M \<subseteq> X \<longrightarrow> X = M"
@@ -253,10 +253,11 @@
     with \<open>v \<notin> span B\<close> have False
       by (auto intro: span_base) }
   ultimately show ?thesis
-    by (auto intro!: exI[of _ B])
+    by (meson that)
 qed
 
-lemma maximal_independent_subset: "\<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
+lemma maximal_independent_subset:
+  obtains B where "B \<subseteq> V" "independent B" "V \<subseteq> span B"
   by (metis maximal_independent_subset_extend[of "{}"] empty_subsetI independent_empty)
 
 text \<open>Extends a basis from B to a basis of the entire space.\<close>
@@ -337,107 +338,105 @@
 qed
 
 lemma exchange_lemma:
-  assumes f: "finite t"
-  and i: "independent s"
-  and sp: "s \<subseteq> span t"
-  shows "\<exists>t'. card t' = card t \<and> finite t' \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
+  assumes f: "finite T"
+    and i: "independent S"
+    and sp: "S \<subseteq> span T"
+  shows "\<exists>t'. card t' = card T \<and> finite t' \<and> S \<subseteq> t' \<and> t' \<subseteq> S \<union> T \<and> S \<subseteq> span t'"
   using f i sp
-proof (induct "card (t - s)" arbitrary: s t rule: less_induct)
+proof (induct "card (T - S)" arbitrary: S T rule: less_induct)
   case less
-  note ft = `finite t` and s = `independent s` and sp = `s \<subseteq> span t`
-  let ?P = "\<lambda>t'. card t' = card t \<and> finite t' \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
-  let ?ths = "\<exists>t'. ?P t'"
-
-  {
-    assume st: "t \<subseteq> s"
-    from spanning_subset_independent[OF st s sp] st ft span_mono[OF st]
-    have ?ths by (auto intro: span_base)
-  }
-  moreover
-  {
-    assume st:"\<not> t \<subseteq> s"
-    from st obtain b where b: "b \<in> t" "b \<notin> s"
+  note ft = \<open>finite T\<close> and S = \<open>independent S\<close> and sp = \<open>S \<subseteq> span T\<close>
+  let ?P = "\<lambda>t'. card t' = card T \<and> finite t' \<and> S \<subseteq> t' \<and> t' \<subseteq> S \<union> T \<and> S \<subseteq> span t'"
+  show ?case
+  proof (cases "S \<subseteq> T \<or> T \<subseteq> S")
+    case True
+    then show ?thesis
+    proof
+      assume "S \<subseteq> T" then show ?thesis
+        by (metis ft Un_commute sp sup_ge1)
+    next
+      assume "T \<subseteq> S" then show ?thesis
+        by (metis Un_absorb sp spanning_subset_independent[OF _ S sp] ft)
+    qed
+  next
+    case False
+    then have st: "\<not> S \<subseteq> T" "\<not> T \<subseteq> S"
+      by auto
+    from st(2) obtain b where b: "b \<in> T" "b \<notin> S"
       by blast
-    from b have "t - {b} - s \<subset> t - s"
+    from b have "T - {b} - S \<subset> T - S"
       by blast
-    then have cardlt: "card (t - {b} - s) < card (t - s)"
+    then have cardlt: "card (T - {b} - S) < card (T - S)"
       using ft by (auto intro: psubset_card_mono)
-    from b ft have ct0: "card t \<noteq> 0"
+    from b ft have ct0: "card T \<noteq> 0"
       by auto
-    have ?ths
-    proof cases
-      assume stb: "s \<subseteq> span (t - {b})"
-      from ft have ftb: "finite (t - {b})"
+    show ?thesis
+    proof (cases "S \<subseteq> span (T - {b})")
+      case True
+      from ft have ftb: "finite (T - {b})"
         by auto
-      from less(1)[OF cardlt ftb s stb]
-      obtain u where u: "card u = card (t - {b})" "s \<subseteq> u" "u \<subseteq> s \<union> (t - {b})" "s \<subseteq> span u"
-        and fu: "finite u" by blast
-      let ?w = "insert b u"
-      have th0: "s \<subseteq> insert b u"
-        using u by blast
-      from u(3) b have "u \<subseteq> s \<union> t"
-        by blast
-      then have th1: "insert b u \<subseteq> s \<union> t"
-        using u b by blast
-      have bu: "b \<notin> u"
-        using b u by blast
-      from u(1) ft b have "card u = (card t - 1)"
+      from less(1)[OF cardlt ftb S True]
+      obtain U where U: "card U = card (T - {b})" "S \<subseteq> U" "U \<subseteq> S \<union> (T - {b})" "S \<subseteq> span U"
+        and fu: "finite U" by blast
+      let ?w = "insert b U"
+      have th0: "S \<subseteq> insert b U"
+        using U by blast
+      have th1: "insert b U \<subseteq> S \<union> T"
+        using U b by blast
+      have bu: "b \<notin> U"
+        using b U by blast
+      from U(1) ft b have "card U = (card T - 1)"
         by auto
-      then have th2: "card (insert b u) = card t"
+      then have th2: "card (insert b U) = card T"
         using card_insert_disjoint[OF fu bu] ct0 by auto
-      from u(4) have "s \<subseteq> span u" .
-      also have "\<dots> \<subseteq> span (insert b u)"
+      from U(4) have "S \<subseteq> span U" .
+      also have "\<dots> \<subseteq> span (insert b U)"
         by (rule span_mono) blast
-      finally have th3: "s \<subseteq> span (insert b u)" .
+      finally have th3: "S \<subseteq> span (insert b U)" .
       from th0 th1 th2 th3 fu have th: "?P ?w"
         by blast
       from th show ?thesis by blast
     next
-      assume stb: "\<not> s \<subseteq> span (t - {b})"
-      from stb obtain a where a: "a \<in> s" "a \<notin> span (t - {b})"
+      case False
+      then obtain a where a: "a \<in> S" "a \<notin> span (T - {b})"
         by blast
       have ab: "a \<noteq> b"
         using a b by blast
-      have at: "a \<notin> t"
-        using a ab span_base[of a "t- {b}"] by auto
-      have mlt: "card ((insert a (t - {b})) - s) < card (t - s)"
+      have at: "a \<notin> T"
+        using a ab span_superset[of a "T- {b}"] by auto
+      have mlt: "card ((insert a (T - {b})) - S) < card (T - S)"
         using cardlt ft a b by auto
-      have ft': "finite (insert a (t - {b}))"
+      have ft': "finite (insert a (T - {b}))"
         using ft by auto
-      {
+      have sp': "S \<subseteq> span (insert a (T - {b}))"
+      proof
         fix x
-        assume xs: "x \<in> s"
-        have t: "t \<subseteq> insert b (insert a (t - {b}))"
+        assume xs: "x \<in> S"
+        have T: "T \<subseteq> insert b (insert a (T - {b}))"
           using b by auto
-        have bs: "b \<in> span (insert a (t - {b}))"
-          apply (rule in_span_delete)
-          using a sp unfolding subset_eq
-          apply auto
-          done
-        from xs sp have "x \<in> span t"
+        have bs: "b \<in> span (insert a (T - {b}))"
+          by (rule in_span_delete) (use a sp in auto)
+        from xs sp have "x \<in> span T"
           by blast
-        with span_mono[OF t] have x: "x \<in> span (insert b (insert a (t - {b})))" ..
-        from span_trans[OF bs x] have "x \<in> span (insert a (t - {b}))" .
-      }
-      then have sp': "s \<subseteq> span (insert a (t - {b}))"
-        by blast
-      from less(1)[OF mlt ft' s sp'] obtain u where u:
-        "card u = card (insert a (t - {b}))"
-        "finite u" "s \<subseteq> u" "u \<subseteq> s \<union> insert a (t - {b})"
-        "s \<subseteq> span u" by blast
-      from u a b ft at ct0 have "?P u"
+        with span_mono[OF T] have x: "x \<in> span (insert b (insert a (T - {b})))" ..
+        from span_trans[OF bs x] show "x \<in> span (insert a (T - {b}))" .
+      qed
+      from less(1)[OF mlt ft' S sp'] obtain U where U:
+        "card U = card (insert a (T - {b}))"
+        "finite U" "S \<subseteq> U" "U \<subseteq> S \<union> insert a (T - {b})"
+        "S \<subseteq> span U" by blast
+      from U a b ft at ct0 have "?P U"
         by auto
       then show ?thesis by blast
     qed
-  }
-  ultimately show ?ths by blast
+  qed
 qed
 
 lemma independent_span_bound:
-  assumes f: "finite t"
-    and i: "independent s"
-    and sp: "s \<subseteq> span t"
-  shows "finite s \<and> card s \<le> card t"
+  assumes f: "finite T"
+    and i: "independent S"
+    and sp: "S \<subseteq> span T"
+  shows "finite S \<and> card S \<le> card T"
   by (metis exchange_lemma[OF f i sp] finite_subset card_mono)
 
 lemma independent_explicit_finite_subsets:
@@ -544,7 +543,8 @@
 lemma basis_card_eq_dim: "B \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = dim V"
   using dim_eq_card[of B V] span_mono[of B V] span_minimal[OF _ subspace_span, of V B] by auto
 
-lemma basis_exists: "\<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> card B = dim V"
+lemma basis_exists:
+  obtains B where "B \<subseteq> V" "independent B" "V \<subseteq> span B" "card B = dim V"
   by (meson basis_card_eq_dim empty_subsetI independent_empty maximal_independent_subset_extend)
 
 lemma dim_eq_card_independent: "independent B \<Longrightarrow> dim B = card B"
@@ -651,35 +651,28 @@
     and fx: "f x = 0"
   shows "x = 0"
   using fB ifB fi xsB fx
-proof (induct arbitrary: x rule: finite_induct[OF fB])
-  case 1
+proof (induction B arbitrary: x rule: finite_induct)
+  case empty
   then show ?case by auto
 next
-  case (2 a b x)
+  case (insert a b x)
   have fb: "finite b" using "2.prems" by simp
   have th0: "f ` b \<subseteq> f ` (insert a b)"
-    apply (rule image_mono)
-    apply blast
-    done
-  from vs2.independent_mono[ OF "2.prems"(2) th0]
-  have ifb: "vs2.independent (f ` b)"  .
+    by (simp add: subset_insertI)
+  have ifb: "vs2.independent (f ` b)"
+    using independent_mono insert.prems(1) th0 by blast
   have fib: "inj_on f b"
-    apply (rule subset_inj_on [OF "2.prems"(3)])
-    apply blast
-    done
-  from vs1.span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)]
+    using insert.prems(2) by blast
+  from vs1.span_breakdown[of a "insert a b", simplified, OF insert.prems(3)]
   obtain k where k: "x - k *a a \<in> vs1.span (b - {a})"
     by blast
   have "f (x - k *a a) \<in> vs2.span (f ` b)"
     unfolding linear_span_image[OF lf]
-    apply (rule imageI)
-    using k vs1.span_mono[of "b - {a}" b]
-    apply blast
-    done
+    using "insert.hyps"(2) k by auto
   then have "f x - k *b f a \<in> vs2.span (f ` b)"
     by (simp add: linear_diff linear_scale lf)
   then have th: "-k *b f a \<in> vs2.span (f ` b)"
-    using "2.prems"(5) by simp
+    using insert.prems(4) by simp
   have xsb: "x \<in> vs1.span b"
   proof (cases "k = 0")
     case True
@@ -688,19 +681,18 @@
       by blast
   next
     case False
-    with vs2.span_scale[OF th, of "- 1/ k"]
-    have th1: "f a \<in> vs2.span (f ` b)"
-      by auto
-    from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric]
-    have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast
-    from "2.prems"(2) [unfolded vs2.dependent_def bex_simps(8), rule_format, of "f a"]
-    have "f a \<notin> vs2.span (f ` b)" using tha
-      using "2.hyps"(2)
-      "2.prems"(3) by auto
-    with th1 have False by blast
+    from inj_on_image_set_diff[OF insert.prems(2), of "insert a b " "{a}", symmetric]
+    have "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast
+    then have "f a \<notin> span (f ` b)" 
+      using dependent_def insert.hyps(2) insert.prems(1) by fastforce
+    moreover have "f a \<in> span (f ` b)"
+      using False span_mul[OF th, of "- 1/ k"] by auto
+    ultimately have False
+      by blast
     then show ?thesis by blast
   qed
-  from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)] show "x = 0" .
+  show "x = 0"
+    using ifb fib xsb insert.IH insert.prems(4) by blast
 qed
 
 lemma linear_eq_on:
@@ -943,7 +935,7 @@
   by (auto simp: linear_iff_module_hom)
 
 lemma linear_injective_left_inverse: "linear s1 s2 f \<Longrightarrow> inj f \<Longrightarrow> \<exists>g. linear s2 s1 g \<and> g \<circ> f = id"
-  using linear_inj_on_left_inverse[of f UNIV] by (auto simp: fun_eq_iff vs1.span_UNIV)
+  using linear_inj_on_left_inverse[of f UNIV] by (auto simp: fun_eq_iff)
 
 lemma linear_surj_right_inverse:
   assumes lf: "linear s1 s2 f"
@@ -979,15 +971,13 @@
       assume xy: "\<not> ?thesis"
       have th: "card S \<le> card (f ` (S - {y}))"
         unfolding c
-        apply (rule card_mono)
-        apply (rule finite_imageI)
-        using fS apply simp
-        using h xy x y f unfolding subset_eq image_iff
-        apply auto
-        apply (case_tac "xa = f x")
-        apply (rule bexI[where x=x])
-        apply auto
-        done
+      proof (rule card_mono)
+        show "finite (f ` (S - {y}))"
+          by (simp add: fS)
+        show "T \<subseteq> f ` (S - {y})"
+          using h xy x y f unfolding subset_eq image_iff
+          by (metis member_remove remove_def)
+      qed
       also have " \<dots> \<le> card (S - {y})"
         apply (rule card_image_le)
         using fS by simp
@@ -1000,10 +990,7 @@
 next
   assume h: ?rhs
   have "f ` S = T"
-    apply (rule card_subset_eq[OF fT ST])
-    unfolding card_image[OF h]
-    apply (rule c)
-    done
+    by (simp add: ST c card_image card_subset_eq fT h)
   then show ?lhs by blast
 qed
 
@@ -1087,10 +1074,9 @@
 qed
 
 lemma basis_subspace_exists:
-   "subspace S
-        \<Longrightarrow> \<exists>b. finite b \<and> b \<subseteq> S \<and>
-                independent b \<and> span b = S \<and> card b = dim S"
-  by (meson basis_exists finiteI_independent span_subspace)
+  assumes "subspace S"
+  obtains B where "finite B" "B \<subseteq> S" "independent B" "span B = S" "card B = dim S"
+by (metis assms span_subspace basis_exists independent_imp_finite)
 
 lemma dim_mono: assumes "V \<subseteq> span W" shows "dim V \<le> dim W"
 proof -
@@ -1190,7 +1176,7 @@
   shows "S = T"
 proof -
   obtain B where B: "B \<le> S" "independent B \<and> S \<subseteq> span B" "card B = dim S"
-    using basis_exists[of S] by auto
+    using basis_exists[of S] by metis
   then have "span B \<subseteq> S"
     using span_mono[of B S] span_eq_iff[of S] assms by metis
   then have "span B = S"
@@ -1242,7 +1228,7 @@
   obtain B where B: "B \<subseteq> S \<inter> T" "S \<inter> T \<subseteq> span B"
              and indB: "independent B"
              and cardB: "card B = dim (S \<inter> T)"
-    using basis_exists by blast
+    using basis_exists by metis
   then obtain C D where "B \<subseteq> C" "C \<subseteq> S" "independent C" "S \<subseteq> span C"
                     and "B \<subseteq> D" "D \<subseteq> T" "independent D" "T \<subseteq> span D"
     using maximal_independent_subset_extend
@@ -1575,8 +1561,7 @@
 context finite_dimensional_vector_space begin
 
 lemma linear_surj_imp_inj:
-  assumes lf: "linear scale scale f"
-    and sf: "surj f"
+  assumes lf: "linear scale scale f" and sf: "surj f"
   shows "inj f"
 proof -
   interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales
@@ -1586,46 +1571,29 @@
     by blast
   {
     fix x
-    assume x: "x \<in> span B"
-    assume fx: "f x = 0"
+    assume x: "x \<in> span B" and fx: "f x = 0"
     from B(2) have fB: "finite B"
       using finiteI_independent by auto
+    have Uspan: "UNIV \<subseteq> span (f ` B)"
+      by (simp add: B(3) lf sf spanning_surjective_image)
     have fBi: "independent (f ` B)"
-      apply (rule card_le_dim_spanning[of "f ` B" ?U])
-      apply blast
-      using sf B(3)
-      unfolding linear_span_image[OF lf] surj_def subset_eq image_iff
-      apply blast
-      using fB apply blast
-      unfolding d[symmetric]
-      apply (rule card_image_le)
-      apply (rule fB)
-      done
+    proof (rule card_le_dim_spanning)
+      show "card (f ` B) \<le> dim ?U"
+        using card_image_le d fB by fastforce
+    qed (use fB Uspan in auto)
     have th0: "dim ?U \<le> card (f ` B)"
-      apply (rule span_card_ge_dim)
-      apply blast
-      unfolding linear_span_image[OF lf]
-      apply (rule subset_trans[where B = "f ` UNIV"])
-      using sf unfolding surj_def
-      apply blast
-      apply (rule image_mono)
-      apply (rule B(3))
-      apply (metis finite_imageI fB)
-      done
+      by (rule span_card_ge_dim) (use Uspan fB in auto)
     moreover have "card (f ` B) \<le> card B"
       by (rule card_image_le, rule fB)
     ultimately have th1: "card B = card (f ` B)"
       unfolding d by arith
     have fiB: "inj_on f B"
-      unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric]
-      by blast
+      by (simp add: eq_card_imp_inj_on fB th1)
     from linear_indep_image_lemma[OF lf fB fBi fiB x] fx
     have "x = 0" by blast
   }
   then show ?thesis
-    unfolding linear_inj_on_iff_eq_0[OF lf subspace_UNIV]
-    using B(3)
-    by blast
+    unfolding linear_injective_0[OF lf] using B(3) by blast
 qed
 
 lemma linear_inverse_left:
@@ -1638,9 +1606,7 @@
     assume lf: "linear scale scale f" "linear scale scale f'"
     assume f: "f \<circ> f' = id"
     from f have sf: "surj f"
-      apply (auto simp add: o_def id_def surj_def)
-      apply metis
-      done
+      by (auto simp add: o_def id_def surj_def) metis
     interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales
     from linear_surjective_isomorphism[OF lf(1) sf] lf f
     have "f' \<circ> f = id"
@@ -1656,19 +1622,14 @@
   shows "linear scale scale g"
 proof -
   from gf have fi: "inj f"
-    apply (auto simp add: inj_on_def o_def id_def fun_eq_iff)
-    apply metis
-    done
+    by (auto simp add: inj_on_def o_def id_def fun_eq_iff) metis
   interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales
   from linear_injective_isomorphism[OF lf fi]
-  obtain h :: "'b \<Rightarrow> 'b" where h: "linear scale scale h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x"
+  obtain h :: "'b \<Rightarrow> 'b" where "linear scale scale h" and h: "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x"
     by blast
   have "h = g"
-    apply (rule ext) using gf h(2,3)
-    apply (simp add: o_def id_def fun_eq_iff)
-    apply metis
-    done
-  with h(1) show ?thesis by blast
+    by (metis gf h isomorphism_expand left_right_inverse_eq)
+  with \<open>linear h\<close> show ?thesis by blast
 qed
 
 lemma inj_linear_imp_inv_linear:
--- a/src/HOL/ex/Cartouche_Examples.thy	Wed May 02 13:49:38 2018 +0200
+++ b/src/HOL/ex/Cartouche_Examples.thy	Thu May 03 15:07:14 2018 +0200
@@ -61,7 +61,7 @@
           if Symbol.is_ascii s then ord s
           else if s = "\<newline>" then 10
           else error ("String literal contains illegal symbol: " ^ quote s ^ Position.here pos);
-      in Syntax.const @{const_syntax Char} $ HOLogic.mk_numeral c end;
+      in list_comb (Syntax.const @{const_syntax Char}, String_Syntax.mk_bits_syntax 8 c) end;
 
     fun mk_string [] = Const (@{const_syntax Nil}, @{typ string})
       | mk_string (s :: ss) =
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Residue_Ring.thy	Thu May 03 15:07:14 2018 +0200
@@ -0,0 +1,89 @@
+(*  Author:  Florian Haftmann, TUM
+*)
+
+section \<open>Proof of concept for residue rings over int using type numerals\<close>
+
+theory Residue_Ring
+  imports Main "HOL-Library.Type_Length"
+begin
+
+class len2 = len0 +
+  assumes len_ge_2 [iff]: "2 \<le> LENGTH('a)"
+begin
+
+subclass len
+proof
+  show "0 < LENGTH('a)" using len_ge_2
+    by arith
+qed
+
+lemma len_not_eq_Suc_0 [simp]:
+  "LENGTH('a) \<noteq> Suc 0"
+  using len_ge_2 by arith
+
+end
+
+instance bit0 and bit1 :: (len) len2
+  by standard (simp_all add: Suc_le_eq)
+
+quotient_type (overloaded) 'a residue_ring = int / "\<lambda>k l. k mod int (LENGTH('a)) = l mod int (LENGTH('a::len2))"
+  by (auto intro!: equivpI reflpI sympI transpI)
+
+context semiring_1
+begin
+
+lift_definition repr :: "'b::len2 residue_ring \<Rightarrow> 'a"
+  is "\<lambda>k. of_nat (nat (k mod int (LENGTH('b))))"
+  by simp
+
+end
+
+instantiation residue_ring :: (len2) comm_ring_1
+begin
+
+lift_definition zero_residue_ring :: "'a residue_ring"
+  is 0 .
+
+lift_definition one_residue_ring :: "'a residue_ring"
+  is 1 .
+
+lift_definition plus_residue_ring :: "'a residue_ring \<Rightarrow> 'a residue_ring \<Rightarrow> 'a residue_ring"
+  is plus
+  by (fact mod_add_cong)
+
+lift_definition uminus_residue_ring :: "'a residue_ring \<Rightarrow> 'a residue_ring"
+  is uminus
+  by (fact mod_minus_cong)
+
+lift_definition minus_residue_ring :: "'a residue_ring \<Rightarrow> 'a residue_ring \<Rightarrow> 'a residue_ring"
+  is minus
+  by (fact mod_diff_cong)
+
+lift_definition times_residue_ring :: "'a residue_ring \<Rightarrow> 'a residue_ring \<Rightarrow> 'a residue_ring"
+  is times
+  by (fact mod_mult_cong)
+
+instance
+  by (standard; transfer) (simp_all add: algebra_simps mod_eq_0_iff_dvd)
+
+end
+
+context
+  includes lifting_syntax
+begin
+
+lemma [transfer_rule]:
+  "((\<longleftrightarrow>) ===> pcr_residue_ring) of_bool of_bool"
+  by (unfold of_bool_def [abs_def]; transfer_prover)
+
+lemma [transfer_rule]:
+  "((=) ===> pcr_residue_ring) numeral numeral"
+  by (rule transfer_rule_numeral; transfer_prover)
+
+lemma [transfer_rule]:
+  "((=) ===> pcr_residue_ring) int of_nat"
+  by (rule transfer_rule_of_nat; transfer_prover)
+
+end
+
+end
--- a/src/Pure/Admin/build_log.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/Admin/build_log.scala	Thu May 03 15:07:14 2018 +0200
@@ -675,13 +675,18 @@
       errors = log_file.filter_lines(ERROR_MESSAGE_MARKER).map(Library.decode_lines(_)))
   }
 
-  def compress_errors(errors: List[String]): Option[Bytes] =
+  def compress_errors(errors: List[String], cache: XZ.Cache = XZ.cache()): Option[Bytes] =
     if (errors.isEmpty) None
-    else Some(Bytes(YXML.string_of_body(XML.Encode.list(XML.Encode.string)(errors))).compress())
+    else {
+      Some(Bytes(YXML.string_of_body(XML.Encode.list(XML.Encode.string)(errors))).
+        compress(cache = cache))
+    }
 
-  def uncompress_errors(bytes: Bytes): List[String] =
+  def uncompress_errors(bytes: Bytes, cache: XZ.Cache = XZ.cache()): List[String] =
     if (bytes.isEmpty) Nil
-    else XML.Decode.list(YXML.string_of_body(_))(YXML.parse_body(bytes.uncompress().text))
+    else {
+      XML.Decode.list(YXML.string_of_body(_))(YXML.parse_body(bytes.uncompress(cache = cache).text))
+    }
 
 
 
@@ -877,6 +882,7 @@
   class Store private[Build_Log](options: Options)
   {
     val xml_cache: XML.Cache = new XML.Cache()
+    val xz_cache: XZ.Cache = XZ.make_cache()
 
     def open_database(
       user: String = options.string("build_log_database_user"),
@@ -1011,7 +1017,7 @@
           stmt.double(13) = session.ml_timing.factor
           stmt.long(14) = session.heap_size
           stmt.string(15) = session.status.map(_.toString)
-          stmt.bytes(16) = compress_errors(session.errors)
+          stmt.bytes(16) = compress_errors(session.errors, cache = xz_cache)
           stmt.string(17) = session.sources
           stmt.execute()
         }
@@ -1049,7 +1055,7 @@
       {
         val ml_stats: List[(String, Option[Bytes])] =
           Par_List.map[(String, Session_Entry), (String, Option[Bytes])](
-            { case (a, b) => (a, Properties.compress(b.ml_statistics).proper) },
+            { case (a, b) => (a, Properties.compress(b.ml_statistics, cache = xz_cache).proper) },
             build_info.sessions.iterator.filter(p => p._2.ml_statistics.nonEmpty).toList)
         val entries = if (ml_stats.nonEmpty) ml_stats else List("" -> None)
         for ((session_name, ml_statistics) <- entries) {
@@ -1178,10 +1184,12 @@
                 sources = res.get_string(Data.sources),
                 heap_size = res.get_long(Data.heap_size),
                 status = res.get_string(Data.status).map(Session_Status.withName(_)),
-                errors = uncompress_errors(res.bytes(Data.errors)),
+                errors = uncompress_errors(res.bytes(Data.errors), cache = xz_cache),
                 ml_statistics =
-                  if (ml_statistics)
-                    Properties.uncompress(res.bytes(Data.ml_statistics), Some(xml_cache))
+                  if (ml_statistics) {
+                    Properties.uncompress(
+                      res.bytes(Data.ml_statistics), cache = xz_cache, Some(xml_cache))
+                  }
                   else Nil)
             session_name -> session_entry
           }).toMap
--- a/src/Pure/Admin/build_status.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/Admin/build_status.scala	Thu May 03 15:07:14 2018 +0200
@@ -204,6 +204,7 @@
       data_hosts.getOrElse(data_name, Set.empty)
 
     val store = Build_Log.store(options)
+
     using(store.open_database())(db =>
     {
       for (profile <- profiles.sortBy(_.description)) {
@@ -272,8 +273,10 @@
 
             val ml_stats =
               ML_Statistics(
-                if (ml_statistics)
-                  Properties.uncompress(res.bytes(Build_Log.Data.ml_statistics))
+                if (ml_statistics) {
+                  Properties.uncompress(
+                    res.bytes(Build_Log.Data.ml_statistics), cache = store.xz_cache)
+                }
                 else Nil,
                 domain = ml_statistics_domain,
                 heading = session_name + print_version(isabelle_version, afp_version, chapter))
@@ -300,7 +303,9 @@
                 average_heap = ml_stats.average_heap_size,
                 stored_heap = ML_Statistics.heap_scale(res.long(Build_Log.Data.heap_size)),
                 status = Build_Log.Session_Status.withName(res.string(Build_Log.Data.status)),
-                errors = Build_Log.uncompress_errors(res.bytes(Build_Log.Data.errors)))
+                errors =
+                  Build_Log.uncompress_errors(res.bytes(Build_Log.Data.errors),
+                    cache = store.xz_cache))
 
             val sessions = data_entries.getOrElse(data_name, Map.empty)
             val session =
--- a/src/Pure/Admin/ci_profile.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/Admin/ci_profile.scala	Thu May 03 15:07:14 2018 +0200
@@ -9,7 +9,7 @@
 
 import java.time.{Instant, ZoneId}
 import java.time.format.DateTimeFormatter
-import java.util.{Properties => JProperties}
+import java.util.{Properties => JProperties, Map => JMap}
 
 
 abstract class CI_Profile extends Isabelle_Tool.Body
@@ -88,7 +88,7 @@
     print_section("CONFIGURATION")
     println(Build_Log.Settings.show())
     val props = load_properties()
-    System.getProperties().putAll(props)
+    System.getProperties().asInstanceOf[JMap[AnyRef, AnyRef]].putAll(props)
 
     val options =
       with_documents(Options.init())
--- a/src/Pure/Concurrent/multithreading.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/Concurrent/multithreading.ML	Thu May 03 15:07:14 2018 +0200
@@ -10,6 +10,8 @@
   val max_threads_update: int -> unit
   val enabled: unit -> bool
   val relevant: 'a list -> bool
+  val parallel_proofs: int ref
+  val parallel_proofs_enabled: int -> bool
   val sync_wait: Time.time option -> ConditionVar.conditionVar -> Mutex.mutex -> bool Exn.result
   val trace: int ref
   val tracing: int -> (unit -> string) -> unit
@@ -47,6 +49,14 @@
 end;
 
 
+(* parallel_proofs *)
+
+val parallel_proofs = ref 1;
+
+fun parallel_proofs_enabled n =
+  enabled () andalso ! parallel_proofs >= n;
+
+
 (* synchronous wait *)
 
 fun sync_wait time cond lock =
--- a/src/Pure/General/bytes.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/General/bytes.scala	Thu May 03 15:07:14 2018 +0200
@@ -178,13 +178,13 @@
 
   /* XZ data compression */
 
-  def uncompress(): Bytes =
-    using(new XZInputStream(stream()))(Bytes.read_stream(_, hint = length))
+  def uncompress(cache: XZ.Cache = XZ.cache()): Bytes =
+    using(new XZInputStream(stream(), cache))(Bytes.read_stream(_, hint = length))
 
-  def compress(options: XZ.Options = XZ.options()): Bytes =
+  def compress(options: XZ.Options = XZ.options(), cache: XZ.Cache = XZ.cache()): Bytes =
   {
     val result = new ByteArrayOutputStream(length)
-    using(new XZOutputStream(result, options))(write_stream(_))
+    using(new XZOutputStream(result, options, cache))(write_stream(_))
     new Bytes(result.toByteArray, 0, result.size)
   }
 }
--- a/src/Pure/General/integer.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/General/integer.ML	Thu May 03 15:07:14 2018 +0200
@@ -20,6 +20,8 @@
   val lcm: int -> int -> int
   val gcds: int list -> int
   val lcms: int list -> int
+  val radicify: int -> int -> int -> int list (* base -> number of positions -> value -> coefficients *)
+  val eval_radix: int -> int list -> int (* base -> coefficients -> value *)
 end;
 
 structure Integer : INTEGER =
@@ -64,6 +66,16 @@
 fun lcms [] = 1
   | lcms (x :: xs) = abs (Library.foldl PolyML.IntInf.lcm (x, xs));
 
+fun radicify base len k =
+  let
+    val _ = if base < 2
+      then error ("Bad radix base: " ^ string_of_int base) else ();
+    fun shift i = swap (div_mod i base);
+  in funpow_yield len shift k |> fst end;
+
+fun eval_radix base ks =
+  fold_rev (fn k => fn i => k + i * base) ks 0;
+
 end;
 
 (*slightly faster than Poly/ML 5.7.1 library implementation, notably on 32bit multicore*)
--- a/src/Pure/General/properties.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/General/properties.scala	Thu May 03 15:07:14 2018 +0200
@@ -46,17 +46,25 @@
     }
   }
 
-  def compress(ps: List[T], options: XZ.Options = XZ.options()): Bytes =
+  def compress(ps: List[T],
+    options: XZ.Options = XZ.options(),
+    cache: XZ.Cache = XZ.cache()): Bytes =
   {
     if (ps.isEmpty) Bytes.empty
-    else Bytes(YXML.string_of_body(XML.Encode.list(XML.Encode.properties)(ps))).compress(options)
+    else {
+      Bytes(YXML.string_of_body(XML.Encode.list(XML.Encode.properties)(ps))).
+        compress(options = options, cache = cache)
+    }
   }
 
-  def uncompress(bs: Bytes, xml_cache: Option[XML.Cache] = None): List[T] =
+  def uncompress(bs: Bytes,
+    cache: XZ.Cache = XZ.cache(),
+    xml_cache: Option[XML.Cache] = None): List[T] =
   {
     if (bs.isEmpty) Nil
     else {
-      val ps = XML.Decode.list(XML.Decode.properties)(YXML.parse_body(bs.uncompress().text))
+      val ps =
+        XML.Decode.list(XML.Decode.properties)(YXML.parse_body(bs.uncompress(cache = cache).text))
       xml_cache match {
         case None => ps
         case Some(cache) => ps.map(cache.props(_))
--- a/src/Pure/General/xz.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/General/xz.scala	Thu May 03 15:07:14 2018 +0200
@@ -7,11 +7,13 @@
 package isabelle
 
 
-import org.tukaani.xz.LZMA2Options
+import org.tukaani.xz.{LZMA2Options, ArrayCache, BasicArrayCache}
 
 
 object XZ
 {
+  /* options */
+
   type Options = LZMA2Options
 
   def options(preset: Int = 3): Options =
@@ -20,4 +22,12 @@
     opts.setPreset(preset)
     opts
   }
+
+
+  /* cache */
+
+  type Cache = ArrayCache
+
+  def cache(): ArrayCache = ArrayCache.getDefaultCache()
+  def make_cache(): ArrayCache = new BasicArrayCache
 }
--- a/src/Pure/Isar/locale.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/Isar/locale.ML	Thu May 03 15:07:14 2018 +0200
@@ -420,7 +420,6 @@
 
 fun make_notes kind = map (fn ((b, atts), facts) =>
   if null atts andalso forall (null o #2) facts
-    andalso not (Proofterm.proofs_enabled ()) (* FIXME *)
   then Lazy_Notes (kind, (b, Lazy.value (maps #1 facts)))
   else Notes (kind, [((b, atts), facts)]));
 
--- a/src/Pure/PIDE/document.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/PIDE/document.scala	Thu May 03 15:07:14 2018 +0200
@@ -412,6 +412,19 @@
     def purge(f: Version => Option[Version], versions: Map[Document_ID.Version, Version])
       : Map[Document_ID.Version, Version] =
       (versions /: (for ((id, v) <- versions.iterator; v1 <- f(v)) yield (id, v1)))(_ + _)
+
+    def purge_future(versions: Map[Document_ID.Version, Version], future: Future[Version])
+      : Future[Version] =
+    {
+      if (future.is_finished) {
+        val version = future.join
+        versions.get(version.id) match {
+          case Some(version1) if !(version eq version1) => Future.value(version1)
+          case _ => future
+        }
+      }
+      else future
+    }
   }
 
   final class Version private(
@@ -449,6 +462,14 @@
       version.is_finished
 
     def truncate: Change = new Change(None, Nil, version)
+
+    def purge(versions: Map[Document_ID.Version, Version]): Option[Change] =
+    {
+      val previous1 = previous.map(Version.purge_future(versions, _))
+      val version1 = Version.purge_future(versions, version)
+      if ((previous eq previous1) && (version eq version1)) None
+      else Some(new Change(previous1, rev_edits, version1))
+    }
   }
 
 
@@ -477,6 +498,13 @@
         case _ => None
       }
     }
+
+    def purge(versions: Map[Document_ID.Version, Version]): History =
+    {
+      val undo_list1 = undo_list.map(_.purge(versions))
+      if (undo_list1.forall(_.isEmpty)) this
+      else new History(for ((a, b) <- undo_list1 zip undo_list) yield a.getOrElse(b))
+    }
   }
 
 
@@ -786,13 +814,16 @@
         }
       }
 
+      val versions2 = Version.purge(_.purge_blobs(blobs1_names), versions1)
+
       copy(
-        versions = Version.purge(_.purge_blobs(blobs1_names), versions1),
+        versions = versions2,
         blobs = blobs1,
         commands = commands1,
         execs = execs1,
         commands_redirection = commands_redirection.restrict(commands1.isDefinedAt(_)),
         assignments = assignments1,
+        history = history.purge(versions2),
         removing_versions = false)
     }
 
--- a/src/Pure/System/isabelle_process.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/System/isabelle_process.ML	Thu May 03 15:07:14 2018 +0200
@@ -218,14 +218,14 @@
   Future.ML_statistics := Options.default_bool "ML_statistics";
   Multithreading.trace := Options.default_int "threads_trace";
   Multithreading.max_threads_update (Options.default_int "threads");
-  Goal.parallel_proofs := Options.default_int "parallel_proofs";
+  Multithreading.parallel_proofs := Options.default_int "parallel_proofs";
   let val proofs = Options.default_int "record_proofs"
   in if proofs < 0 then () else Proofterm.proofs := proofs end;
   Printer.show_markup_default := false);
 
 fun init_options_interactive () =
  (init_options ();
-  Goal.parallel_proofs := (if Options.default_int "parallel_proofs" > 0 then 3 else 0);
+  Multithreading.parallel_proofs := (if Options.default_int "parallel_proofs" > 0 then 3 else 0);
   Printer.show_markup_default := true);
 
 end;
--- a/src/Pure/Thy/sessions.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/Thy/sessions.scala	Thu May 03 15:07:14 2018 +0200
@@ -974,6 +974,7 @@
     /* SQL database content */
 
     val xml_cache = new XML.Cache()
+    val xz_cache = XZ.make_cache()
 
     def read_bytes(db: SQL.Database, name: String, column: SQL.Column): Bytes =
       db.using_statement(Session_Info.table.select(List(column),
@@ -984,7 +985,8 @@
       })
 
     def read_properties(db: SQL.Database, name: String, column: SQL.Column): List[Properties.T] =
-      Properties.uncompress(read_bytes(db, name, column), Some(xml_cache))
+      Properties.uncompress(
+        read_bytes(db, name, column), cache = xz_cache, xml_cache = Some(xml_cache))
 
 
     /* output */
@@ -1040,11 +1042,11 @@
         {
           stmt.string(1) = name
           stmt.bytes(2) = Properties.encode(build_log.session_timing)
-          stmt.bytes(3) = Properties.compress(build_log.command_timings)
-          stmt.bytes(4) = Properties.compress(build_log.theory_timings)
-          stmt.bytes(5) = Properties.compress(build_log.ml_statistics)
-          stmt.bytes(6) = Properties.compress(build_log.task_statistics)
-          stmt.bytes(7) = Build_Log.compress_errors(build_log.errors)
+          stmt.bytes(3) = Properties.compress(build_log.command_timings, cache = xz_cache)
+          stmt.bytes(4) = Properties.compress(build_log.theory_timings, cache = xz_cache)
+          stmt.bytes(5) = Properties.compress(build_log.ml_statistics, cache = xz_cache)
+          stmt.bytes(6) = Properties.compress(build_log.task_statistics, cache = xz_cache)
+          stmt.bytes(7) = Build_Log.compress_errors(build_log.errors, cache = xz_cache)
           stmt.string(8) = build.sources
           stmt.string(9) = cat_lines(build.input_heaps)
           stmt.string(10) = build.output_heap getOrElse ""
@@ -1070,7 +1072,7 @@
       read_properties(db, name, Session_Info.task_statistics)
 
     def read_errors(db: SQL.Database, name: String): List[String] =
-      Build_Log.uncompress_errors(read_bytes(db, name, Session_Info.errors))
+      Build_Log.uncompress_errors(read_bytes(db, name, Session_Info.errors), cache = xz_cache)
 
     def read_build(db: SQL.Database, name: String): Option[Build.Session_Info] =
     {
--- a/src/Pure/goal.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/goal.ML	Thu May 03 15:07:14 2018 +0200
@@ -7,7 +7,6 @@
 signature BASIC_GOAL =
 sig
   val quick_and_dirty: bool Config.T
-  val parallel_proofs: int Unsynchronized.ref
   val SELECT_GOAL: tactic -> int -> tactic
   val PREFER_GOAL: tactic -> int -> tactic
   val CONJUNCTS: tactic -> int -> tactic
@@ -112,10 +111,8 @@
     else skip
   end;
 
-val parallel_proofs = Unsynchronized.ref 1;
-
 fun future_enabled n =
-  Multithreading.enabled () andalso ! parallel_proofs >= n andalso
+  Multithreading.parallel_proofs_enabled n andalso
   is_some (Future.worker_task ());
 
 fun future_enabled_timing t =
--- a/src/Pure/more_thm.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/more_thm.ML	Thu May 03 15:07:14 2018 +0200
@@ -603,10 +603,9 @@
 
 (* unofficial theorem names *)
 
+fun has_name_hint thm = AList.defined (op =) (Thm.get_tags thm) Markup.nameN;
 fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN);
-
-val has_name_hint = can the_name_hint;
-val get_name_hint = the_default "??.unknown" o try the_name_hint;
+fun get_name_hint thm = if has_name_hint thm then the_name_hint thm else "??.unknown";
 
 fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name);
 
--- a/src/Pure/proofterm.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/proofterm.ML	Thu May 03 15:07:14 2018 +0200
@@ -1613,11 +1613,15 @@
     val body0 =
       if not (proofs_enabled ()) then Future.value (make_body0 MinProof)
       else
-        (singleton o Future.cond_forks)
-          {name = "Proofterm.prepare_thm_proof", group = NONE,
-            deps = [], pri = 0, interrupts = true}
-          (fn () =>
-            make_body0 (shrink_proof (rew_proof thy (fold_rev implies_intr_proof hyps prf))));
+        let
+          val rew = rew_proof thy;
+          val prf' = fold_rev implies_intr_proof hyps prf;
+        in
+          (singleton o Future.cond_forks)
+            {name = "Proofterm.prepare_thm_proof", group = NONE,
+              deps = [], pri = 1, interrupts = true}
+            (fn () => make_body0 (shrink_proof (rew prf')))
+        end;
 
     fun new_prf () = (serial (), fulfill_proof_future thy promises postproc body0);
     val (i, body') =
--- a/src/Pure/raw_simplifier.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/raw_simplifier.ML	Thu May 03 15:07:14 2018 +0200
@@ -90,6 +90,7 @@
   val prems_of: Proof.context -> thm list
   val add_simp: thm -> Proof.context -> Proof.context
   val del_simp: thm -> Proof.context -> Proof.context
+  val reorient_simp: thm -> Proof.context -> Proof.context
   val init_simpset: thm list -> Proof.context -> Proof.context
   val add_eqcong: thm -> Proof.context -> Proof.context
   val del_eqcong: thm -> Proof.context -> Proof.context
@@ -560,16 +561,21 @@
     else rrule_eq_True ctxt thm name lhs elhs rhs thm
   end;
 
-fun extract_rews ctxt thms =
-  let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt
-  in maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk ctxt thm)) thms end;
+fun extract_rews ctxt sym thms =
+  let
+    val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt;
+    val mk =
+      if sym then fn ctxt => fn th => (mk ctxt th) RL [Drule.symmetric_thm]
+      else mk  
+  in maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk ctxt thm)) thms
+  end;
 
 fun extract_safe_rrules ctxt thm =
-  maps (orient_rrule ctxt) (extract_rews ctxt [thm]);
+  maps (orient_rrule ctxt) (extract_rews ctxt false [thm]);
 
 fun mk_rrules ctxt thms =
   let
-    val rews = extract_rews ctxt thms
+    val rews = extract_rews ctxt false thms
     val raw_rrules = flat (map (mk_rrule ctxt) rews)
   in map mk_rrule2 raw_rrules end
 
@@ -578,20 +584,24 @@
 
 local
 
-fun comb_simps ctxt comb mk_rrule thms =
-  let val rews = extract_rews ctxt (map (Thm.transfer' ctxt) thms);
+fun comb_simps ctxt comb mk_rrule sym thms =
+  let val rews = extract_rews ctxt sym (map (Thm.transfer' ctxt) thms);
   in fold (fold comb o mk_rrule) rews ctxt end;
 
 in
 
 fun ctxt addsimps thms =
-  comb_simps ctxt insert_rrule (mk_rrule ctxt) thms;
+  comb_simps ctxt insert_rrule (mk_rrule ctxt) false thms;
+
+fun addsymsimps ctxt thms =
+  comb_simps ctxt insert_rrule (mk_rrule ctxt) true thms;
 
 fun ctxt delsimps thms =
-  comb_simps ctxt del_rrule (map mk_rrule2 o mk_rrule ctxt) thms;
+  comb_simps ctxt del_rrule (map mk_rrule2 o mk_rrule ctxt) false thms;
 
 fun add_simp thm ctxt = ctxt addsimps [thm];
 fun del_simp thm ctxt = ctxt delsimps [thm];
+fun reorient_simp thm ctxt = addsymsimps (ctxt delsimps [thm]) [thm];
 
 end;
 
--- a/src/Pure/simplifier.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Pure/simplifier.ML	Thu May 03 15:07:14 2018 +0200
@@ -32,6 +32,7 @@
   val attrib: (thm -> Proof.context -> Proof.context) -> attribute
   val simp_add: attribute
   val simp_del: attribute
+  val simp_reorient: attribute
   val cong_add: attribute
   val cong_del: attribute
   val check_simproc: Proof.context -> xstring * Position.T -> string
@@ -89,6 +90,7 @@
 
 val simp_add = attrib add_simp;
 val simp_del = attrib del_simp;
+val simp_reorient = attrib reorient_simp;
 val cong_add = attrib add_cong;
 val cong_del = attrib del_cong;
 
@@ -267,6 +269,7 @@
 (* add / del *)
 
 val simpN = "simp";
+val reorientN = "reorient"
 val congN = "cong";
 val onlyN = "only";
 val no_asmN = "no_asm";
@@ -340,6 +343,7 @@
  [Args.$$$ simpN -- Args.colon >> K (Method.modifier simp_add \<^here>),
   Args.$$$ simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>),
   Args.$$$ simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>),
+  Args.$$$ reorientN -- Args.colon >> K (Method.modifier simp_reorient \<^here>),
   Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >>
     K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = \<^here>}]
    @ cong_modifiers;
@@ -347,6 +351,7 @@
 val simp_modifiers' =
  [Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>),
   Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>),
+  Args.$$$ reorientN -- Args.colon >> K (Method.modifier simp_reorient \<^here>),
   Args.$$$ onlyN -- Args.colon >>
     K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = \<^here>}]
    @ cong_modifiers;
--- a/src/Tools/Adhoc_Overloading.thy	Wed May 02 13:49:38 2018 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,16 +0,0 @@
-(* Author: Alexander Krauss, TU Muenchen
-   Author: Christian Sternagel, University of Innsbruck
-*)
-
-section \<open>Adhoc overloading of constants based on their types\<close>
-
-theory Adhoc_Overloading
-imports Pure
-keywords
-  "adhoc_overloading" "no_adhoc_overloading" :: thy_decl
-begin
-
-ML_file "adhoc_overloading.ML"
-
-end
-
--- a/src/Tools/Code/code_haskell.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Tools/Code/code_haskell.ML	Thu May 03 15:07:14 2018 +0200
@@ -35,6 +35,16 @@
 
 (** Haskell serializer **)
 
+val print_haskell_string =
+  let
+    fun char c =
+      let
+        val _ = if Symbol.is_ascii c then ()
+          else error "non-ASCII byte in Haskell string literal";
+        val s = ML_Syntax.print_char c;
+      in if s = "'" then "\\'" else s end;
+  in quote o translate_string char end;
+
 fun print_haskell_stmt class_syntax tyco_syntax const_syntax
     reserved deresolve deriving_show =
   let
@@ -134,7 +144,7 @@
                 :: replicate n "_"
                 @ "="
                 :: "error"
-                @@ ML_Syntax.print_string const
+                @@ print_haskell_string const
               );
             fun print_eqn ((ts, t), (some_thm, _)) =
               let
@@ -426,18 +436,12 @@
 
 fun print_numeral typ = Library.enclose "(" (" :: " ^ typ ^ ")") o signed_string_of_int;
 
-val literals = let
-  fun char_haskell c =
-    let
-      val s = ML_Syntax.print_char c;
-    in if s = "'" then "\\'" else s end;
-in Literals {
-  literal_char = Library.enclose "'" "'" o char_haskell,
-  literal_string = quote o translate_string char_haskell,
+val literals = Literals {
+  literal_string = print_haskell_string,
   literal_numeral = print_numeral "Integer",
   literal_list = enum "," "[" "]",
   infix_cons = (5, ":")
-} end;
+};
 
 
 (** optional monad syntax **)
--- a/src/Tools/Code/code_ml.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Tools/Code/code_ml.ML	Thu May 03 15:07:14 2018 +0200
@@ -51,10 +51,14 @@
 
 (** SML serializer **)
 
-fun print_char_any_ml s =
-  if Symbol.is_char s then ML_Syntax.print_char s else "\\092" ^ unprefix "\\" s;
+fun print_sml_char c =
+  if c = "\\"
+  then "\\092" (*produce strings suitable for both SML as well as Isabelle/ML*)
+  else if Symbol.is_ascii c
+  then ML_Syntax.print_char c
+  else error "non-ASCII byte in SML string literal";
 
-val print_string_any_ml = quote o implode o map print_char_any_ml o Symbol.explode;
+val print_sml_string = quote o translate_string print_sml_char;
 
 fun print_sml_stmt tyco_syntax const_syntax reserved is_constr deresolve =
   let
@@ -255,7 +259,7 @@
             @ "="
             :: "raise"
             :: "Fail"
-            @@ print_string_any_ml const
+            @@ print_sml_string const
           ))
       | print_stmt _ (ML_Val binding) =
           let
@@ -358,8 +362,7 @@
   );
 
 val literals_sml = Literals {
-  literal_char = prefix "#" o quote o ML_Syntax.print_char,
-  literal_string = print_string_any_ml,
+  literal_string = print_sml_string,
   literal_numeral = fn k => "(" ^ string_of_int k ^ " : IntInf.int)",
   literal_list = enum "," "[" "]",
   infix_cons = (7, "::")
@@ -368,6 +371,23 @@
 
 (** OCaml serializer **)
 
+val print_ocaml_string =
+  let
+    fun chr i =
+      let
+        val xs = string_of_int i;
+        val ys = replicate_string (3 - length (raw_explode xs)) "0";
+      in "\\" ^ ys ^ xs end;
+    fun char c =
+      let
+        val i = ord c;
+        val s =
+          if i >= 128 then error "non-ASCII byte in OCaml string literal"
+          else if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126
+          then chr i else c
+      in s end;
+  in quote o translate_string char end;
+
 fun print_ocaml_stmt tyco_syntax const_syntax reserved is_constr deresolve =
   let
     val deresolve_const = deresolve o Constant;
@@ -600,7 +620,7 @@
             :: replicate n "_"
             @ "="
             :: "failwith"
-            @@ ML_Syntax.print_string const
+            @@ print_ocaml_string const
           ))
       | print_stmt _ (ML_Val binding) =
           let
@@ -696,25 +716,13 @@
   );
 
 val literals_ocaml = let
-  fun chr i =
-    let
-      val xs = string_of_int i;
-      val ys = replicate_string (3 - length (raw_explode xs)) "0";
-    in "\\" ^ ys ^ xs end;
-  fun char_ocaml c =
-    let
-      val i = ord c;
-      val s = if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126
-        then chr i else c
-    in s end;
   fun numeral_ocaml k = if k < 0
     then "(Big_int.minus_big_int " ^ numeral_ocaml (~ k) ^ ")"
     else if k <= 1073741823
       then "(Big_int.big_int_of_int " ^ string_of_int k ^ ")"
       else "(Big_int.big_int_of_string " ^ quote (string_of_int k) ^ ")"
 in Literals {
-  literal_char = Library.enclose "'" "'" o char_ocaml,
-  literal_string = quote o translate_string char_ocaml,
+  literal_string = print_ocaml_string,
   literal_numeral = numeral_ocaml,
   literal_list = enum ";" "[" "]",
   infix_cons = (6, "::")
--- a/src/Tools/Code/code_printer.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Tools/Code/code_printer.ML	Thu May 03 15:07:14 2018 +0200
@@ -39,11 +39,10 @@
   val aux_params: var_ctxt -> iterm list list -> string list
 
   type literals
-  val Literals: { literal_char: string -> string, literal_string: string -> string,
+  val Literals: { literal_string: string -> string,
         literal_numeral: int -> string,
         literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string }
     -> literals
-  val literal_char: literals -> string -> string
   val literal_string: literals -> string -> string
   val literal_numeral: literals -> int -> string
   val literal_list: literals -> Pretty.T list -> Pretty.T
@@ -215,7 +214,6 @@
 (** pretty literals **)
 
 datatype literals = Literals of {
-  literal_char: string -> string,
   literal_string: string -> string,
   literal_numeral: int -> string,
   literal_list: Pretty.T list -> Pretty.T,
@@ -224,7 +222,6 @@
 
 fun dest_Literals (Literals lits) = lits;
 
-val literal_char = #literal_char o dest_Literals;
 val literal_string = #literal_string o dest_Literals;
 val literal_numeral = #literal_numeral o dest_Literals;
 val literal_list = #literal_list o dest_Literals;
--- a/src/Tools/Code/code_scala.ML	Wed May 02 13:49:38 2018 +0200
+++ b/src/Tools/Code/code_scala.ML	Thu May 03 15:07:14 2018 +0200
@@ -23,6 +23,24 @@
 
 val target = "Scala";
 
+val print_scala_string =
+  let
+    fun chr i = "\\u" ^ align_right "0" 4 (Int.fmt StringCvt.HEX i)
+    fun char c = if c = "'" then "\\'"
+      else if c = "\"" then "\\\""
+      else if c = "\\" then "\\\\"
+      else
+        let
+          val i = ord c
+        in
+          if i < 32 orelse i > 126
+          then chr i
+          else if i >= 128
+          then error "non-ASCII byte in Scala string literal"
+          else c
+        end
+  in quote o translate_string char end;
+
 datatype scala_stmt = Fun of typscheme * ((iterm list * iterm) * (thm option * bool)) list
   | Datatype of vname list * ((string * vname list) * itype list) list
   | Class of (vname * ((class * class) list * (string * itype) list))
@@ -189,7 +207,7 @@
             val vars = intro_vars params reserved;
           in
             concat [print_defhead tyvars vars const vs params tys ty',
-              str ("sys.error(\"" ^ const ^ "\")")]
+              str ("sys.error(" ^ print_scala_string const ^ ")")]
           end
       | print_def const (vs, ty) eqs =
           let
@@ -432,19 +450,12 @@
     >> (fn case_insensitive => serialize_scala case_insensitive));
 
 val literals = let
-  fun char_scala c = if c = "'" then "\\'"
-    else if c = "\"" then "\\\""
-    else if c = "\\" then "\\\\"
-    else let val k = ord c
-    in if k < 32 orelse k > 126
-    then "\\u" ^ align_right "0" 4 (Int.fmt StringCvt.HEX k) else c end
   fun numeral_scala k =
     if ~2147483647 < k andalso k <= 2147483647
     then signed_string_of_int k
     else quote (signed_string_of_int k)
 in Literals {
-  literal_char = Library.enclose "'" "'" o char_scala,
-  literal_string = quote o translate_string char_scala,
+  literal_string = print_scala_string,
   literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")",
   literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps],
   infix_cons = (6, "::")
--- a/src/Tools/adhoc_overloading.ML	Wed May 02 13:49:38 2018 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,246 +0,0 @@
-(*  Author:     Alexander Krauss, TU Muenchen
-    Author:     Christian Sternagel, University of Innsbruck
-
-Adhoc overloading of constants based on their types.
-*)
-
-signature ADHOC_OVERLOADING =
-sig
-  val is_overloaded: Proof.context -> string -> bool
-  val generic_add_overloaded: string -> Context.generic -> Context.generic
-  val generic_remove_overloaded: string -> Context.generic -> Context.generic
-  val generic_add_variant: string -> term -> Context.generic -> Context.generic
-  (*If the list of variants is empty at the end of "generic_remove_variant", then
-  "generic_remove_overloaded" is called implicitly.*)
-  val generic_remove_variant: string -> term -> Context.generic -> Context.generic
-  val show_variants: bool Config.T
-end
-
-structure Adhoc_Overloading: ADHOC_OVERLOADING =
-struct
-
-val show_variants = Attrib.setup_config_bool @{binding show_variants} (K false);
-
-
-(* errors *)
-
-fun err_duplicate_variant oconst =
-  error ("Duplicate variant of " ^ quote oconst);
-
-fun err_not_a_variant oconst =
-  error ("Not a variant of " ^  quote oconst);
-
-fun err_not_overloaded oconst =
-  error ("Constant " ^ quote oconst ^ " is not declared as overloaded");
-
-fun err_unresolved_overloading ctxt0 (c, T) t instances =
-  let
-    val ctxt = Config.put show_variants true ctxt0
-    val const_space = Proof_Context.const_space ctxt
-    val prt_const =
-      Pretty.block [Name_Space.pretty ctxt const_space c, Pretty.str " ::", Pretty.brk 1,
-        Pretty.quote (Syntax.pretty_typ ctxt T)]
-  in
-    error (Pretty.string_of (Pretty.chunks [
-      Pretty.block [
-        Pretty.str "Unresolved adhoc overloading of constant", Pretty.brk 1,
-        prt_const, Pretty.brk 1, Pretty.str "in term", Pretty.brk 1,
-        Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t)]],
-      Pretty.block (
-        (if null instances then [Pretty.str "no instances"]
-        else Pretty.fbreaks (
-          Pretty.str "multiple instances:" ::
-          map (Pretty.mark Markup.item o Syntax.pretty_term ctxt) instances)))]))
-  end;
-
-
-(* generic data *)
-
-fun variants_eq ((v1, T1), (v2, T2)) =
-  Term.aconv_untyped (v1, v2) andalso T1 = T2;
-
-structure Overload_Data = Generic_Data
-(
-  type T =
-    {variants : (term * typ) list Symtab.table,
-     oconsts : string Termtab.table};
-  val empty = {variants = Symtab.empty, oconsts = Termtab.empty};
-  val extend = I;
-
-  fun merge
-    ({variants = vtab1, oconsts = otab1},
-     {variants = vtab2, oconsts = otab2}) : T =
-    let
-      fun merge_oconsts _ (oconst1, oconst2) =
-        if oconst1 = oconst2 then oconst1
-        else err_duplicate_variant oconst1;
-    in
-      {variants = Symtab.merge_list variants_eq (vtab1, vtab2),
-       oconsts = Termtab.join merge_oconsts (otab1, otab2)}
-    end;
-);
-
-fun map_tables f g =
-  Overload_Data.map (fn {variants = vtab, oconsts = otab} =>
-    {variants = f vtab, oconsts = g otab});
-
-val is_overloaded = Symtab.defined o #variants o Overload_Data.get o Context.Proof;
-val get_variants = Symtab.lookup o #variants o Overload_Data.get o Context.Proof;
-val get_overloaded = Termtab.lookup o #oconsts o Overload_Data.get o Context.Proof;
-
-fun generic_add_overloaded oconst context =
-  if is_overloaded (Context.proof_of context) oconst then context
-  else map_tables (Symtab.update (oconst, [])) I context;
-
-fun generic_remove_overloaded oconst context =
-  let
-    fun remove_oconst_and_variants context oconst =
-      let
-        val remove_variants =
-          (case get_variants (Context.proof_of context) oconst of
-            NONE => I
-          | SOME vs => fold (Termtab.remove (op =) o rpair oconst o fst) vs);
-      in map_tables (Symtab.delete_safe oconst) remove_variants context end;
-  in
-    if is_overloaded (Context.proof_of context) oconst then remove_oconst_and_variants context oconst
-    else err_not_overloaded oconst
-  end;
-
-local
-  fun generic_variant add oconst t context =
-    let
-      val ctxt = Context.proof_of context;
-      val _ = if is_overloaded ctxt oconst then () else err_not_overloaded oconst;
-      val T = t |> fastype_of;
-      val t' = Term.map_types (K dummyT) t;
-    in
-      if add then
-        let
-          val _ =
-            (case get_overloaded ctxt t' of
-              NONE => ()
-            | SOME oconst' => err_duplicate_variant oconst');
-        in
-          map_tables (Symtab.cons_list (oconst, (t', T))) (Termtab.update (t', oconst)) context
-        end
-      else
-        let
-          val _ =
-            if member variants_eq (the (get_variants ctxt oconst)) (t', T) then ()
-            else err_not_a_variant oconst;
-        in
-          map_tables (Symtab.map_entry oconst (remove1 variants_eq (t', T)))
-            (Termtab.delete_safe t') context
-          |> (fn context =>
-            (case get_variants (Context.proof_of context) oconst of
-              SOME [] => generic_remove_overloaded oconst context
-            | _ => context))
-        end
-    end;
-in
-  val generic_add_variant = generic_variant true;
-  val generic_remove_variant = generic_variant false;
-end;
-
-
-(* check / uncheck *)
-
-fun unifiable_with thy T1 T2 =
-  let
-    val maxidx1 = Term.maxidx_of_typ T1;
-    val T2' = Logic.incr_tvar (maxidx1 + 1) T2;
-    val maxidx2 = Term.maxidx_typ T2' maxidx1;
-  in can (Sign.typ_unify thy (T1, T2')) (Vartab.empty, maxidx2) end;
-
-fun get_candidates ctxt (c, T) =
-  get_variants ctxt c
-  |> Option.map (map_filter (fn (t, T') =>
-    if unifiable_with (Proof_Context.theory_of ctxt) T T' then SOME t
-    else NONE));
-
-fun insert_variants ctxt t (oconst as Const (c, T)) =
-      (case get_candidates ctxt (c, T) of
-        SOME [] => err_unresolved_overloading ctxt (c, T) t []
-      | SOME [variant] => variant
-      | _ => oconst)
-  | insert_variants _ _ oconst = oconst;
-
-fun insert_overloaded ctxt =
-  let
-    fun proc t =
-      Term.map_types (K dummyT) t
-      |> get_overloaded ctxt
-      |> Option.map (Const o rpair (Term.type_of t));
-  in
-    Pattern.rewrite_term_top (Proof_Context.theory_of ctxt) [] [proc]
-  end;
-
-fun check ctxt =
-  map (fn t => Term.map_aterms (insert_variants ctxt t) t);
-
-fun uncheck ctxt ts =
-  if Config.get ctxt show_variants orelse exists (is_none o try Term.type_of) ts then ts
-  else map (insert_overloaded ctxt) ts;
-
-fun reject_unresolved ctxt =
-  let
-    val the_candidates = the o get_candidates ctxt;
-    fun check_unresolved t =
-      (case filter (is_overloaded ctxt o fst) (Term.add_consts t []) of
-        [] => t
-      | (cT :: _) => err_unresolved_overloading ctxt cT t (the_candidates cT));
-  in map check_unresolved end;
-
-
-(* setup *)
-
-val _ = Context.>>
-  (Syntax_Phases.term_check 0 "adhoc_overloading" check
-   #> Syntax_Phases.term_check 1 "adhoc_overloading_unresolved_check" reject_unresolved
-   #> Syntax_Phases.term_uncheck 0 "adhoc_overloading" uncheck);
-
-
-(* commands *)
-
-fun generic_adhoc_overloading_cmd add =
-  if add then
-    fold (fn (oconst, ts) =>
-      generic_add_overloaded oconst
-      #> fold (generic_add_variant oconst) ts)
-  else
-    fold (fn (oconst, ts) =>
-      fold (generic_remove_variant oconst) ts);
-
-fun adhoc_overloading_cmd' add args phi =
-  let val args' = args
-    |> map (apsnd (map_filter (fn t =>
-         let val t' = Morphism.term phi t;
-         in if Term.aconv_untyped (t, t') then SOME t' else NONE end)));
-  in generic_adhoc_overloading_cmd add args' end;
-
-fun adhoc_overloading_cmd add raw_args lthy =
-  let
-    fun const_name ctxt =
-      fst o dest_Const o Proof_Context.read_const {proper = false, strict = false} ctxt;  (* FIXME {proper = true, strict = true} (!?) *)
-    fun read_term ctxt = singleton (Variable.polymorphic ctxt) o Syntax.read_term ctxt;
-    val args =
-      raw_args
-      |> map (apfst (const_name lthy))
-      |> map (apsnd (map (read_term lthy)));
-  in
-    Local_Theory.declaration {syntax = true, pervasive = false}
-      (adhoc_overloading_cmd' add args) lthy
-  end;
-
-val _ =
-  Outer_Syntax.local_theory @{command_keyword adhoc_overloading}
-    "add adhoc overloading for constants / fixed variables"
-    (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd true);
-
-val _ =
-  Outer_Syntax.local_theory @{command_keyword no_adhoc_overloading}
-    "delete adhoc overloading for constants / fixed variables"
-    (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd false);
-
-end;
-
--- a/src/Tools/jEdit/src/Isabelle.props	Wed May 02 13:49:38 2018 +0200
+++ b/src/Tools/jEdit/src/Isabelle.props	Thu May 03 15:07:14 2018 +0200
@@ -30,6 +30,9 @@
 #menu actions and dockables
 plugin.isabelle.jedit.Plugin.menu.label=Isabelle
 plugin.isabelle.jedit.Plugin.menu= \
+  isabelle.preview \
+  isabelle.draft \
+  - \
   isabelle-debugger \
   isabelle-documentation \
   isabelle-monitor \
--- a/src/Tools/jEdit/src/jEdit.props	Wed May 02 13:49:38 2018 +0200
+++ b/src/Tools/jEdit/src/jEdit.props	Thu May 03 15:07:14 2018 +0200
@@ -215,6 +215,7 @@
 isabelle.decrease-font-size.shortcut2=C+SUBTRACT
 isabelle.decrease-font-size.shortcut=C+MINUS
 isabelle.decrease-font-size2.label=Decrease font size (clone)
+isabelle.draft.label=Show draft in browser
 isabelle.exclude-word-permanently.label=Exclude word permanently
 isabelle.exclude-word.label=Exclude word
 isabelle.include-word-permanently.label=Include word permanently
@@ -227,7 +228,7 @@
 isabelle.newline.label=Newline with indentation of Isabelle keywords
 isabelle.newline.shortcut=ENTER
 isabelle.options.label=Isabelle options
-isabelle.preview.label=HTML preview of PIDE document
+isabelle.preview.label=Show preview in browser
 isabelle.reset-continuous-checking.label=Reset continuous checking
 isabelle.reset-font-size.label=Reset font size
 isabelle.reset-node-required.label=Reset node required
--- a/src/Tools/jEdit/src/jedit_lib.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Tools/jEdit/src/jedit_lib.scala	Thu May 03 15:07:14 2018 +0200
@@ -129,6 +129,13 @@
 
   def buffer_file(buffer: Buffer): Option[JFile] = check_file(buffer_name(buffer))
 
+  def buffer_undo_in_progress[A](buffer: JEditBuffer, body: => A): A =
+  {
+    val undo_in_progress = buffer.isUndoInProgress
+    def set(b: Boolean) { Untyped.set[Boolean](buffer, "undoInProgress", b) }
+    try { set(true); body } finally { set(undo_in_progress) }
+  }
+
 
   /* main jEdit components */
 
--- a/src/Tools/jEdit/src/pretty_text_area.scala	Wed May 02 13:49:38 2018 +0200
+++ b/src/Tools/jEdit/src/pretty_text_area.scala	Thu May 03 15:07:14 2018 +0200
@@ -144,7 +144,7 @@
             JEdit_Lib.buffer_edit(getBuffer) {
               rich_text_area.active_reset()
               getBuffer.setFoldHandler(new Fold_Handling.Document_Fold_Handler(rendering))
-              setText(text)
+              JEdit_Lib.buffer_undo_in_progress(getBuffer, setText(text))
               setCaretPosition(0)
             }
           }