merged
authorwenzelm
Wed, 15 Apr 2020 14:31:21 +0200
changeset 71753 65b7d9ec05f5
parent 71752 5d360394f292 (current diff)
parent 71749 77232ff6b8f6 (diff)
child 71754 2006be3cb98b
merged
NEWS
src/Pure/Concurrent/standard_thread.ML
src/Pure/Concurrent/standard_thread.scala
--- a/Admin/components/components.sha1	Wed Apr 15 12:33:42 2020 +0200
+++ b/Admin/components/components.sha1	Wed Apr 15 14:31:21 2020 +0200
@@ -266,6 +266,7 @@
 3eca4b80710996fff87ed1340dcea2c5f6ebf4f7  scala-2.11.8.tar.gz
 0004e53f885fb165b50c95686dec40d99ab0bdbd  scala-2.12.0.tar.gz
 059cbdc58d36e3ac1fffcccd9139ecd34f271882  scala-2.12.10.tar.gz
+82056106aa6fd37c159ea76d16096c20a749cccd  scala-2.12.11.tar.gz
 74a8c3dab3a25a87357996ab3e95d825dc820fd0  scala-2.12.2.tar.gz
 d66796a68ec3254b46b17b1f8ee5bcc56a93aacf  scala-2.12.3.tar.gz
 1636556167dff2c191baf502c23f12e09181ef78  scala-2.12.4.tar.gz
--- a/Admin/components/main	Wed Apr 15 12:33:42 2020 +0200
+++ b/Admin/components/main	Wed Apr 15 14:31:21 2020 +0200
@@ -14,7 +14,7 @@
 opam-2.0.6
 polyml-5.8.1-20200228
 postgresql-42.2.9
-scala-2.12.10
+scala-2.12.11
 smbc-0.4.1
 spass-3.8ds-1
 sqlite-jdbc-3.30.1
--- a/CONTRIBUTORS	Wed Apr 15 12:33:42 2020 +0200
+++ b/CONTRIBUTORS	Wed Apr 15 14:31:21 2020 +0200
@@ -3,6 +3,10 @@
 listed as an author in one of the source files of this Isabelle distribution.
 
 
+Contributions to this Isabelle version
+--------------------------------------
+
+
 Contributions to Isabelle2020
 -----------------------------
 
--- a/NEWS	Wed Apr 15 12:33:42 2020 +0200
+++ b/NEWS	Wed Apr 15 14:31:21 2020 +0200
@@ -4,6 +4,27 @@
 (Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.)
 
 
+New in this Isabelle version
+----------------------------
+
+*** System ***
+
+* The command-line tool "isabelle console" now supports interrupts
+properly (on Linux and macOS).
+
+* The Isabelle/Scala "Progress" interface changed slightly and
+"No_Progress" has been discontinued. INCOMPATIBILITY, use "new Progress"
+instead.
+
+* General support for Isabelle/Scala system services, configured via the
+shell function "isabelle_scala_service" in etc/settings (e.g. of an
+Isabelle component); see implementations of class
+Isabelle_System.Service in Isabelle/Scala. This supersedes former
+"isabelle_scala_tools" and "isabelle_file_format": minor
+INCOMPATIBILITY.
+
+
+
 New in Isabelle2020 (April 2020)
 --------------------------------
 
--- a/README_REPOSITORY	Wed Apr 15 12:33:42 2020 +0200
+++ b/README_REPOSITORY	Wed Apr 15 14:31:21 2020 +0200
@@ -34,8 +34,6 @@
 
     ./bin/isabelle jedit -l HOL
 
-    ./bin/isabelle jedit -b -f  #optional: force fresh build of Isabelle/Scala
-
 4. Access documentation (bash shell commands):
 
     ./bin/isabelle build_doc -a
@@ -48,11 +46,9 @@
 
 Mercurial https://www.mercurial-scm.org belongs to source code
 management systems that follow the so-called paradigm of "distributed
-version control".  This means plain revision control without the
-legacy of CVS or SVN (and without the extra complexity introduced by
-git).  See also http://hginit.com/ for an introduction to the main
-ideas.  The Mercurial book http://hgbook.red-bean.com/ explains many
-more details.
+version control".  This means plain versioning without the legacy of
+SVN and the extra complexity of GIT.  See also
+https://www.mercurial-scm.org/learn
 
 Mercurial offers some flexibility in organizing the flow of changes,
 both between individual developers and designated pull/push areas that
@@ -100,7 +96,7 @@
 as follows:
 
   [ui]
-  username = XXX
+  username = ABC
 
 Isabelle contributors are free to choose either a short "login name"
 (for accounts at TU Munich) or a "full name" -- with or without mail
@@ -136,8 +132,7 @@
 quite easy to publish changed clones again on the web, using the
 ad-hoc command "hg serve -v", or the hgweb.cgi or hgwebdir.cgi scripts
 that are included in the Mercurial distribution, and send a "pull
-request" to someone else.  There are also public hosting services for
-Mercurial repositories, notably Bitbucket.
+request" to someone else.
 
 The downstream/upstream mode of operation is quite common in the
 distributed version control community, and works well for occasional
--- a/etc/options	Wed Apr 15 12:33:42 2020 +0200
+++ b/etc/options	Wed Apr 15 14:31:21 2020 +0200
@@ -148,6 +148,15 @@
   -- "ML process command prefix (process policy)"
 
 
+section "PIDE Build"
+
+option pide_build : bool = false
+  -- "build session heaps via PIDE"
+
+option pide_reports : bool = true
+  -- "report PIDE markup"
+
+
 section "Editor Session"
 
 public option editor_load_delay : real = 0.5
--- a/etc/settings	Wed Apr 15 12:33:42 2020 +0200
+++ b/etc/settings	Wed Apr 15 14:31:21 2020 +0200
@@ -20,10 +20,10 @@
 
 classpath "$ISABELLE_HOME/lib/classes/Pure.jar"
 
-isabelle_scala_tools 'isabelle.Tools'
-[ -d "$ISABELLE_HOME/Admin" ] && isabelle_scala_tools 'isabelle.Admin_Tools'
+isabelle_scala_service 'isabelle.Tools'
+[ -d "$ISABELLE_HOME/Admin" ] && isabelle_scala_service 'isabelle.Admin_Tools'
 
-isabelle_file_format 'isabelle.Bibtex$File_Format'
+isabelle_scala_service 'isabelle.Bibtex$File_Format'
 
 #paranoia settings -- avoid intrusion of alien options
 unset "_JAVA_OPTIONS"
--- a/lib/scripts/getfunctions	Wed Apr 15 12:33:42 2020 +0200
+++ b/lib/scripts/getfunctions	Wed Apr 15 14:31:21 2020 +0200
@@ -113,22 +113,6 @@
 }
 export -f classpath
 
-#Isabelle/Scala tools
-function isabelle_scala_tools ()
-{
-  local X=""
-  for X in "$@"
-  do
-    if [ -z "$ISABELLE_SCALA_TOOLS" ]; then
-      ISABELLE_SCALA_TOOLS="$X"
-    else
-      ISABELLE_SCALA_TOOLS="$ISABELLE_SCALA_TOOLS:$X"
-    fi
-  done
-  export ISABELLE_SCALA_TOOLS
-}
-export -f isabelle_scala_tools
-
 #Isabelle fonts
 function isabelle_fonts ()
 {
@@ -160,21 +144,21 @@
 }
 export -f isabelle_fonts_hidden
 
-#file formats
-function isabelle_file_format ()
+#Isabelle/Scala services
+function isabelle_scala_service ()
 {
   local X=""
   for X in "$@"
   do
-    if [ -z "$ISABELLE_FILE_FORMATS" ]; then
-      ISABELLE_FILE_FORMATS="$X"
+    if [ -z "$ISABELLE_SCALA_SERVICES" ]; then
+      ISABELLE_SCALA_SERVICES="$X"
     else
-      ISABELLE_FILE_FORMATS="$ISABELLE_FILE_FORMATS:$X"
+      ISABELLE_SCALA_SERVICES="$ISABELLE_SCALA_SERVICES:$X"
     fi
   done
-  export ISABELLE_FILE_FORMATS
+  export ISABELLE_SCALA_SERVICES
 }
-export -f isabelle_file_format
+export -f isabelle_scala_service
 
 #administrative build
 if [ -e "$ISABELLE_HOME/Admin/build" ]; then
--- a/src/Doc/System/Environment.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Doc/System/Environment.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -390,8 +390,8 @@
   \<^medskip>
   This is how to invoke a function body with proper return code and printing
   of errors, and without printing of a redundant \<^verbatim>\<open>val it = (): unit\<close> result:
-  @{verbatim [display] \<open>isabelle process -e 'Command_Line.tool0 (fn () => writeln "OK")'\<close>}
-  @{verbatim [display] \<open>isabelle process -e 'Command_Line.tool0 (fn () => error "Bad")'\<close>}
+  @{verbatim [display] \<open>isabelle process -e 'Command_Line.tool (fn () => writeln "OK")'\<close>}
+  @{verbatim [display] \<open>isabelle process -e 'Command_Line.tool (fn () => error "Bad")'\<close>}
 \<close>
 
 
--- a/src/HOL/Analysis/Abstract_Euclidean_Space.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Abstract_Euclidean_Space.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -158,7 +158,7 @@
   by (simp add: contractible_imp_path_connected_space)
 
 lemma connected_Euclidean_space: "connected_space (Euclidean_space n)"
-  by (simp add: contractible_Euclidean_space contractible_imp_connected_space)
+  by (simp add: contractible_imp_connected_space)
 
 lemma locally_path_connected_Euclidean_space:
    "locally_path_connected_space (Euclidean_space n)"
--- a/src/HOL/Analysis/Abstract_Topology.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Abstract_Topology.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -2383,12 +2383,6 @@
     by (simp add: closed_map_id continuous_closed_imp_quotient_map)
 qed
 
-lemma homeomorphic_maps_i [simp]:"homeomorphic_maps X Y id id \<longleftrightarrow> Y = X"
-  by (metis (full_types) eq_id_iff homeomorphic_maps_id)
-
-lemma homeomorphic_map_i [simp]: "homeomorphic_map X Y id \<longleftrightarrow> Y = X"
-  by (metis (no_types) eq_id_iff homeomorphic_map_id)
-
 lemma homeomorphic_map_compose:
   assumes "homeomorphic_map X Y f" "homeomorphic_map Y X'' g"
   shows "homeomorphic_map X X'' (g \<circ> f)"
--- a/src/HOL/Analysis/Arcwise_Connected.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Arcwise_Connected.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1789,13 +1789,13 @@
             have "False" if "\<xi> \<in> F n" "u < \<xi>" "\<xi> < v" for \<xi>
             proof -
               have "\<xi> \<notin> {z..v}"
-                by (metis DiffI disjoint_iff_not_equal less_irrefl singletonD that v(3))
+                by (metis DiffI disjoint_iff_not_equal less_irrefl singletonD that(1,3) v(3))
               moreover have "\<xi> \<notin> {w..z} \<inter> F n"
                 by (metis equals0D wzF_null)
               ultimately have "\<xi> \<in> {u..w}"
                 using that by auto
               then show ?thesis
-                by (metis DiffI disjoint_iff_not_equal less_eq_real_def not_le singletonD that u(3))
+                by (metis DiffI disjoint_iff_not_equal less_eq_real_def not_le singletonD that(1,2) u(3))
             qed
             moreover
             have "\<lbrakk>\<xi> \<in> F n; v < \<xi>; \<xi> < u\<rbrakk> \<Longrightarrow> False" for \<xi>
@@ -1811,7 +1811,7 @@
           by (metis dist_norm dist_triangle_half_r less_irrefl)
       qed (auto simp: open_segment_commute)
       show ?thesis
-        unfolding \<phi>_def by (metis (no_types, hide_lams) INT_I Inf_lower2 rangeI that F01 subsetCE pqF)
+        unfolding \<phi>_def by (metis (no_types, hide_lams) INT_I Inf_lower2 rangeI that(3) F01 subsetCE pqF)
     qed
     show "closed {0..1::real}" by auto
   qed (meson \<phi>_def)
@@ -1972,7 +1972,7 @@
     show "pathstart (subpath u v g) = a" "pathfinish (subpath u v g) = b"
       by (simp_all add: \<open>a = g u\<close> \<open>b = g v\<close>)
     show "path_image (subpath u v g) \<subseteq> path_image p"
-      by (metis \<open>arc g\<close> \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> arc_imp_path order_trans pag path_image_def path_image_subpath_subset ph)
+      by (metis \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> order_trans pag path_image_def path_image_subpath_subset ph)
     show "arc (subpath u v g)"
       using \<open>arc g\<close> \<open>a = g u\<close> \<open>b = g v\<close> \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> arc_subpath_arc \<open>a \<noteq> b\<close> by blast
   qed
@@ -2049,7 +2049,7 @@
                                  and g: "pathstart g = y" "pathfinish g = z"
     using \<open>y \<noteq> z\<close> by (force simp: path_connected_arcwise)
   have "compact (g -` frontier S \<inter> {0..1})"
-    apply (simp add: compact_eq_bounded_closed bounded_Int bounded_closed_interval)
+    apply (simp add: compact_eq_bounded_closed bounded_Int)
      apply (rule closed_vimage_Int)
     using \<open>arc g\<close> apply (auto simp: arc_def path_def)
     done
--- a/src/HOL/Analysis/Ball_Volume.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Ball_Volume.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -46,7 +46,7 @@
        (auto intro!: derivative_eq_intros continuous_intros simp: set_borel_measurable_def)
   also have "\<dots> = (\<integral>\<^sup>+ x. 2 * ennreal ((1 - x\<^sup>2) powr (real n / 2) * indicator {0..1} x) \<partial>lborel)"
     by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) 
-       (auto simp: indicator_def powr_minus powr_half_sqrt field_split_simps ennreal_mult' mult_ac)
+       (auto simp: indicator_def powr_minus powr_half_sqrt field_split_simps ennreal_mult')
   also have "\<dots> = (\<integral>\<^sup>+ x. ennreal ((1 - x\<^sup>2) powr (real n / 2) * indicator {0..1} x) \<partial>lborel) +
                     (\<integral>\<^sup>+ x. ennreal ((1 - x\<^sup>2) powr (real n / 2) * indicator {0..1} x) \<partial>lborel)"
     (is "_ = ?I + _") by (simp add: mult_2 nn_integral_add)
@@ -137,7 +137,7 @@
       proof (cases "y \<in> {-r<..<r}")
         case True
         hence "y\<^sup>2 < r\<^sup>2" by (subst real_sqrt_less_iff [symmetric]) auto
-        thus ?thesis by (subst insert.IH) (auto simp: real_sqrt_less_iff)
+        thus ?thesis by (subst insert.IH) (auto)
       qed (insert elim, auto)
     qed
   qed
--- a/src/HOL/Analysis/Bochner_Integration.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Bochner_Integration.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -403,7 +403,7 @@
   from g have "simple_bochner_integral M (\<lambda>x. f (g x)) =
     (\<Sum>y\<in>g ` space M. measure M {x \<in> space M. g x = y} *\<^sub>R f y)"
     by (intro simple_bochner_integral_partition)
-       (auto simp: simple_bochner_integrable_compose2[where p="\<lambda>x y. f x"] zero
+       (auto simp: simple_bochner_integrable_compose2[where p="\<lambda>x y. f x"]
              elim: simple_bochner_integrable.cases)
   also have "\<dots> = f (simple_bochner_integral M g)"
     by (simp add: simple_bochner_integral_def sum scale)
@@ -1379,7 +1379,7 @@
 lemma integral_minus_iff[simp]:
   "integrable M (\<lambda>x. - f x ::'a::{banach, second_countable_topology}) \<longleftrightarrow> integrable M f"
   unfolding integrable_iff_bounded
-  by (auto intro: borel_measurable_uminus[of "\<lambda>x. - f x" M, simplified])
+  by (auto)
 
 lemma integrable_indicator_iff:
   "integrable M (indicator A::_ \<Rightarrow> real) \<longleftrightarrow> A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
@@ -1634,7 +1634,7 @@
       with seq show "\<exists>N. \<forall>n\<ge>N. 0 \<le> integral\<^sup>L M (U n)"
         by auto
     qed
-  qed (auto simp: integrable_mult_left_iff)
+  qed (auto)
   also have "\<dots> = integral\<^sup>L M f"
     using nonneg by (auto intro!: integral_cong_AE)
   finally show ?thesis .
@@ -1657,7 +1657,7 @@
         by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure)
     next
       case (mult f c) then show ?case
-        by (auto simp add: integrable_mult_left_iff nn_integral_cmult ennreal_mult integral_nonneg_AE)
+        by (auto simp add: nn_integral_cmult ennreal_mult integral_nonneg_AE)
     next
       case (add g f)
       then have "integrable M f" "integrable M g"
@@ -2756,7 +2756,7 @@
       by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong)
   qed
   also have "\<dots> = (\<Sum>a\<in>A. f a * measure M {a})"
-    using finite by (subst integral_sum) (auto simp add: integrable_mult_left_iff)
+    using finite by (subst integral_sum) (auto)
   finally show ?thesis .
 qed
 
--- a/src/HOL/Analysis/Borel_Space.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Borel_Space.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1758,7 +1758,7 @@
       also have "\<dots> < ?I i / 2 + ?I i / 2"
         by (intro add_strict_mono d less_trans[OF _ j] *)
       also have "\<dots> \<le> ?I i"
-        by (simp add: field_simps of_nat_Suc)
+        by (simp add: field_simps)
       finally show "dist (f y) (f z) \<le> ?I i"
         by simp
     qed
--- a/src/HOL/Analysis/Bounded_Linear_Function.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Bounded_Linear_Function.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -420,7 +420,7 @@
     using b
     by (simp add: inner_sum_left inner_Basis if_distrib cong: if_cong) (simp add: sum.swap)
   also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> b)))"
-    using b by (simp add: sum.delta)
+    using b by (simp)
   also have "\<dots> = f x \<bullet> b"
     by (metis (mono_tags, lifting) Linear_Algebra.linear_componentwise linear_axioms)
   finally show "(\<Sum>j\<in>Basis. \<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j)) *\<^sub>R j) \<bullet> b = f x \<bullet> b" .
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -149,7 +149,7 @@
     apply (rule_tac x=r in exI, simp)
      apply (rule homotopic_with_trans, assumption)
      apply (rule_tac f = "r \<circ> f" and g="r \<circ> id" in homotopic_with_eq)
-        apply (rule_tac Y=S in homotopic_compose_continuous_left)
+        apply (rule_tac Y=S in homotopic_with_compose_continuous_left)
          apply (auto simp: homotopic_with_sym)
     done
 qed
--- a/src/HOL/Analysis/Cartesian_Space.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Cartesian_Space.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1436,7 +1436,7 @@
       using idplus [OF \<open>m \<noteq> n\<close>] by simp
   qed
   then show ?thesis
-    by (metis \<open>linear f\<close> matrix_vector_mul)
+    by (metis \<open>linear f\<close> matrix_vector_mul(2))
 qed
 
 end
\ No newline at end of file
--- a/src/HOL/Analysis/Change_Of_Vars.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Change_Of_Vars.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -543,7 +543,7 @@
                 also have "\<dots> = norm (f' x (y - x) - (f y - f x)) / r"
                   using \<open>r > 0\<close> by (simp add: divide_simps scale_right_diff_distrib [symmetric])
                 also have "\<dots> \<le> norm (f y - (f x + f' x (y - x))) / norm (y - x)"
-                  using that \<open>r > 0\<close> False by (simp add: algebra_simps field_split_simps dist_norm norm_minus_commute mult_right_mono)
+                  using that \<open>r > 0\<close> False by (simp add: field_split_simps dist_norm norm_minus_commute mult_right_mono)
                 also have "\<dots> < k"
                   using that \<open>0 < \<zeta>\<close> by (simp add: dist_commute r_def  \<zeta> [OF \<open>y \<in> S\<close> False])
                 finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" .
@@ -1656,7 +1656,7 @@
             show "linear ((*v) (matrix (f' x) - B))"
               by (rule matrix_vector_mul_linear)
             have "((\<lambda>y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within S)"
-              using tendsto_minus [OF lim_df] by (simp add: algebra_simps field_split_simps)
+              using tendsto_minus [OF lim_df] by (simp add: field_split_simps)
             then show "((\<lambda>y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within S)"
             proof (rule Lim_transform)
               have "((\<lambda>y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \<longlongrightarrow> 0) (at x within S)"
@@ -1724,7 +1724,7 @@
               then show "((\<lambda>y. (matrix (f' x) - B) *v (y - x) /\<^sub>R
                            norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0)
                           (at x within S)"
-                by (simp add: algebra_simps diff lin_df matrix_vector_mul_linear scalar_mult_eq_scaleR)
+                by (simp add: algebra_simps diff lin_df scalar_mult_eq_scaleR)
             qed
           qed (use x in \<open>simp; auto simp: not_less\<close>)
           ultimately have "f' x = (*v) B"
@@ -2168,7 +2168,7 @@
         proof (rule order_trans [OF measT])
           have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n"
             using \<open>c > 0\<close>
-            apply (simp add: algebra_simps power_mult_distrib)
+            apply (simp add: algebra_simps)
             by (metis Suc_pred power_Suc zero_less_card_finite)
           also have "\<dots> \<le> (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n"
             by (rule mult_right_mono [OF d]) auto
@@ -2239,7 +2239,7 @@
           using \<open>c > 0\<close> by (simp add: content_cbox_if_cart)
         finally have "sum ?\<mu> \<F> \<le> (2 ^ ?m * c ^ ?m)" .
         then show ?thesis
-          using \<open>e > 0\<close> \<open>c > 0\<close> by (auto simp: field_split_simps mult_less_0_iff)
+          using \<open>e > 0\<close> \<open>c > 0\<close> by (auto simp: field_split_simps)
       qed
       finally show ?thesis .
     qed
@@ -2384,7 +2384,7 @@
         done
       ultimately show ?thesis
         using \<open>y > 0\<close> integral_restrict_Int [of S "{t. h n (g t) = y}" "\<lambda>t. \<bar>det (matrix (g' t))\<bar> * y"]
-        apply (simp add: integrable_on_indicator integrable_on_cmult_iff integral_indicator)
+        apply (simp add: integrable_on_indicator integral_indicator)
         apply (simp add: indicator_def if_distrib cong: if_cong)
         done
     qed
--- a/src/HOL/Analysis/Complex_Analysis_Basics.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -489,7 +489,7 @@
 
 lemma analytic_on_sum [analytic_intros]:
   "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on S) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) analytic_on S"
-  by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add)
+  by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_add)
 
 lemma deriv_left_inverse:
   assumes "f holomorphic_on S" and "g holomorphic_on T"
@@ -579,11 +579,11 @@
 
 lemma deriv_cmult_at:
   "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. c * f w) z = c * deriv f z"
-by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const)
+by (auto simp: complex_derivative_mult_at)
 
 lemma deriv_cmult_right_at:
   "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. f w * c) z = deriv f z * c"
-by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const)
+by (auto simp: complex_derivative_mult_at)
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Complex differentiation of sequences and series\<close>
 
@@ -758,7 +758,7 @@
     apply (rule field_differentiable_bound
       [where f' = "\<lambda>w. f (Suc n) w * (z - w)^n / (fact n)"
          and S = "closed_segment w z", OF convex_closed_segment])
-    apply (auto simp: ends_in_segment DERIV_subset [OF sum_deriv wzs]
+    apply (auto simp: DERIV_subset [OF sum_deriv wzs]
                   norm_divide norm_mult norm_power divide_le_cancel cmod_bound)
     done
   also have "...  \<le> B * norm (z - w) ^ Suc n / (fact n)"
--- a/src/HOL/Analysis/Complex_Transcendental.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Complex_Transcendental.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -612,7 +612,7 @@
   apply (simp add: cos_add cmod_power2 cos_of_real sin_of_real Complex_eq)
   apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
   apply (simp only: left_diff_distrib [symmetric] power_mult_distrib sin_squared_eq)
-  apply (simp add: power2_eq_square algebra_simps field_split_simps)
+  apply (simp add: power2_eq_square field_split_simps)
   done
 
 lemma norm_sin_squared:
@@ -621,7 +621,7 @@
   apply (simp add: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double Complex_eq)
   apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
   apply (simp only: left_diff_distrib [symmetric] power_mult_distrib cos_squared_eq)
-  apply (simp add: power2_eq_square algebra_simps field_split_simps)
+  apply (simp add: power2_eq_square field_split_simps)
   done
 
 lemma exp_uminus_Im: "exp (- Im z) \<le> exp (cmod z)"
@@ -1102,7 +1102,7 @@
   by auto
 
 lemma Arg2pi_cnj_eq_inverse: "z\<noteq>0 \<Longrightarrow> Arg2pi (cnj z) = Arg2pi (inverse z)"
-  apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric] mult.commute)
+  apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric])
   by (metis of_real_power zero_less_norm_iff zero_less_power)
 
 lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \<in> \<real> then Arg2pi z else 2*pi - Arg2pi z)"
@@ -1946,7 +1946,7 @@
   by (metis Arg_eq_0 Arg_inverse inverse_inverse_eq)
 
 lemma Arg_cnj_eq_inverse: "z\<noteq>0 \<Longrightarrow> Arg (cnj z) = Arg (inverse z)"
-  apply (simp add: Arg_eq_iff field_split_simps complex_norm_square [symmetric] mult.commute)
+  apply (simp add: Arg_eq_iff field_split_simps complex_norm_square [symmetric])
   by (metis of_real_power zero_less_norm_iff zero_less_power)
 
 lemma Arg_cnj: "Arg(cnj z) = (if z \<in> \<real> then Arg z else - Arg z)"
@@ -2936,7 +2936,7 @@
     unfolding Arctan_def scaleR_conv_of_real
     apply (intro derivative_eq_intros | simp add: nz0 *)+
     using nz0 nz1 zz
-    apply (simp add: algebra_simps field_split_simps power2_eq_square)
+    apply (simp add: field_split_simps power2_eq_square)
     apply algebra
     done
 qed
@@ -3014,7 +3014,7 @@
   have "\<exists>c. \<forall>u\<in>ball 0 1. Arctan u - G u = c"
   proof (rule has_field_derivative_zero_constant)
     fix u :: complex assume "u \<in> ball 0 1"
-    hence u: "norm u < 1" by (simp add: dist_0_norm)
+    hence u: "norm u < 1" by (simp)
     define K where "K = (norm u + 1) / 2"
     from u and abs_Im_le_cmod[of u] have Im_u: "\<bar>Im u\<bar> < 1" by linarith
     from u have K: "0 \<le> K" "norm u < K" "K < 1" by (simp_all add: K_def)
@@ -3634,7 +3634,7 @@
   assumes "z\<^sup>2 \<noteq> 1" shows "cos(Arcsin z) \<noteq> 0"
 proof -
   have eq: "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = z\<^sup>2 * (z\<^sup>2 - 1)"
-    by (simp add: power_mult_distrib algebra_simps)
+    by (simp add: algebra_simps)
   have "\<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> z\<^sup>2 - 1"
   proof
     assume "\<i> * z * (csqrt (1 - z\<^sup>2)) = z\<^sup>2 - 1"
@@ -3665,7 +3665,7 @@
   assumes "z\<^sup>2 \<noteq> 1" shows "sin(Arccos z) \<noteq> 0"
 proof -
   have eq: "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = -(z\<^sup>2) * (1 - z\<^sup>2)"
-    by (simp add: power_mult_distrib algebra_simps)
+    by (simp add: algebra_simps)
   have "\<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> 1 - z\<^sup>2"
   proof
     assume "\<i> * z * (csqrt (1 - z\<^sup>2)) = 1 - z\<^sup>2"
@@ -3971,7 +3971,7 @@
     shows "exp(2 * of_real pi * \<i> * of_nat j / of_nat n)^n = 1"
 proof -
   have *: "of_nat j * (complex_of_real pi * 2) = complex_of_real (2 * real j * pi)"
-    by (simp add: of_real_numeral)
+    by (simp)
   then show ?thesis
     apply (simp add: exp_of_nat_mult [symmetric] mult_ac exp_Euler)
     apply (simp only: * cos_of_real sin_of_real)
@@ -4006,7 +4006,7 @@
    note * = this
   show ?thesis
     using assms
-    by (simp add: exp_eq field_split_simps mult_ac of_real_numeral *)
+    by (simp add: exp_eq field_split_simps *)
 qed
 
 corollary bij_betw_roots_unity:
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -346,7 +346,7 @@
 proof-
   have "{a..b} = {a..} \<inter> {..b}" by auto
   also have "interior \<dots> = {a<..} \<inter> {..<b}"
-    by (simp add: interior_real_atLeast interior_real_atMost)
+    by (simp)
   also have "\<dots> = {a<..<b}" by auto
   finally show ?thesis .
 qed
--- a/src/HOL/Analysis/Cross3.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Cross3.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -100,7 +100,7 @@
 
 lemma  cross_basis_nonzero:
   "u \<noteq> 0 \<Longrightarrow> u \<times> axis 1 1 \<noteq> 0 \<or> u \<times> axis 2 1 \<noteq> 0 \<or> u \<times> axis 3 1 \<noteq> 0"
-  by (clarsimp simp add: axis_def cross3_simps) (metis vector_3 exhaust_3)
+  by (clarsimp simp add: axis_def cross3_simps) (metis exhaust_3)
 
 lemma  cross_dot_cancel:
   fixes x::"real^3"
--- a/src/HOL/Analysis/Derivative.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Derivative.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1056,7 +1056,7 @@
       by (rule openE)
     then have "\<exists>c. \<forall>x\<in>ball a e. f x = c"
       by (intro has_derivative_zero_constant)
-         (auto simp: at_within_open[OF _ open_ball] f convex_ball)
+         (auto simp: at_within_open[OF _ open_ball] f)
     with \<open>0<e\<close> have "\<forall>x\<in>ball a e. f a = f x"
       by auto
     then show "eventually (\<lambda>b. f a = f b) (at a within s)"
@@ -2187,7 +2187,7 @@
   show "((\<lambda>y. (exp (y *\<^sub>R A) - exp (t *\<^sub>R A) - (y - t) *\<^sub>R (exp (t *\<^sub>R A) * A)) /\<^sub>R norm (y - t)) \<longlongrightarrow> 0)
       (at t within T)"
     by (rule Lim_transform_eventually)
-      (auto simp: algebra_simps field_split_simps exp_add_commuting[symmetric])
+      (auto simp: field_split_simps exp_add_commuting[symmetric])
 qed (rule bounded_linear_scaleR_left)
 
 lemma exp_times_scaleR_commute: "exp (t *\<^sub>R A) * A = A * exp (t *\<^sub>R A)"
--- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -109,7 +109,7 @@
 qed
 
 lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"
-  by (simp add: open_dist subset_eq mem_ball Ball_def dist_commute)
+  by (simp add: open_dist subset_eq Ball_def dist_commute)
 
 lemma openI [intro?]: "(\<And>x. x\<in>S \<Longrightarrow> \<exists>e>0. ball x e \<subseteq> S) \<Longrightarrow> open S"
   by (auto simp: open_contains_ball)
--- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -798,7 +798,7 @@
 corollary cobounded_imp_unbounded:
     fixes S :: "'a::{real_normed_vector, perfect_space} set"
     shows "bounded (- S) \<Longrightarrow> \<not> bounded S"
-  using bounded_Un [of S "-S"]  by (simp add: sup_compl_top)
+  using bounded_Un [of S "-S"]  by (simp)
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Relations among convergence and absolute convergence for power series\<close>
 
--- a/src/HOL/Analysis/Embed_Measure.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Embed_Measure.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -251,7 +251,7 @@
   have fg[simp]: "\<And>A. inj_on (map_prod f g) A" "\<And>A. inj_on f A" "\<And>A. inj_on g A"
     using f g by (auto simp: inj_on_def)
 
-  note complete_lattice_class.Sup_insert[simp del] ccSup_insert[simp del] SUP_insert[simp del]
+  note complete_lattice_class.Sup_insert[simp del] ccSup_insert[simp del]
      ccSUP_insert[simp del]
   show sets: "sets ?L = sets (embed_measure (M \<Otimes>\<^sub>M N) (map_prod f g))"
     unfolding map_prod_def[symmetric]
--- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -3357,7 +3357,7 @@
   have eq: "?g =
         (\<lambda>x. \<Sum>i\<in>Basis. ((\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0) \<circ>
                (\<lambda>x. norm(\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f) x)"
-    by (simp add: sum.delta)
+    by (simp)
   have *: "(\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0) \<circ>
            (\<lambda>x. norm (\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f
            absolutely_integrable_on S"
--- a/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -742,7 +742,7 @@
       show "continuous_on UNIV (g \<circ> (\<theta> n))" for n :: nat
         unfolding \<theta>_def
         apply (intro continuous_on_compose2 [OF contg] continuous_intros conth)
-        apply (auto simp: aibi * mem_box less_max_iff_disj min_less_iff_disj algebra_simps field_split_simps)
+        apply (auto simp: aibi * mem_box less_max_iff_disj min_less_iff_disj field_split_simps)
         done
       show "(\<lambda>n. (g \<circ> \<theta> n) x) \<longlonglongrightarrow> g (f x)"
         if "x \<notin> N" for x
--- a/src/HOL/Analysis/Extended_Real_Limits.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Extended_Real_Limits.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -304,7 +304,7 @@
   fix P d
   assume "0 < d" and "\<forall>y. y \<in> S \<longrightarrow> y \<noteq> x \<and> dist y x < d \<longrightarrow> P y"
   then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
-    by (auto simp: zero_less_dist_iff dist_commute)
+    by (auto simp: dist_commute)
   then show "\<exists>r>0. Inf (f ` (Collect P)) \<le> Inf (f ` (S \<inter> ball x r - {x}))"
     by (intro exI[of _ d] INF_mono conjI \<open>0 < d\<close>) auto
 next
@@ -324,7 +324,7 @@
   fix P d
   assume "0 < d" and "\<forall>y. y \<in> S \<longrightarrow> y \<noteq> x \<and> dist y x < d \<longrightarrow> P y"
   then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
-    by (auto simp: zero_less_dist_iff dist_commute)
+    by (auto simp: dist_commute)
   then show "\<exists>r>0. Sup (f ` (S \<inter> ball x r - {x})) \<le> Sup (f ` (Collect P))"
     by (intro exI[of _ d] SUP_mono conjI \<open>0 < d\<close>) auto
 next
--- a/src/HOL/Analysis/FPS_Convergence.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/FPS_Convergence.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -429,7 +429,7 @@
   have "(\<lambda>n. norm (c * z) ^ n /\<^sub>R fact n) sums exp (norm (c * z))"
     by (rule exp_converges)
   also have "(\<lambda>n. norm (c * z) ^ n /\<^sub>R fact n) = (\<lambda>n. norm (fps_nth (fps_exp c) n * z ^ n))"
-    by (rule ext) (simp add: norm_divide norm_mult norm_power field_split_simps power_mult_distrib)
+    by (rule ext) (simp add: norm_divide norm_mult norm_power field_split_simps)
   finally have "summable \<dots>" by (simp add: sums_iff)
   thus "summable (\<lambda>n. fps_nth (fps_exp c) n * z ^ n)"
     by (rule summable_norm_cancel)
@@ -585,7 +585,7 @@
 lemma eval_fps_exp [simp]:
   fixes c :: "'a :: {banach, real_normed_field}"
   shows "eval_fps (fps_exp c) z = exp (c * z)" unfolding eval_fps_def exp_def
-  by (simp add: eval_fps_def exp_def scaleR_conv_of_real field_split_simps power_mult_distrib)
+  by (simp add: eval_fps_def exp_def scaleR_conv_of_real field_split_simps)
 
 text \<open>
   The case of division is more complicated and will therefore not be handled here.
--- a/src/HOL/Analysis/Fashoda_Theorem.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Fashoda_Theorem.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -677,10 +677,10 @@
       path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union>
       path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union>
       path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2)
-      by(auto simp add: path_image_join path_linepath)
+      by(auto simp add: path_image_join)
   have abab: "cbox a b \<subseteq> cbox ?a ?b"
     unfolding interval_cbox_cart[symmetric]
-    by (auto simp add:less_eq_vec_def forall_2 vector_2)
+    by (auto simp add:less_eq_vec_def forall_2)
   obtain z where
     "z \<in> path_image
           (linepath (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) +++
--- a/src/HOL/Analysis/Function_Metric.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Function_Metric.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -187,7 +187,7 @@
     have "dist x y \<le> 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} + (1/2)^N"
       by (rule dist_fun_le_dist_first_terms)
     also have "... \<le> 2 * f + e / 8"
-      apply (rule add_mono) using ** \<open>2^N > 8/e\<close> by(auto simp add: algebra_simps field_split_simps)
+      apply (rule add_mono) using ** \<open>2^N > 8/e\<close> by(auto simp add: field_split_simps)
     also have "... \<le> e/2 + e/8"
       unfolding f_def by auto
     also have "... < e"
@@ -331,7 +331,7 @@
       also have "... < (1/2)^k * min e 1"
         using C \<open>m \<ge> N\<close> \<open>n \<ge> N\<close> by auto
       finally have "min (dist (u m i) (u n i)) 1 < min e 1"
-        by (auto simp add: algebra_simps field_split_simps)
+        by (auto simp add: field_split_simps)
       then show "dist (u m i) (u n i) < e" by auto
     qed
     then show "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m i) (u n i) < e"
@@ -371,7 +371,7 @@
         using dist_fun_le_dist_first_terms by auto
       also have "... \<le> 2 * e/4 + e/4"
         apply (rule add_mono)
-        using ** \<open>2^N > 4/e\<close> less_imp_le by (auto simp add: algebra_simps field_split_simps)
+        using ** \<open>2^N > 4/e\<close> less_imp_le by (auto simp add: field_split_simps)
       also have "... < e" by auto
       finally show ?thesis by simp
     qed
--- a/src/HOL/Analysis/Function_Topology.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Function_Topology.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -198,7 +198,7 @@
   have "(\<Pi>\<^sub>E i\<in>I. topspace (T i)) \<in> {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
     using openin_topspace not_finite_existsD by auto
   then show "(\<Pi>\<^sub>E i\<in>I. topspace (T i)) \<subseteq> topspace (product_topology T I)"
-    unfolding product_topology_def using PiE_def by (auto simp add: topology_generated_by_topspace)
+    unfolding product_topology_def using PiE_def by (auto)
 qed
 
 lemma product_topology_basis:
--- a/src/HOL/Analysis/Further_Topology.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Further_Topology.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -239,14 +239,14 @@
   obtain c where homhc: "homotopic_with_canon (\<lambda>z. True) (sphere 0 1 \<inter> S) (sphere 0 1 \<inter> T) h (\<lambda>x. c)"
     apply (rule_tac c="-d" in that)
     apply (rule homotopic_with_eq)
-       apply (rule homotopic_compose_continuous_left [OF hom_hd conT0 sub0T])
+       apply (rule homotopic_with_compose_continuous_left [OF hom_hd conT0 sub0T])
     using d apply (auto simp: h_def)
     done
   show ?thesis
     apply (rule_tac x=c in exI)
     apply (rule homotopic_with_trans [OF _ homhc])
     apply (rule homotopic_with_eq)
-       apply (rule homotopic_compose_continuous_left [OF homfg conT0 sub0T])
+       apply (rule homotopic_with_compose_continuous_left [OF homfg conT0 sub0T])
       apply (auto simp: h_def)
     done
 qed
@@ -3682,7 +3682,7 @@
       qed
       show "disjoint \<V>"
         apply (clarsimp simp add: \<V>_def pairwise_def disjnt_def add.commute [of _ "x*y" for x y]
-                        image_add_ball ball_eq_ball_iff)
+                        ball_eq_ball_iff)
         apply (rule disjoint_ballI)
         apply (auto simp: dist_norm neq_iff)
         by (metis norm_minus_commute xy)+
@@ -4221,7 +4221,7 @@
       using fim him by force
   qed auto
   then have "homotopic_with_canon (\<lambda>f. True) (sphere a r) (sphere b s) (k \<circ> (h \<circ> f)) (k \<circ> (\<lambda>x. c))"
-    by (rule homotopic_compose_continuous_left [OF _ contk kim])
+    by (rule homotopic_with_compose_continuous_left [OF _ contk kim])
   then have "homotopic_with_canon (\<lambda>z. True) (sphere a r) (sphere b s) f (\<lambda>x. k c)"
     apply (rule homotopic_with_eq, auto)
     by (metis fim hk homeomorphism_def image_subset_iff mem_sphere)
@@ -5321,7 +5321,7 @@
   assumes "connected S" and C: "C \<in> components(- S)" shows "connected(frontier C)"
   apply (rule connected_frontier_simple)
   using C in_components_connected apply blast
-  by (metis Compl_eq_Diff_UNIV connected_UNIV assms top_greatest component_complement_connected)
+  by (metis assms component_complement_connected)
 
 lemma connected_frontier_disjoint:
   fixes S :: "'a :: euclidean_space set"
--- a/src/HOL/Analysis/Gamma_Function.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Gamma_Function.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -395,7 +395,7 @@
     from d have "\<lceil>2 * (cmod z + d)\<rceil> \<ge> \<lceil>0::real\<rceil>"
       by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all
     hence "2 * (norm z + d) \<le> of_nat (nat \<lceil>2 * (norm z + d)\<rceil>)" unfolding N_def
-      by (simp_all add: le_of_int_ceiling)
+      by (simp_all)
     also have "... \<le> of_nat N" unfolding N_def
       by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1)
     finally have "of_nat N \<ge> 2 * (norm z + d)" .
@@ -540,7 +540,7 @@
     also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))"
       by (simp add: harm_def sum_subtractf sum_distrib_left divide_inverse)
     also from n have "\<dots> - ?g n = 0"
-      by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps Ln_of_nat)
+      by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps)
     finally show "(\<Sum>k<n. ?f k) - ?g n = 0" .
   qed
   show "(\<lambda>n. (\<Sum>k<n. ?f k) - ?g n) \<longlonglongrightarrow> 0" by (subst tendsto_cong[OF A]) simp_all
@@ -583,7 +583,7 @@
     with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))"
       by (simp add: field_split_simps eq_neg_iff_add_eq_0 del: of_nat_Suc)
     also have "norm \<dots> = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))"
-      by (simp add: norm_divide norm_mult field_split_simps add_ac del: of_nat_Suc)
+      by (simp add: norm_divide norm_mult field_split_simps del: of_nat_Suc)
     also {
       from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \<le> of_nat (Suc n) / (2 * norm t)"
         by (intro divide_left_mono mult_pos_pos) simp_all
@@ -1110,7 +1110,7 @@
              pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
       by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
     also from n have "\<dots> = ?f n * rGamma_series z n"
-      by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def add_ac)
+      by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def)
     finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
   qed
   moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
@@ -1175,7 +1175,7 @@
   have "1/2 > (0::real)" by simp
   from tendstoD[OF assms, OF this]
     show "eventually (\<lambda>n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially"
-    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
+    by (force elim!: eventually_mono simp: dist_real_def)
   from assms have "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> 1 * Gamma z"
     by (intro tendsto_intros)
   thus "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> Gamma z" by simp
@@ -1189,7 +1189,7 @@
 proof (rule Lim_transform_eventually)
   have "1/2 > (0::real)" by simp
   from tendstoD[OF assms(2), OF this] show "eventually (\<lambda>n. g n * f n / f n = g n) sequentially"
-    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
+    by (force elim!: eventually_mono simp: dist_real_def)
   from assms have "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> 1 / rGamma z"
     by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff)
   thus "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> Gamma z" by (simp add: Gamma_def divide_inverse)
@@ -1409,13 +1409,13 @@
                        \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma z +
               rGamma z * d * (y - z)) /\<^sub>R  cmod (y - z)) \<midarrow>z\<rightarrow> 0"
       by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
-                    netlimit_at of_real_def[symmetric] suminf_def)
+                    of_real_def[symmetric] suminf_def)
 next
   fix n :: nat
   from has_field_derivative_rGamma_complex_nonpos_Int[of n]
   show "let z = - of_nat n in (\<lambda>y. (rGamma y - rGamma z - (- 1) ^ n * prod of_nat {1..n} *
                   (y - z)) /\<^sub>R cmod (y - z)) \<midarrow>z\<rightarrow> 0"
-    by (simp add: has_field_derivative_def has_derivative_def fact_prod netlimit_at Let_def)
+    by (simp add: has_field_derivative_def has_derivative_def fact_prod Let_def)
 next
   fix z :: complex
   from rGamma_series_complex_converges[of z] have "rGamma_series z \<longlonglongrightarrow> rGamma z"
@@ -1535,7 +1535,7 @@
   fix n :: nat assume n: "n > 0"
   have "Re (rGamma_series (of_real x) n) =
           Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))"
-    using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real)
+    using n by (simp add: rGamma_series_def powr_def pochhammer_of_real)
   also from n have "\<dots> = Re (of_real ((pochhammer x (Suc n)) /
                               (fact n * (exp (x * ln (real_of_nat n))))))"
     by (subst exp_of_real) simp
@@ -1569,7 +1569,7 @@
                        \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma x +
               rGamma x * d * (y - x)) /\<^sub>R  norm (y - x)) \<midarrow>x\<rightarrow> 0"
       by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
-                    netlimit_at of_real_def[symmetric] suminf_def)
+                    of_real_def[symmetric] suminf_def)
 next
   fix n :: nat
   have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))"
@@ -1577,7 +1577,7 @@
                   simp: Polygamma_of_real rGamma_real_def [abs_def])
   thus "let x = - of_nat n in (\<lambda>y. (rGamma y - rGamma x - (- 1) ^ n * prod of_nat {1..n} *
                   (y - x)) /\<^sub>R norm (y - x)) \<midarrow>x::real\<rightarrow> 0"
-    by (simp add: has_field_derivative_def has_derivative_def fact_prod netlimit_at Let_def)
+    by (simp add: has_field_derivative_def has_derivative_def fact_prod Let_def)
 next
   fix x :: real
   have "rGamma_series x \<longlonglongrightarrow> rGamma x"
@@ -1619,7 +1619,7 @@
   assume xn: "x > 0" "n > 0"
   have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \<ge> 1" for k
     using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps)
-  with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real)
+  with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_real)
 qed
 
 lemma ln_Gamma_real_converges:
@@ -2061,7 +2061,7 @@
     by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all
   with lim[of "1/2 :: 'a"] have "?h \<longlonglongrightarrow> 2 * Gamma (1/2 :: 'a)" by (simp add: exp_of_real)
   from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis
-    by (simp add: field_split_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps)
+    by (simp add: field_split_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real)
 qed
 
 text \<open>
@@ -2126,13 +2126,13 @@
 
   have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t
   proof -
-    from that have t: "t \<in> \<int> \<longleftrightarrow> t = 0" by (auto elim!: Ints_cases simp: dist_0_norm)
+    from that have t: "t \<in> \<int> \<longleftrightarrow> t = 0" by (auto elim!: Ints_cases)
     hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)"
       unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def)
     also have "a*cot (a*t) - 1/t = (F t) / (G t)"
       using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def)
     also have "\<dots> = (\<Sum>n. f n * (a*t)^n) / (\<Sum>n. g n * (a*t)^n)"
-      using sums[of t] that by (simp add: sums_iff dist_0_norm)
+      using sums[of t] that by (simp add: sums_iff)
     finally show "h t = h2 t" by (simp only: h2_def)
   qed
 
@@ -2167,7 +2167,7 @@
     have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add)
     have "eventually (\<lambda>z. h z = h2 z) (nhds z)"
       using eventually_nhds_in_nhd[of z ?A] using h_eq z
-      by (auto elim!: eventually_mono simp: dist_0_norm)
+      by (auto elim!: eventually_mono)
 
     moreover from sums(2)[OF z] z have nz: "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0"
       unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def)
@@ -2326,7 +2326,7 @@
         show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z"
         proof (cases "z = 0")
           assume z': "z \<noteq> 0"
-          with z have z'': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z \<notin> \<int>" by (auto elim!: Ints_cases simp: dist_0_norm)
+          with z have z'': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z \<notin> \<int>" by (auto elim!: Ints_cases)
           from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp
           with z'' z' show ?thesis by (simp add: g_def ac_simps)
         qed (simp add: g_def)
@@ -2471,7 +2471,7 @@
       show "\<forall>z\<in>B. norm (h' z) \<le> M/2"
       proof
         fix t :: complex assume t: "t \<in> B"
-        from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm)
+        from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp)
         also have "norm \<dots> = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp
         also have "norm (h' (t/2) + h' ((t+1)/2)) \<le> norm (h' (t/2)) + norm (h' ((t+1)/2))"
           by (rule norm_triangle_ineq)
@@ -2482,7 +2482,7 @@
         also have "(M + M) / 4 = M / 2" by simp
         finally show "norm (h' t) \<le> M/2" by - simp_all
       qed
-    qed (insert bdd, auto simp: cball_eq_empty)
+    qed (insert bdd, auto)
     hence "M \<le> 0" by simp
     finally show "h' z = 0" by simp
   qed
@@ -2512,7 +2512,7 @@
   with c have A: "h z * g z = 0" for z by simp
   hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp
   hence "\<exists>c'. \<forall>z\<in>UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all
-  then obtain c' where c: "\<And>z. g z = c'" by (force simp: dist_0_norm)
+  then obtain c' where c: "\<And>z. g z = c'" by (force)
   from this[of 0] have "c' = pi" unfolding g_def by simp
   with c have "g z = pi" by simp
 
@@ -2547,7 +2547,7 @@
 
 lemma Gamma_reflection_complex':
   "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))"
-  using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def field_split_simps mult_ac)
+  using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def field_split_simps)
 
 
 
@@ -2675,7 +2675,7 @@
   from z have z': "z \<noteq> 0" by auto
 
   have "eventually (\<lambda>n. ?r' n = ?r n) sequentially"
-    using z by (auto simp: field_split_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac
+    using z by (auto simp: field_split_simps Gamma_series_def ring_distribs exp_diff ln_div
                      intro: eventually_mono eventually_gt_at_top[of "0::nat"] dest: pochhammer_eq_0_imp_nonpos_Int)
   moreover have "?r' \<longlonglongrightarrow> exp (z * of_real (ln 1))"
     by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all
@@ -2819,7 +2819,7 @@
   have "(\<lambda>n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n))))
             \<longlonglongrightarrow> 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \<longlonglongrightarrow> _")
     using Gamma_gbinomial[of "of_nat k :: 'a"]
-    by (simp add: binomial_gbinomial add_ac Gamma_def field_split_simps exp_of_real [symmetric] exp_minus)
+    by (simp add: binomial_gbinomial Gamma_def field_split_simps exp_of_real [symmetric] exp_minus)
   also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact)
   finally show "?f \<longlonglongrightarrow> 1 / fact k" .
 
@@ -2862,7 +2862,7 @@
     next
       case False
       with \<open>z = 0\<close> show ?thesis
-        by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff gbinomial_1)
+        by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff)
     qed
   next
     case False
@@ -2885,7 +2885,7 @@
   have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))"
     by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac)
   also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)"
-    using Gamma_plus1[of "z+1"] by (auto simp add: field_split_simps mult_ac add_ac)
+    using Gamma_plus1[of "z+1"] by (auto simp add: field_split_simps)
   finally show ?thesis .
 qed
 
@@ -3073,7 +3073,7 @@
   have "eventually (\<lambda>n. Gamma_series z n =
           of_nat n powr z * fact n / pochhammer z (n+1)) sequentially"
     using eventually_gt_at_top[of "0::nat"]
-    by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def)
+    by eventually_elim (simp add: powr_def algebra_simps Gamma_series_def)
   from this and Gamma_series_LIMSEQ[of z]
     have C: "(\<lambda>k. of_nat k powr z * fact k / pochhammer z (k+1)) \<longlonglongrightarrow> Gamma z"
     by (blast intro: Lim_transform_eventually)
@@ -3441,7 +3441,7 @@
       fix n :: nat and x :: real assume x: "x \<in> ball 0 1"
       {
         fix k :: nat assume k: "k \<ge> 1"
-        from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1)
+        from x have "x^2 < 1" by (auto simp: abs_square_less_1)
         also from k have "\<dots> \<le> of_nat k^2" by simp
         finally have "(1 - x^2 / of_nat k^2) \<in> {0..1}" using k
           by (simp_all add: field_simps del: of_nat_Suc)
@@ -3470,7 +3470,7 @@
         by (simp add: eval_nat_numeral)
       from sums_divide[OF this, of "x^3 * pi"] x
         have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)"
-        by (simp add: field_split_simps eval_nat_numeral power_mult_distrib mult_ac)
+        by (simp add: field_split_simps eval_nat_numeral)
       with x have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)"
         by (simp add: g_def)
       hence "f' x = g x / x^2" by (simp add: sums_iff f'_def)
@@ -3485,7 +3485,7 @@
         assume x: "x \<noteq> 0"
         from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF
                sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x
-          show ?thesis by (simp add: mult_ac power_mult_distrib field_split_simps eval_nat_numeral)
+          show ?thesis by (simp add: field_split_simps eval_nat_numeral)
       qed (simp only: summable_0_powser)
     qed
     hence "f' \<midarrow> 0 \<rightarrow> f' 0" by (simp add: isCont_def)
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1526,7 +1526,7 @@
     apply (rule sum.cong)
     using assms
     apply simp
-    apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4))
+    apply (metis abs_of_nonneg content_pos_le)
     done
   have e: "0 \<le> e"
     using assms(2) norm_ge_zero order_trans by blast
--- a/src/HOL/Analysis/Homeomorphism.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Homeomorphism.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1993,7 +1993,7 @@
       have "q' t = (h \<circ> (*\<^sub>R) 2) t" if "0 \<le> t" "t \<le> 1/2" for t
       proof (rule covering_space_lift_unique [OF cov, of q' 0 "h \<circ> (*\<^sub>R) 2" "{0..1/2}" "f \<circ> g \<circ> (*\<^sub>R) 2" t])
         show "q' 0 = (h \<circ> (*\<^sub>R) 2) 0"
-          by (metis \<open>pathstart q' = pathstart q\<close> comp_def h pastq pathstart_def pth_4(2))
+          by (metis \<open>pathstart q' = pathstart q\<close> comp_def h(3) pastq pathstart_def pth_4(2))
         show "continuous_on {0..1/2} (f \<circ> g \<circ> (*\<^sub>R) 2)"
           apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \<open>path g\<close>] continuous_on_subset [OF contf])
           using g(2) path_image_def by fastforce+
--- a/src/HOL/Analysis/Homotopy.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Homotopy.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -11,7 +11,7 @@
 definition\<^marker>\<open>tag important\<close> homotopic_with
 where
  "homotopic_with P X Y f g \<equiv>
-   (\<exists>h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \<and>
+   (\<exists>h. continuous_map (prod_topology (top_of_set {0..1::real}) X) Y h \<and>
        (\<forall>x. h(0, x) = f x) \<and>
        (\<forall>x. h(1, x) = g x) \<and>
        (\<forall>t \<in> {0..1}. P(\<lambda>x. h(t,x))))"
@@ -46,9 +46,7 @@
 lemma continuous_map_o_Pair: 
   assumes h: "continuous_map (prod_topology X Y) Z h" and t: "t \<in> topspace X"
   shows "continuous_map Y Z (h \<circ> Pair t)"
-  apply (intro continuous_map_compose [OF _ h] continuous_map_id [unfolded id_def] continuous_intros)
-  apply (simp add: t)
-  done
+  by (intro continuous_map_compose [OF _ h] continuous_intros; simp add: t)
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Trivial properties\<close>
 
@@ -74,18 +72,15 @@
   assumes hom: "homotopic_with P X Y f g"
     and Q: "\<And>h. \<lbrakk>continuous_map X Y h; P h\<rbrakk> \<Longrightarrow> Q h"
   shows "homotopic_with Q X Y f g"
-  using hom
-  apply (simp add: homotopic_with_def)
-  apply (erule ex_forward)
-  apply (force simp: o_def dest: continuous_map_o_Pair intro: Q)
-  done
+  using hom unfolding homotopic_with_def
+  by (force simp: o_def dest: continuous_map_o_Pair intro: Q)
 
 lemma homotopic_with_imp_continuous_maps:
     assumes "homotopic_with P X Y f g"
     shows "continuous_map X Y f \<and> continuous_map X Y g"
 proof -
-  obtain h
-    where conth: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h"
+  obtain h :: "real \<times> 'a \<Rightarrow> 'b"
+    where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) Y h"
       and h: "\<forall>x. h (0, x) = f x" "\<forall>x. h (1, x) = g x"
     using assms by (auto simp: homotopic_with_def)
   have *: "t \<in> {0..1} \<Longrightarrow> continuous_map X Y (h \<circ> (\<lambda>x. (t,x)))" for t
@@ -115,9 +110,9 @@
   unfolding homotopic_with_def
 proof (intro exI conjI allI ballI)
   let ?h = "\<lambda>(t::real,x). if t = 1 then g x else f x"
-  show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y ?h"
+  show "continuous_map (prod_topology (top_of_set {0..1}) X) Y ?h"
   proof (rule continuous_map_eq)
-    show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y (f \<circ> snd)"
+    show "continuous_map (prod_topology (top_of_set {0..1}) X) Y (f \<circ> snd)"
       by (simp add: contf continuous_map_of_snd)
   qed (auto simp: fg)
   show "P (\<lambda>x. ?h (t, x))" if "t \<in> {0..1}" for t
@@ -134,15 +129,11 @@
 
 lemma homotopic_with_subset_left:
      "\<lbrakk>homotopic_with_canon P X Y f g; Z \<subseteq> X\<rbrakk> \<Longrightarrow> homotopic_with_canon P Z Y f g"
-  apply (simp add: homotopic_with_def)
-  apply (fast elim!: continuous_on_subset ex_forward)
-  done
+  unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward)
 
 lemma homotopic_with_subset_right:
      "\<lbrakk>homotopic_with_canon P X Y f g; Y \<subseteq> Z\<rbrakk> \<Longrightarrow> homotopic_with_canon P X Z f g"
-  apply (simp add: homotopic_with_def)
-  apply (fast elim!: continuous_on_subset ex_forward)
-  done
+  unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward)
 
 subsection\<open>Homotopy with P is an equivalence relation\<close>
 
@@ -158,9 +149,7 @@
   let ?I01 = "subtopology euclideanreal {0..1}"
   let ?j = "\<lambda>y. (1 - fst y, snd y)"
   have 1: "continuous_map (prod_topology ?I01 X) (prod_topology euclideanreal X) ?j"
-    apply (intro continuous_intros)
-    apply (simp_all add: prod_topology_subtopology continuous_map_from_subtopology [OF continuous_map_fst])
-    done
+    by (intro continuous_intros; simp add: continuous_map_subtopology_fst prod_topology_subtopology)
   have *: "continuous_map (prod_topology ?I01 X) (prod_topology ?I01 X) ?j"
   proof -
     have "continuous_map (prod_topology ?I01 X) (subtopology (prod_topology euclideanreal X) ({0..1} \<times> topspace X)) ?j"
@@ -171,9 +160,8 @@
   show ?thesis
     using assms
     apply (clarsimp simp add: homotopic_with_def)
-    apply (rename_tac h)
-    apply (rule_tac x="h \<circ> (\<lambda>y. (1 - fst y, snd y))" in exI)
-    apply (simp add: continuous_map_compose [OF *])
+    subgoal for h
+      by (rule_tac x="h \<circ> (\<lambda>y. (1 - fst y, snd y))" in exI) (simp add: continuous_map_compose [OF *])
     done
 qed
 
@@ -209,16 +197,12 @@
         by simp
       show "continuous_map (subtopology ?X01 {y \<in> topspace ?X01. fst y \<le> 1/2}) Y
                (k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x)))"
-        apply (rule fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology | simp)+
-          apply (intro continuous_intros fst continuous_map_from_subtopology)
-         apply (force simp: prod_topology_subtopology)
-        using continuous_map_snd continuous_map_from_subtopology by blast
+        apply (intro fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology fst continuous_map_from_subtopology | simp)+
+        by (force simp: prod_topology_subtopology)
       show "continuous_map (subtopology ?X01 {y \<in> topspace ?X01. 1/2 \<le> fst y}) Y
                (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x)))"
-        apply (rule fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology | simp)+
-          apply (rule continuous_intros fst continuous_map_from_subtopology | simp)+
-         apply (force simp: prod_topology_subtopology)
-        using continuous_map_snd  continuous_map_from_subtopology by blast
+        apply (intro fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology fst continuous_map_from_subtopology | simp)+
+        by (force simp: prod_topology_subtopology)
       show "(k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y = (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y"
         if "y \<in> topspace ?X01" and "fst y = 1/2" for y
         using that by (simp add: keq)
@@ -228,10 +212,10 @@
     show "\<forall>x. k (1, x) = h x"
       by (simp add: k12 k_def)
     show "\<forall>t\<in>{0..1}. P (\<lambda>x. k (t, x))"
-      using P
-      apply (clarsimp simp add: k_def)
-      apply (case_tac "t \<le> 1/2", auto)
-      done
+    proof 
+      fix t show "t\<in>{0..1} \<Longrightarrow> P (\<lambda>x. k (t, x))"
+        by (cases "t \<le> 1/2") (auto simp add: k_def P)
+    qed
   qed
 qed
 
@@ -246,16 +230,10 @@
    \<Longrightarrow> homotopic_with q X1 X3 (h \<circ> f) (h \<circ> g)"
   unfolding homotopic_with_def
   apply clarify
-  apply (rename_tac k)
-  apply (rule_tac x="h \<circ> k" in exI)
-  apply (rule conjI continuous_map_compose | simp add: o_def)+
+  subgoal for k
+    by (rule_tac x="h \<circ> k" in exI) (rule conjI continuous_map_compose | simp add: o_def)+
   done
 
-lemma homotopic_compose_continuous_map_left:
-   "\<lbrakk>homotopic_with (\<lambda>k. True) X1 X2 f g; continuous_map X2 X3 h\<rbrakk>
-        \<Longrightarrow> homotopic_with (\<lambda>k. True) X1 X3 (h \<circ> f) (h \<circ> g)"
-  by (simp add: homotopic_with_compose_continuous_map_left)
-
 lemma homotopic_with_compose_continuous_map_right:
   assumes hom: "homotopic_with p X2 X3 f g" and conth: "continuous_map X1 X2 h"
     and q: "\<And>j. p j \<Longrightarrow> q(j \<circ> h)"
@@ -281,19 +259,15 @@
   qed (auto simp: k)
 qed
 
-lemma homotopic_compose_continuous_map_right:
-   "\<lbrakk>homotopic_with (\<lambda>k. True) X2 X3 f g; continuous_map X1 X2 h\<rbrakk>
-        \<Longrightarrow> homotopic_with (\<lambda>k. True) X1 X3 (f \<circ> h) (g \<circ> h)"
-  by (meson homotopic_with_compose_continuous_map_right)
-
 corollary homotopic_compose:
-      shows "\<lbrakk>homotopic_with (\<lambda>x. True) X Y f f'; homotopic_with (\<lambda>x. True) Y Z g g'\<rbrakk>
-             \<Longrightarrow> homotopic_with (\<lambda>x. True) X Z (g \<circ> f) (g' \<circ> f')"
-  apply (rule homotopic_with_trans [where g = "g \<circ> f'"])
-  apply (simp add: homotopic_compose_continuous_map_left homotopic_with_imp_continuous_maps)
-  by (simp add: homotopic_compose_continuous_map_right homotopic_with_imp_continuous_maps)
-
-
+  assumes "homotopic_with (\<lambda>x. True) X Y f f'" "homotopic_with (\<lambda>x. True) Y Z g g'"
+  shows "homotopic_with (\<lambda>x. True) X Z (g \<circ> f) (g' \<circ> f')"
+proof (rule homotopic_with_trans [where g = "g \<circ> f'"])
+  show "homotopic_with (\<lambda>x. True) X Z (g \<circ> f) (g \<circ> f')"
+    using assms by (simp add: homotopic_with_compose_continuous_map_left homotopic_with_imp_continuous_maps)
+  show "homotopic_with (\<lambda>x. True) X Z (g \<circ> f') (g' \<circ> f')"
+    using assms by (simp add: homotopic_with_compose_continuous_map_right homotopic_with_imp_continuous_maps)
+qed
 
 proposition homotopic_with_compose_continuous_right:
     "\<lbrakk>homotopic_with_canon (\<lambda>f. p (f \<circ> h)) X Y f g; continuous_on W h; h ` W \<subseteq> X\<rbrakk>
@@ -301,16 +275,10 @@
   apply (clarsimp simp add: homotopic_with_def)
   apply (rename_tac k)
   apply (rule_tac x="k \<circ> (\<lambda>y. (fst y, h (snd y)))" in exI)
-  apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+
-  apply (erule continuous_on_subset)
-  apply (fastforce simp: o_def)+
+  apply (rule conjI continuous_intros continuous_on_compose2 [where f=snd and g=h] | simp)+
+  apply (fastforce simp: o_def elim: continuous_on_subset)+
   done
 
-proposition homotopic_compose_continuous_right:
-     "\<lbrakk>homotopic_with_canon (\<lambda>f. True) X Y f g; continuous_on W h; h ` W \<subseteq> X\<rbrakk>
-      \<Longrightarrow> homotopic_with_canon (\<lambda>f. True) W Y (f \<circ> h) (g \<circ> h)"
-  using homotopic_with_compose_continuous_right by fastforce
-
 proposition homotopic_with_compose_continuous_left:
      "\<lbrakk>homotopic_with_canon (\<lambda>f. p (h \<circ> f)) X Y f g; continuous_on Y h; h ` Y \<subseteq> Z\<rbrakk>
       \<Longrightarrow> homotopic_with_canon p X Z (h \<circ> f) (h \<circ> g)"
@@ -318,21 +286,13 @@
   apply (rename_tac k)
   apply (rule_tac x="h \<circ> k" in exI)
   apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+
-  apply (erule continuous_on_subset)
-  apply (fastforce simp: o_def)+
+  apply (fastforce simp: o_def elim: continuous_on_subset)+
   done
 
-proposition homotopic_compose_continuous_left:
-   "\<lbrakk>homotopic_with_canon (\<lambda>_. True) X Y f g;
-     continuous_on Y h; h ` Y \<subseteq> Z\<rbrakk>
-    \<Longrightarrow> homotopic_with_canon (\<lambda>f. True) X Z (h \<circ> f) (h \<circ> g)"
-  using homotopic_with_compose_continuous_left by fastforce
-
 lemma homotopic_from_subtopology:
    "homotopic_with P X X' f g \<Longrightarrow> homotopic_with P (subtopology X s) X' f g"
   unfolding homotopic_with_def
-  apply (erule ex_forward)
-  by (simp add: continuous_map_from_subtopology prod_topology_subtopology(2))
+  by (force simp add: continuous_map_from_subtopology prod_topology_subtopology(2) elim!: ex_forward)
 
 lemma homotopic_on_emptyI:
     assumes "topspace X = {}" "P f" "P g"
@@ -365,9 +325,7 @@
         and h: "\<And>x. h (0, x) = a" "\<And>x. h (1, x) = b"
       using hom by (auto simp: homotopic_with_def)
     have cont: "continuous_map (top_of_set {0..1}) X' (h \<circ> (\<lambda>t. (t, c)))"
-      apply (rule continuous_map_compose [OF _ conth])
-      apply (rule continuous_intros c | simp)+
-      done
+      by (rule continuous_map_compose [OF _ conth] continuous_intros c | simp)+
     then show ?thesis
       by (force simp: h)
   qed
@@ -449,10 +407,13 @@
     let ?h = "\<lambda>(t,z). \<lambda>i\<in>I. h i (t,z i)"
     have "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I))
                          (Y i) (\<lambda>x. h i (fst x, snd x i))" if "i \<in> I" for i
-      unfolding continuous_map_pairwise case_prod_unfold
-      apply (intro conjI that  continuous_intros continuous_map_compose [OF _ h, unfolded o_def])
-      using continuous_map_componentwise continuous_map_snd that apply fastforce
-      done
+    proof -
+      have \<section>: "continuous_map (prod_topology (top_of_set {0..1}) (product_topology X I)) (X i) (\<lambda>x. snd x i)"
+        using continuous_map_componentwise continuous_map_snd that by fastforce
+      show ?thesis
+        unfolding continuous_map_pairwise case_prod_unfold
+        by (intro conjI that \<section> continuous_intros continuous_map_compose [OF _ h, unfolded o_def])
+    qed
     then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I))
          (product_topology Y I) ?h"
       by (auto simp: continuous_map_componentwise case_prod_beta)
@@ -570,10 +531,9 @@
 
 proposition homotopic_paths_eq:
      "\<lbrakk>path p; path_image p \<subseteq> s; \<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t\<rbrakk> \<Longrightarrow> homotopic_paths s p q"
-  apply (simp add: homotopic_paths_def)
-  apply (rule homotopic_with_eq)
-  apply (auto simp: path_def pathstart_def pathfinish_def path_image_def elim: continuous_on_eq)
-  done
+  unfolding homotopic_paths_def
+  by (rule homotopic_with_eq)
+     (auto simp: path_def pathstart_def pathfinish_def path_image_def elim: continuous_on_eq)
 
 proposition homotopic_paths_reparametrize:
   assumes "path p"
@@ -618,25 +578,25 @@
 
 
 text\<open> A slightly ad-hoc but useful lemma in constructing homotopies.\<close>
-lemma homotopic_join_lemma:
+lemma continuous_on_homotopic_join_lemma:
   fixes q :: "[real,real] \<Rightarrow> 'a::topological_space"
-  assumes p: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. p (fst y) (snd y))"
-      and q: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. q (fst y) (snd y))"
+  assumes p: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. p (fst y) (snd y))" (is "continuous_on ?A ?p")
+      and q: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. q (fst y) (snd y))" (is "continuous_on ?A ?q")
       and pf: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish(p t) = pathstart(q t)"
     shows "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. (p(fst y) +++ q(fst y)) (snd y))"
 proof -
-  have 1: "(\<lambda>y. p (fst y) (2 * snd y)) = (\<lambda>y. p (fst y) (snd y)) \<circ> (\<lambda>y. (fst y, 2 * snd y))"
-    by (rule ext) (simp)
-  have 2: "(\<lambda>y. q (fst y) (2 * snd y - 1)) = (\<lambda>y. q (fst y) (snd y)) \<circ> (\<lambda>y. (fst y, 2 * snd y - 1))"
-    by (rule ext) (simp)
+  have \<section>: "(\<lambda>t. p (fst t) (2 * snd t)) = ?p \<circ> (\<lambda>y. (fst y, 2 * snd y))"
+          "(\<lambda>t. q (fst t) (2 * snd t - 1)) = ?q \<circ> (\<lambda>y. (fst y, 2 * snd y - 1))"
+    by force+
   show ?thesis
-    apply (simp add: joinpaths_def)
-    apply (rule continuous_on_cases_le)
-    apply (simp_all only: 1 2)
-    apply (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+
-    using pf
-    apply (auto simp: mult.commute pathstart_def pathfinish_def)
-    done
+    unfolding joinpaths_def
+  proof (rule continuous_on_cases_le)
+    show "continuous_on {y \<in> ?A. snd y \<le> 1/2} (\<lambda>t. p (fst t) (2 * snd t))" 
+         "continuous_on {y \<in> ?A. 1/2 \<le> snd y} (\<lambda>t. q (fst t) (2 * snd t - 1))"
+         "continuous_on ?A snd"
+      unfolding \<section>
+      by (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+
+  qed (use pf in \<open>auto simp: mult.commute pathstart_def pathfinish_def\<close>)
 qed
 
 text\<open> Congruence properties of homotopy w.r.t. path-combining operations.\<close>
@@ -658,11 +618,10 @@
 
 proposition homotopic_paths_join:
     "\<lbrakk>homotopic_paths s p p'; homotopic_paths s q q'; pathfinish p = pathstart q\<rbrakk> \<Longrightarrow> homotopic_paths s (p +++ q) (p' +++ q')"
-  apply (simp add: homotopic_paths_def homotopic_with_def, clarify)
+  apply (clarsimp simp add: homotopic_paths_def homotopic_with_def)
   apply (rename_tac k1 k2)
   apply (rule_tac x="(\<lambda>y. ((k1 \<circ> Pair (fst y)) +++ (k2 \<circ> Pair (fst y))) (snd y))" in exI)
-  apply (rule conjI continuous_intros homotopic_join_lemma)+
-  apply (auto simp: joinpaths_def pathstart_def pathfinish_def path_image_def)
+  apply (intro conjI continuous_intros continuous_on_homotopic_join_lemma; force simp: joinpaths_def pathstart_def pathfinish_def path_image_def)
   done
 
 proposition homotopic_paths_continuous_image:
@@ -676,13 +635,19 @@
 text\<^marker>\<open>tag important\<close>\<open>So taking equivalence classes under homotopy would give the fundamental group\<close>
 
 proposition homotopic_paths_rid:
-    "\<lbrakk>path p; path_image p \<subseteq> s\<rbrakk> \<Longrightarrow> homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p"
-  apply (subst homotopic_paths_sym)
-  apply (rule homotopic_paths_reparametrize [where f = "\<lambda>t. if  t \<le> 1 / 2 then 2 *\<^sub>R t else 1"])
-  apply (simp_all del: le_divide_eq_numeral1)
-  apply (subst split_01)
-  apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+
-  done
+  assumes "path p" "path_image p \<subseteq> s"
+  shows "homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p"
+proof -
+  have \<section>: "continuous_on {0..1} (\<lambda>t::real. if t \<le> 1/2 then 2 *\<^sub>R t else 1)"
+    unfolding split_01
+    by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+
+  show ?thesis
+    using assms
+    apply (subst homotopic_paths_sym)
+     apply (rule homotopic_paths_reparametrize [where f = "\<lambda>t. if t \<le> 1/2 then 2 *\<^sub>R t else 1"])
+           apply (rule \<section> continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+
+    done
+qed
 
 proposition homotopic_paths_lid:
    "\<lbrakk>path p; path_image p \<subseteq> s\<rbrakk> \<Longrightarrow> homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p) p"
@@ -696,35 +661,35 @@
     \<Longrightarrow> homotopic_paths s (p +++ (q +++ r)) ((p +++ q) +++ r)"
   apply (subst homotopic_paths_sym)
   apply (rule homotopic_paths_reparametrize
-           [where f = "\<lambda>t. if  t \<le> 1 / 2 then inverse 2 *\<^sub>R t
+           [where f = "\<lambda>t. if  t \<le> 1/2 then inverse 2 *\<^sub>R t
                            else if  t \<le> 3 / 4 then t - (1 / 4)
                            else 2 *\<^sub>R t - 1"])
   apply (simp_all del: le_divide_eq_numeral1)
   apply (simp add: subset_path_image_join)
-  apply (rule continuous_on_cases_1 continuous_intros)+
-  apply (auto simp: joinpaths_def)
+  apply (rule continuous_on_cases_1 continuous_intros | auto simp: joinpaths_def)+
   done
 
 proposition homotopic_paths_rinv:
   assumes "path p" "path_image p \<subseteq> s"
     shows "homotopic_paths s (p +++ reversepath p) (linepath (pathstart p) (pathstart p))"
 proof -
-  have "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))"
-    using assms
-    apply (simp add: joinpaths_def subpath_def reversepath_def path_def del: le_divide_eq_numeral1)
-    apply (rule continuous_on_cases_le)
-    apply (rule_tac [2] continuous_on_compose [of _ _ p, unfolded o_def])
-    apply (rule continuous_on_compose [of _ _ p, unfolded o_def])
-    apply (auto intro!: continuous_intros simp del: eq_divide_eq_numeral1)
-    apply (force elim!: continuous_on_subset simp add: mult_le_one)+
-    done
+  have p: "continuous_on {0..1} p" 
+    using assms by (auto simp: path_def)
+  let ?A = "{0..1} \<times> {0..1}"
+  have "continuous_on ?A (\<lambda>x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))"
+    unfolding joinpaths_def subpath_def reversepath_def path_def add_0_right diff_0_right
+  proof (rule continuous_on_cases_le)
+    show "continuous_on {x \<in> ?A. snd x \<le> 1/2} (\<lambda>t. p (fst t * (2 * snd t)))"
+         "continuous_on {x \<in> ?A. 1/2 \<le> snd x} (\<lambda>t. p (fst t * (1 - (2 * snd t - 1))))"
+         "continuous_on ?A snd"
+      by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp add: mult_le_one)+
+  qed (auto simp add: algebra_simps)
   then show ?thesis
     using assms
     apply (subst homotopic_paths_sym_eq)
     unfolding homotopic_paths_def homotopic_with_def
     apply (rule_tac x="(\<lambda>y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI)
-    apply (simp add: path_defs joinpaths_def subpath_def reversepath_def)
-    apply (force simp: mult_le_one)
+    apply (force simp: mult_le_one path_defs joinpaths_def subpath_def reversepath_def)
     done
 qed
 
@@ -786,8 +751,7 @@
    "\<lbrakk>path p; path_image p \<subseteq> s; pathfinish p = pathstart p; \<And>t. t \<in> {0..1} \<Longrightarrow> p(t) = q(t)\<rbrakk>
           \<Longrightarrow> homotopic_loops s p q"
   unfolding homotopic_loops_def
-  apply (rule homotopic_with_eq)
-  apply (rule homotopic_with_refl [where f = p, THEN iffD2])
+  apply (rule homotopic_with_eq [OF homotopic_with_refl [where f = p, THEN iffD2]])
   apply (simp_all add: path_image_def path_def pathstart_def pathfinish_def)
   done
 
@@ -810,29 +774,31 @@
   have "path p" by (metis assms homotopic_loops_imp_path)
   have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop)
   have pip: "path_image p \<subseteq> s" by (metis assms homotopic_loops_imp_subset)
-  obtain h where conth: "continuous_on ({0..1::real} \<times> {0..1}) h"
-             and hs: "h ` ({0..1} \<times> {0..1}) \<subseteq> s"
+  let ?A = "{0..1::real} \<times> {0..1::real}"
+  obtain h where conth: "continuous_on ?A h"
+             and hs: "h ` ?A \<subseteq> s"
              and [simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(0,x) = p x"
              and [simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(1,x) = a"
              and ends: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish (h \<circ> Pair t) = pathstart (h \<circ> Pair t)"
     using assms by (auto simp: homotopic_loops homotopic_with)
   have conth0: "path (\<lambda>u. h (u, 0))"
     unfolding path_def
-    apply (rule continuous_on_compose [of _ _ h, unfolded o_def])
-    apply (force intro: continuous_intros continuous_on_subset [OF conth])+
-    done
+  proof (rule continuous_on_compose [of _ _ h, unfolded o_def])
+    show "continuous_on ((\<lambda>x. (x, 0)) ` {0..1}) h"
+      by (force intro: continuous_on_subset [OF conth])
+  qed (force intro: continuous_intros)
   have pih0: "path_image (\<lambda>u. h (u, 0)) \<subseteq> s"
     using hs by (force simp: path_image_def)
-  have c1: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. h (fst x * snd x, 0))"
-    apply (rule continuous_on_compose [of _ _ h, unfolded o_def])
-    apply (force simp: mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+
-    done
-  have c2: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. h (fst x - fst x * snd x, 0))"
-    apply (rule continuous_on_compose [of _ _ h, unfolded o_def])
-    apply (force simp: mult_left_le mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+
-    apply (rule continuous_on_subset [OF conth])
-    apply (auto simp: algebra_simps add_increasing2 mult_left_le)
-    done
+  have c1: "continuous_on ?A (\<lambda>x. h (fst x * snd x, 0))"
+  proof (rule continuous_on_compose [of _ _ h, unfolded o_def])
+    show "continuous_on ((\<lambda>x. (fst x * snd x, 0)) ` ?A) h"
+      by (force simp: mult_le_one intro: continuous_on_subset [OF conth])
+  qed (force intro: continuous_intros)+
+  have c2: "continuous_on ?A (\<lambda>x. h (fst x - fst x * snd x, 0))"
+  proof (rule continuous_on_compose [of _ _ h, unfolded o_def])
+    show "continuous_on ((\<lambda>x. (fst x - fst x * snd x, 0)) ` ?A) h"
+      by (auto simp: algebra_simps add_increasing2 mult_left_le intro: continuous_on_subset [OF conth])
+  qed (force intro: continuous_intros)
   have [simp]: "\<And>t. \<lbrakk>0 \<le> t \<and> t \<le> 1\<rbrakk> \<Longrightarrow> h (t, 1) = h (t, 0)"
     using ends by (simp add: pathfinish_def pathstart_def)
   have adhoc_le: "c * 4 \<le> 1 + c * (d * 4)" if "\<not> d * 4 \<le> 3" "0 \<le> c" "c \<le> 1" for c d::real
@@ -840,10 +806,10 @@
     have "c * 3 \<le> c * (d * 4)" using that less_eq_real_def by auto
     with \<open>c \<le> 1\<close> show ?thesis by fastforce
   qed
-  have *: "\<And>p x. (path p \<and> path(reversepath p)) \<and>
-                  (path_image p \<subseteq> s \<and> path_image(reversepath p) \<subseteq> s) \<and>
-                  (pathfinish p = pathstart(linepath a a +++ reversepath p) \<and>
-                   pathstart(reversepath p) = a) \<and> pathstart p = x
+  have *: "\<And>p x. \<lbrakk>path p \<and> path(reversepath p);
+                  path_image p \<subseteq> s \<and> path_image(reversepath p) \<subseteq> s;
+                  pathfinish p = pathstart(linepath a a +++ reversepath p) \<and>
+                   pathstart(reversepath p) = a \<and> pathstart p = x\<rbrakk>
                   \<Longrightarrow> homotopic_paths s (p +++ linepath a a +++ reversepath p) (linepath x x)"
     by (metis homotopic_paths_lid homotopic_paths_join
               homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv)
@@ -854,22 +820,28 @@
     apply (rule homotopic_paths_sym)
     using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" s]
     by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_pathstart homotopic_paths_imp_subset)
-  moreover have "homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))
+  moreover 
+  have "homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))
                                    ((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0)))"
-    apply (simp add: homotopic_paths_def homotopic_with_def)
-    apply (rule_tac x="\<lambda>y. (subpath 0 (fst y) (\<lambda>u. h (u, 0)) +++ (\<lambda>u. h (Pair (fst y) u)) +++ subpath (fst y) 0 (\<lambda>u. h (u, 0))) (snd y)" in exI)
-    apply (simp add: subpath_reversepath)
-    apply (intro conjI homotopic_join_lemma)
-    using ploop
-    apply (simp_all add: path_defs joinpaths_def o_def subpath_def conth c1 c2)
-    apply (force simp: algebra_simps mult_le_one mult_left_le intro: hs [THEN subsetD] adhoc_le)
-    done
+    unfolding homotopic_paths_def homotopic_with_def
+  proof (intro exI strip conjI)
+    let ?h = "\<lambda>y. (subpath 0 (fst y) (\<lambda>u. h (u, 0)) +++ (\<lambda>u. h (Pair (fst y) u)) 
+               +++ subpath (fst y) 0 (\<lambda>u. h (u, 0))) (snd y)" 
+    have "continuous_on ?A ?h"
+      by (intro continuous_on_homotopic_join_lemma; simp add: path_defs joinpaths_def subpath_def conth c1 c2)
+    moreover have "?h ` ?A \<subseteq> s"
+      unfolding joinpaths_def subpath_def
+      by (force simp: algebra_simps mult_le_one mult_left_le intro: hs [THEN subsetD] adhoc_le)
+  ultimately show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1}))
+                         (top_of_set s) ?h"
+    by (simp add: subpath_reversepath)
+  qed (use ploop in \<open>simp_all add: reversepath_def path_defs joinpaths_def o_def subpath_def conth c1 c2\<close>)
   moreover have "homotopic_paths s ((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0)))
                                    (linepath (pathstart p) (pathstart p))"
-    apply (rule *)
-    apply (simp add: pih0 pathstart_def pathfinish_def conth0)
-    apply (simp add: reversepath_def joinpaths_def)
-    done
+  proof (rule *; simp add: pih0 pathstart_def pathfinish_def conth0)
+    show "a = (linepath a a +++ reversepath (\<lambda>u. h (u, 0))) 0 \<and> reversepath (\<lambda>u. h (u, 0)) 0 = a"
+      by (simp_all add: reversepath_def joinpaths_def)
+  qed
   ultimately show ?thesis
     by (blast intro: homotopic_paths_trans)
 qed
@@ -877,23 +849,23 @@
 proposition homotopic_loops_conjugate:
   fixes s :: "'a::real_normed_vector set"
   assumes "path p" "path q" and pip: "path_image p \<subseteq> s" and piq: "path_image q \<subseteq> s"
-      and papp: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q"
+      and pq: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q"
     shows "homotopic_loops s (p +++ q +++ reversepath p) q"
 proof -
   have contp: "continuous_on {0..1} p"  using \<open>path p\<close> [unfolded path_def] by blast
   have contq: "continuous_on {0..1} q"  using \<open>path q\<close> [unfolded path_def] by blast
-  have c1: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. p ((1 - fst x) * snd x + fst x))"
-    apply (rule continuous_on_compose [of _ _ p, unfolded o_def])
-    apply (force simp: mult_le_one intro!: continuous_intros)
-    apply (rule continuous_on_subset [OF contp])
-    apply (auto simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1)
-    done
-  have c2: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. p ((fst x - 1) * snd x + 1))"
-    apply (rule continuous_on_compose [of _ _ p, unfolded o_def])
-    apply (force simp: mult_le_one intro!: continuous_intros)
-    apply (rule continuous_on_subset [OF contp])
-    apply (auto simp: algebra_simps add_increasing2 mult_left_le_one_le)
-    done
+  let ?A = "{0..1::real} \<times> {0..1::real}"
+  have c1: "continuous_on ?A (\<lambda>x. p ((1 - fst x) * snd x + fst x))"
+  proof (rule continuous_on_compose [of _ _ p, unfolded o_def])
+    show "continuous_on ((\<lambda>x. (1 - fst x) * snd x + fst x) ` ?A) p"
+      by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1)
+  qed (force intro: continuous_intros)
+  have c2: "continuous_on ?A (\<lambda>x. p ((fst x - 1) * snd x + 1))"
+  proof (rule continuous_on_compose [of _ _ p, unfolded o_def])
+    show "continuous_on ((\<lambda>x. (fst x - 1) * snd x + 1) ` ?A) p"
+      by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_left_le_one_le)
+  qed (force intro: continuous_intros)
+
   have ps1: "\<And>a b. \<lbrakk>b * 2 \<le> 1; 0 \<le> b; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((1 - a) * (2 * b) + a) \<in> s"
     using sum_le_prod1
     by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD])
@@ -901,21 +873,26 @@
     apply (rule pip [unfolded path_image_def, THEN subsetD])
     apply (rule image_eqI, blast)
     apply (simp add: algebra_simps)
-    by (metis add_mono_thms_linordered_semiring(1) affine_ineq linear mult.commute mult.left_neutral mult_right_mono not_le
+    by (metis add_mono_thms_linordered_semiring(1) affine_ineq linear mult.commute mult.left_neutral mult_right_mono
               add.commute zero_le_numeral)
   have qs: "\<And>a b. \<lbrakk>4 * b \<le> 3; \<not> b * 2 \<le> 1\<rbrakk> \<Longrightarrow> q (4 * b - 2) \<in> s"
     using path_image_def piq by fastforce
   have "homotopic_loops s (p +++ q +++ reversepath p)
                           (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))"
-    apply (simp add: homotopic_loops_def homotopic_with_def)
-    apply (rule_tac x="(\<lambda>y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" in exI)
-    apply (simp add: subpath_refl subpath_reversepath)
-    apply (intro conjI homotopic_join_lemma)
-    using papp qloop
-    apply (simp_all add: path_defs joinpaths_def o_def subpath_def c1 c2)
-    apply (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd)
-    apply (auto simp: ps1 ps2 qs)
-    done
+    unfolding homotopic_loops_def homotopic_with_def
+  proof (intro exI strip conjI)
+    let ?h = "(\<lambda>y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" 
+    have "continuous_on ?A (\<lambda>y. q (snd y))"
+      by (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd)
+    then have "continuous_on ?A ?h"
+      apply (intro continuous_on_homotopic_join_lemma)
+      using pq qloop
+      by (auto simp: path_defs joinpaths_def subpath_def c1 c2)
+    then show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1})) (top_of_set s) ?h"
+      by (auto simp: joinpaths_def subpath_def  ps1 ps2 qs)
+    show "?h (1,x) = (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) x"  for x
+      using pq by (simp add: pathfinish_def subpath_refl)
+  qed (auto simp: subpath_reversepath)
   moreover have "homotopic_loops s (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q"
   proof -
     have "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q) q"
@@ -975,10 +952,10 @@
 
 lemma homotopic_with_linear:
   fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
-  assumes contf: "continuous_on s f"
-      and contg:"continuous_on s g"
-      and sub: "\<And>x. x \<in> s \<Longrightarrow> closed_segment (f x) (g x) \<subseteq> t"
-    shows "homotopic_with_canon (\<lambda>z. True) s t f g"
+  assumes contf: "continuous_on S f"
+      and contg:"continuous_on S g"
+      and sub: "\<And>x. x \<in> S \<Longrightarrow> closed_segment (f x) (g x) \<subseteq> t"
+    shows "homotopic_with_canon (\<lambda>z. True) S t f g"
   apply (simp add: homotopic_with_def)
   apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI)
   apply (intro conjI)
@@ -990,8 +967,8 @@
 lemma homotopic_paths_linear:
   fixes g h :: "real \<Rightarrow> 'a::real_normed_vector"
   assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g"
-          "\<And>t. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> s"
-    shows "homotopic_paths s g h"
+          "\<And>t. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> S"
+    shows "homotopic_paths S g h"
   using assms
   unfolding path_def
   apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def)
@@ -1002,8 +979,8 @@
 lemma homotopic_loops_linear:
   fixes g h :: "real \<Rightarrow> 'a::real_normed_vector"
   assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h"
-          "\<And>t x. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> s"
-    shows "homotopic_loops s g h"
+          "\<And>t x. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> S"
+    shows "homotopic_loops S g h"
   using assms
   unfolding path_def
   apply (simp add: pathstart_def pathfinish_def homotopic_loops_def homotopic_with_def)
@@ -1013,29 +990,33 @@
   done
 
 lemma homotopic_paths_nearby_explicit:
-  assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g"
-      and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> s\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)"
-    shows "homotopic_paths s g h"
-  apply (rule homotopic_paths_linear [OF assms(1-4)])
+  assumes \<section>: "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g"
+      and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> S\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)"
+    shows "homotopic_paths S g h"
+proof (rule homotopic_paths_linear [OF \<section>])
+  show "\<And>t. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> S"
   by (metis no segment_bound(1) subsetI norm_minus_commute not_le)
+qed
 
 lemma homotopic_loops_nearby_explicit:
-  assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h"
-      and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> s\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)"
-    shows "homotopic_loops s g h"
-  apply (rule homotopic_loops_linear [OF assms(1-4)])
+  assumes \<section>: "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h"
+      and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> S\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)"
+    shows "homotopic_loops S g h"
+proof (rule homotopic_loops_linear [OF \<section>])
+  show "\<And>t. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> S"
   by (metis no segment_bound(1) subsetI norm_minus_commute not_le)
+qed
 
 lemma homotopic_nearby_paths:
   fixes g h :: "real \<Rightarrow> 'a::euclidean_space"
-  assumes "path g" "open s" "path_image g \<subseteq> s"
+  assumes "path g" "open S" "path_image g \<subseteq> S"
     shows "\<exists>e. 0 < e \<and>
                (\<forall>h. path h \<and>
                     pathstart h = pathstart g \<and> pathfinish h = pathfinish g \<and>
-                    (\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_paths s g h)"
+                    (\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_paths S g h)"
 proof -
-  obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - s \<Longrightarrow> e \<le> dist x y"
-    using separate_compact_closed [of "path_image g" "-s"] assms by force
+  obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - S \<Longrightarrow> e \<le> dist x y"
+    using separate_compact_closed [of "path_image g" "-S"] assms by force
   show ?thesis
     apply (intro exI conjI)
     using e [unfolded dist_norm]
@@ -1045,13 +1026,13 @@
 
 lemma homotopic_nearby_loops:
   fixes g h :: "real \<Rightarrow> 'a::euclidean_space"
-  assumes "path g" "open s" "path_image g \<subseteq> s" "pathfinish g = pathstart g"
+  assumes "path g" "open S" "path_image g \<subseteq> S" "pathfinish g = pathstart g"
     shows "\<exists>e. 0 < e \<and>
                (\<forall>h. path h \<and> pathfinish h = pathstart h \<and>
-                    (\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_loops s g h)"
+                    (\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_loops S g h)"
 proof -
-  obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - s \<Longrightarrow> e \<le> dist x y"
-    using separate_compact_closed [of "path_image g" "-s"] assms by force
+  obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - S \<Longrightarrow> e \<le> dist x y"
+    using separate_compact_closed [of "path_image g" "-S"] assms by force
   show ?thesis
     apply (intro exI conjI)
     using e [unfolded dist_norm]
@@ -1072,28 +1053,34 @@
   have 2: "t * 2 > 1 \<Longrightarrow> u + (2*t - 1) * v \<le> v + (2*t - 1) * w" for t
     by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \<open>u \<le> v\<close> \<open>v \<le> w\<close>)
   have t2: "\<And>t::real. t*2 = 1 \<Longrightarrow> t = 1/2" by auto
-  show ?thesis
-    apply (rule homotopic_paths_subset [OF _ pag])
-    using assms
-    apply (cases "w = u")
-    using homotopic_paths_rinv [of "subpath u v g" "path_image g"]
-    apply (force simp: closed_segment_eq_real_ivl image_mono path_image_def subpath_refl)
-      apply (rule homotopic_paths_sym)
-      apply (rule homotopic_paths_reparametrize
-             [where f = "\<lambda>t. if  t \<le> 1 / 2
-                             then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t
-                             else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))"])
-      using \<open>path g\<close> path_subpath u w apply blast
-      using \<open>path g\<close> path_image_subpath_subset u w(1) apply blast
-      apply simp_all
-      apply (subst split_01)
-      apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+
-      apply (simp_all add: field_simps not_le)
-      apply (force dest!: t2)
-      apply (force simp: algebra_simps mult_left_mono affine_ineq dest!: 1 2)
-      apply (simp add: joinpaths_def subpath_def)
-      apply (force simp: algebra_simps)
-      done
+  have "homotopic_paths (path_image g) (subpath u v g +++ subpath v w g) (subpath u w g)"
+  proof (cases "w = u")
+    case True
+    then show ?thesis
+      by (metis \<open>path g\<close> homotopic_paths_rinv path_image_subpath_subset path_subpath pathstart_subpath reversepath_subpath subpath_refl u v)
+  next
+    case False
+    let ?f = "\<lambda>t. if  t \<le> 1/2 then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t
+                               else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))"
+    show ?thesis
+    proof (rule homotopic_paths_sym [OF homotopic_paths_reparametrize [where f = ?f]])
+      show "path (subpath u w g)"
+        using assms(1) path_subpath u w(1) by blast
+      show "path_image (subpath u w g) \<subseteq> path_image g"
+        by (meson path_image_subpath_subset u w(1))
+      show "continuous_on {0..1} ?f"
+        unfolding split_01
+        by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def dest!: t2)+
+      show "?f ` {0..1} \<subseteq> {0..1}"
+        using False assms
+        by (force simp: field_simps not_le mult_left_mono affine_ineq dest!: 1 2)
+      show "(subpath u v g +++ subpath v w g) t = subpath u w g (?f t)" if "t \<in> {0..1}" for t 
+        using assms
+        unfolding joinpaths_def subpath_def by (auto simp add: divide_simps add.commute mult.commute mult.left_commute)
+    qed (use False in auto)
+  qed
+  then show ?thesis
+    by (rule homotopic_paths_subset [OF _ pag])
 qed
 
 lemma homotopic_join_subpaths2:
@@ -1108,25 +1095,30 @@
     shows "homotopic_paths s (subpath v w g +++ subpath w u g) (subpath v u g)"
 proof -
   have "homotopic_paths s (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)"
-    apply (rule homotopic_paths_join)
-    using hom homotopic_paths_sym_eq apply blast
-    apply (metis \<open>path g\<close> homotopic_paths_eq pag path_image_subpath_subset path_subpath subset_trans v w, simp)
-    done
+  proof (rule homotopic_paths_join)
+    show "homotopic_paths s (subpath u w g) (subpath u v g +++ subpath v w g)"
+      using hom homotopic_paths_sym_eq by blast
+    show "homotopic_paths s (subpath w v g) (subpath w v g)"
+      by (metis \<open>path g\<close> homotopic_paths_eq pag path_image_subpath_subset path_subpath subset_trans v w)
+  qed auto
   also have "homotopic_paths s ((subpath u v g +++ subpath v w g) +++ subpath w v g) (subpath u v g +++ subpath v w g +++ subpath w v g)"
-    apply (rule homotopic_paths_sym [OF homotopic_paths_assoc])
-    using assms by (simp_all add: path_image_subpath_subset [THEN order_trans])
+    by (rule homotopic_paths_sym [OF homotopic_paths_assoc])
+       (use assms in \<open>simp_all add: path_image_subpath_subset [THEN order_trans]\<close>)
   also have "homotopic_paths s (subpath u v g +++ subpath v w g +++ subpath w v g)
                                (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))"
-    apply (rule homotopic_paths_join)
-    apply (metis \<open>path g\<close> homotopic_paths_eq order.trans pag path_image_subpath_subset path_subpath u v)
-    apply (metis (no_types, lifting) \<open>path g\<close> homotopic_paths_linv order_trans pag path_image_subpath_subset path_subpath pathfinish_subpath reversepath_subpath v w)
-    apply simp
-    done
+  proof (rule homotopic_paths_join; simp)
+    show "path (subpath u v g) \<and> path_image (subpath u v g) \<subseteq> s"
+      by (metis \<open>path g\<close> order.trans pag path_image_subpath_subset path_subpath u v)
+    show "homotopic_paths s (subpath v w g +++ subpath w v g) (linepath (g v) (g v))"
+      by (metis (no_types, lifting) \<open>path g\<close> homotopic_paths_linv order_trans pag path_image_subpath_subset path_subpath pathfinish_subpath reversepath_subpath v w)
+  qed 
   also have "homotopic_paths s (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)"
-    apply (rule homotopic_paths_rid)
-    using \<open>path g\<close> path_subpath u v apply blast
-    apply (meson \<open>path g\<close> order.trans pag path_image_subpath_subset u v)
-    done
+  proof (rule homotopic_paths_rid)
+    show "path (subpath u v g)"
+      using \<open>path g\<close> path_subpath u v by blast
+    show "path_image (subpath u v g) \<subseteq> s"
+      by (meson \<open>path g\<close> order.trans pag path_image_subpath_subset u v)
+  qed
   finally have "homotopic_paths s (subpath u w g +++ subpath w v g) (subpath u v g)" .
   then show ?thesis
     using homotopic_join_subpaths2 by blast
@@ -1135,8 +1127,8 @@
 proposition homotopic_join_subpaths:
    "\<lbrakk>path g; path_image g \<subseteq> s; u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk>
     \<Longrightarrow> homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)"
-  apply (rule le_cases3 [of u v w])
-using homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3 by metis+
+  using le_cases3 [of u v w] homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3 
+  by metis
 
 text\<open>Relating homotopy of trivial loops to path-connectedness.\<close>
 
@@ -1200,11 +1192,10 @@
             (\<forall>p a. path p \<and> path_image p \<subseteq> S \<and>
                   pathfinish p = pathstart p \<and> a \<in> S
                   \<longrightarrow> homotopic_loops S p (linepath a a))"
-apply (simp add: simply_connected_def)
+  unfolding simply_connected_def
 apply (rule iffI, force, clarify)
-apply (rule_tac q = "linepath (pathstart p) (pathstart p)" in homotopic_loops_trans)
-apply (fastforce simp add:)
-using homotopic_loops_sym apply blast
+  apply (rule_tac q = "linepath (pathstart p) (pathstart p)" in homotopic_loops_trans)
+using homotopic_loops_sym apply blast+
 done
 
 lemma simply_connected_eq_contractible_loop_some:
@@ -1213,13 +1204,23 @@
                 path_connected S \<and>
                 (\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p
                     \<longrightarrow> (\<exists>a. a \<in> S \<and> homotopic_loops S p (linepath a a)))"
-apply (rule iffI)
- apply (fastforce simp: simply_connected_imp_path_connected simply_connected_eq_contractible_loop_any)
-apply (clarsimp simp add: simply_connected_eq_contractible_loop_any)
-apply (drule_tac x=p in spec)
-using homotopic_loops_trans path_connected_eq_homotopic_points
-  apply blast
-done
+     (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+  using simply_connected_eq_contractible_loop_any by (blast intro: simply_connected_imp_path_connected)
+next
+  assume r: ?rhs
+  note pa = r [THEN conjunct2, rule_format]
+  show ?lhs
+  proof (clarsimp simp add: simply_connected_eq_contractible_loop_any)
+    fix p a
+    assume "path p" and "path_image p \<subseteq> S" "pathfinish p = pathstart p"
+      and "a \<in> S"
+    with pa [of p] show "homotopic_loops S p (linepath a a)"
+      using homotopic_loops_trans path_connected_eq_homotopic_points r by blast
+  qed
+qed
 
 lemma simply_connected_eq_contractible_loop_all:
   fixes S :: "_::real_normed_vector set"
@@ -1242,9 +1243,9 @@
   next
     assume ?rhs
     then show "simply_connected S"
-      apply (simp add: simply_connected_eq_contractible_loop_any False)
-      by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans
-             path_component_imp_homotopic_points path_component_refl)
+      unfolding simply_connected_eq_contractible_loop_any 
+      by (meson False homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans 
+          path_component_imp_homotopic_points path_component_refl)
   qed
 qed
 
@@ -1254,11 +1255,17 @@
            path_connected S \<and>
            (\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p
             \<longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p)))"
-apply (rule iffI)
- apply (simp add: simply_connected_imp_path_connected)
- apply (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null)
-by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image
-         simply_connected_eq_contractible_loop_some subset_iff)
+     (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    unfolding simply_connected_imp_path_connected
+    by (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null)
+next
+  assume  ?rhs
+  then show ?lhs
+    using homotopic_paths_imp_homotopic_loops simply_connected_eq_contractible_loop_some by fastforce
+qed
 
 lemma simply_connected_eq_homotopic_paths:
   fixes S :: "_::real_normed_vector set"
@@ -1316,38 +1323,25 @@
        for p a b
   proof -
     have "path (fst \<circ> p)"
-      apply (rule Path_Connected.path_continuous_image [OF \<open>path p\<close>])
-      apply (rule continuous_intros)+
-      done
+      by (simp add: continuous_on_fst Path_Connected.path_continuous_image [OF \<open>path p\<close>])
     moreover have "path_image (fst \<circ> p) \<subseteq> S"
       using that apply (simp add: path_image_def) by force
     ultimately have p1: "homotopic_loops S (fst \<circ> p) (linepath a a)"
       using S that
-      apply (simp add: simply_connected_eq_contractible_loop_any)
-      apply (drule_tac x="fst \<circ> p" in spec)
-      apply (drule_tac x=a in spec)
-      apply (auto simp: pathstart_def pathfinish_def)
-      done
+      by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def)
     have "path (snd \<circ> p)"
-      apply (rule Path_Connected.path_continuous_image [OF \<open>path p\<close>])
-      apply (rule continuous_intros)+
-      done
+      by (simp add: continuous_on_snd Path_Connected.path_continuous_image [OF \<open>path p\<close>])
     moreover have "path_image (snd \<circ> p) \<subseteq> T"
       using that apply (simp add: path_image_def) by force
     ultimately have p2: "homotopic_loops T (snd \<circ> p) (linepath b b)"
       using T that
-      apply (simp add: simply_connected_eq_contractible_loop_any)
-      apply (drule_tac x="snd \<circ> p" in spec)
-      apply (drule_tac x=b in spec)
-      apply (auto simp: pathstart_def pathfinish_def)
-      done
+      by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def)
     show ?thesis
-      using p1 p2
-      apply (simp add: homotopic_loops, clarify)
+      using p1 p2 unfolding homotopic_loops
+      apply clarify
       apply (rename_tac h k)
       apply (rule_tac x="\<lambda>z. Pair (h z) (k z)" in exI)
-      apply (intro conjI continuous_intros | assumption)+
-      apply (auto simp: pathstart_def pathfinish_def)
+      apply (force intro: continuous_intros simp: pathstart_def pathfinish_def)+
       done
   qed
   with assms show ?thesis
@@ -1371,14 +1365,15 @@
     using assms by (force simp: contractible_def)
   then have "a \<in> S"
     by (metis False homotopic_constant_maps homotopic_with_symD homotopic_with_trans path_component_in_topspace topspace_euclidean_subtopology)
-  show ?thesis
-    apply (simp add: simply_connected_eq_contractible_loop_all False)
-    apply (rule bexI [OF _ \<open>a \<in> S\<close>])
+  have "\<forall>p. path p \<and>
+            path_image p \<subseteq> S \<and> pathfinish p = pathstart p \<longrightarrow>
+            homotopic_loops S p (linepath a a)"
     using a apply (simp add: homotopic_loops_def homotopic_with_def path_def path_image_def pathfinish_def pathstart_def, clarify)
     apply (rule_tac x="(h \<circ> (\<lambda>y. (fst y, (p \<circ> snd) y)))" in exI)
-    apply (intro conjI continuous_on_compose continuous_intros)
-    apply (erule continuous_on_subset | force)+
+    apply (intro conjI continuous_on_compose continuous_intros; force elim: continuous_on_subset)
     done
+  with \<open>a \<in> S\<close> show ?thesis
+    by (auto simp add: simply_connected_eq_contractible_loop_all False)
 qed
 
 corollary contractible_imp_connected:
@@ -1401,7 +1396,7 @@
   obtain b where b: "homotopic_with_canon (\<lambda>x. True) T T id (\<lambda>x. b)"
     using assms by (force simp: contractible_def)
   have "homotopic_with_canon (\<lambda>f. True) T U (g \<circ> id) (g \<circ> (\<lambda>x. b))"
-    by (metis Abstract_Topology.continuous_map_subtopology_eu b g homotopic_compose_continuous_map_left)
+    by (metis Abstract_Topology.continuous_map_subtopology_eu b g homotopic_with_compose_continuous_map_left)
   then have "homotopic_with_canon (\<lambda>f. True) S U (g \<circ> id \<circ> f) (g \<circ> (\<lambda>x. b) \<circ> f)"
     by (simp add: f homotopic_with_compose_continuous_map_right)
   then show ?thesis
@@ -1412,19 +1407,14 @@
   assumes f: "continuous_on S f" "f ` S \<subseteq> T"
       and T: "contractible T"
     obtains c where "homotopic_with_canon (\<lambda>h. True) S T f (\<lambda>x. c)"
-apply (rule nullhomotopic_through_contractible [OF f, of id T])
-using assms
-apply (auto)
-done
+  by (rule nullhomotopic_through_contractible [OF f, of id T]) (use assms in auto)
 
 lemma nullhomotopic_from_contractible:
   assumes f: "continuous_on S f" "f ` S \<subseteq> T"
       and S: "contractible S"
     obtains c where "homotopic_with_canon (\<lambda>h. True) S T f (\<lambda>x. c)"
-apply (rule nullhomotopic_through_contractible [OF continuous_on_id _ f S, of S])
-using assms
-apply (auto simp: comp_def)
-done
+  apply (rule nullhomotopic_through_contractible [OF continuous_on_id _ f S, of S])
+  using assms by (auto simp: comp_def)
 
 lemma homotopic_through_contractible:
   fixes S :: "_::real_normed_vector set"
@@ -1436,14 +1426,10 @@
    shows "homotopic_with_canon (\<lambda>h. True) S U (g1 \<circ> f1) (g2 \<circ> f2)"
 proof -
   obtain c1 where c1: "homotopic_with_canon (\<lambda>h. True) S U (g1 \<circ> f1) (\<lambda>x. c1)"
-    apply (rule nullhomotopic_through_contractible [of S f1 T g1 U])
-    using assms apply auto
-    done
+    by (rule nullhomotopic_through_contractible [of S f1 T g1 U]) (use assms in auto)
   obtain c2 where c2: "homotopic_with_canon (\<lambda>h. True) S U (g2 \<circ> f2) (\<lambda>x. c2)"
-    apply (rule nullhomotopic_through_contractible [of S f2 T g2 U])
-    using assms apply auto
-    done
-  have *: "S = {} \<or> (\<exists>t. path_connected t \<and> t \<subseteq> U \<and> c2 \<in> t \<and> c1 \<in> t)"
+    by (rule nullhomotopic_through_contractible [of S f2 T g2 U]) (use assms in auto)
+  have "S = {} \<or> (\<exists>t. path_connected t \<and> t \<subseteq> U \<and> c2 \<in> t \<and> c1 \<in> t)"
   proof (cases "S = {}")
     case True then show ?thesis by force
   next
@@ -1452,12 +1438,10 @@
       using homotopic_with_imp_continuous_maps by fastforce+
     with \<open>path_connected U\<close> show ?thesis by blast
   qed
-  show ?thesis
-    apply (rule homotopic_with_trans [OF c1])
-    apply (rule homotopic_with_symD)
-    apply (rule homotopic_with_trans [OF c2])
-    apply (simp add: path_component homotopic_constant_maps *)
-    done
+  then have "homotopic_with_canon (\<lambda>h. True) S U (\<lambda>x. c2) (\<lambda>x. c1)"
+    by (simp add: path_component homotopic_constant_maps)
+  then show ?thesis
+    using c1 c2 homotopic_with_symD homotopic_with_trans by blast
 qed
 
 lemma homotopic_into_contractible:
@@ -1498,16 +1482,14 @@
     by (simp add: assms rel_interior_eq_empty)
   then obtain a where a: "a \<in> rel_interior S"  by blast
   with ST have "a \<in> T"  by blast
-  have *: "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
-    apply (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a])
-    using assms by blast
-  show ?thesis
-    unfolding starlike_def
-    apply (rule bexI [OF _ \<open>a \<in> T\<close>])
-    apply (simp add: closed_segment_eq_open)
+  have "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
+    by (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) (use assms in auto)
+  then have "\<forall>x\<in>T. a \<in> T \<and> open_segment a x \<subseteq> T"
     apply (intro conjI ballI a \<open>a \<in> T\<close> rel_interior_closure_convex_segment [OF \<open>convex S\<close> a])
-    apply (simp add: order_trans [OF * ST])
-    done
+    using ST by blast
+  then show ?thesis
+    unfolding starlike_def using bexI [OF _ \<open>a \<in> T\<close>]
+    by (simp add: closed_segment_eq_open)
 qed
 
 lemma starlike_imp_contractible_gen:
@@ -1523,14 +1505,14 @@
     apply (erule a [unfolded closed_segment_def, THEN subsetD], simp)
     apply (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1))
     done
-  then show ?thesis
-    apply (rule_tac a=a in that)
+  then have "homotopic_with_canon P S S (\<lambda>x. x) (\<lambda>x. a)"
     using \<open>a \<in> S\<close>
     apply (simp add: homotopic_with_def)
     apply (rule_tac x="\<lambda>y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI)
-    apply (intro conjI ballI continuous_on_compose continuous_intros)
-    apply (simp_all add: P)
+    apply (intro conjI ballI continuous_on_compose continuous_intros; simp add: P)
     done
+  then show ?thesis
+    using that by blast
 qed
 
 lemma starlike_imp_contractible:
@@ -1580,9 +1562,7 @@
 next
   case False
   show ?thesis
-    apply (rule starlike_imp_contractible)
-    apply (rule starlike_convex_tweak_boundary_points [OF \<open>convex S\<close> False TS])
-    done
+    by (meson False assms starlike_convex_tweak_boundary_points starlike_imp_contractible)
 qed
 
 lemma convex_imp_contractible:
@@ -1658,14 +1638,10 @@
 lemma locally_open_subset:
   assumes "locally P S" "openin (top_of_set S) t"
     shows "locally P t"
-using assms
-apply (simp add: locally_def)
-apply (erule all_forward)+
-apply (rule impI)
-apply (erule impCE)
- using openin_trans apply blast
-apply (erule ex_forward)
-by (metis (no_types, hide_lams) Int_absorb1 Int_lower1 Int_subset_iff openin_open openin_subtopology_Int_subset)
+  using assms
+  unfolding locally_def
+  apply (elim all_forward)
+  by (meson dual_order.trans openin_imp_subset openin_subset_trans openin_trans)
 
 lemma locally_diff_closed:
     "\<lbrakk>locally P S; closedin (top_of_set S) t\<rbrakk> \<Longrightarrow> locally P (S - t)"
@@ -1682,27 +1658,32 @@
 
 lemma locally_iff:
     "locally P S \<longleftrightarrow>
-     (\<forall>T x. open T \<and> x \<in> S \<inter> T \<longrightarrow> (\<exists>U. open U \<and> (\<exists>v. P v \<and> x \<in> S \<inter> U \<and> S \<inter> U \<subseteq> v \<and> v \<subseteq> S \<inter> T)))"
+     (\<forall>T x. open T \<and> x \<in> S \<inter> T \<longrightarrow> (\<exists>U. open U \<and> (\<exists>V. P V \<and> x \<in> S \<inter> U \<and> S \<inter> U \<subseteq> V \<and> V \<subseteq> S \<inter> T)))"
 apply (simp add: le_inf_iff locally_def openin_open, safe)
 apply (metis IntE IntI le_inf_iff)
 apply (metis IntI Int_subset_iff)
 done
 
 lemma locally_Int:
-  assumes S: "locally P S" and t: "locally P t"
-      and P: "\<And>S t. P S \<and> P t \<Longrightarrow> P(S \<inter> t)"
-    shows "locally P (S \<inter> t)"
-using S t unfolding locally_iff
-apply clarify
-apply (drule_tac x=T in spec)+
-apply (drule_tac x=x in spec)+
-apply clarsimp
-apply (rename_tac U1 U2 V1 V2)
-apply (rule_tac x="U1 \<inter> U2" in exI)
-apply (simp add: open_Int)
-apply (rule_tac x="V1 \<inter> V2" in exI)
-apply (auto intro: P)
-  done
+  assumes S: "locally P S" and T: "locally P T"
+      and P: "\<And>S T. P S \<and> P T \<Longrightarrow> P(S \<inter> T)"
+  shows "locally P (S \<inter> T)"
+  unfolding locally_iff
+proof clarify
+  fix A x
+  assume "open A" "x \<in> A" "x \<in> S" "x \<in> T"
+  then obtain U1 V1 U2 V2 
+    where "open U1" "P V1" "x \<in> S \<inter> U1" "S \<inter> U1 \<subseteq> V1 \<and> V1 \<subseteq> S \<inter> A" 
+          "open U2" "P V2" "x \<in> T \<inter> U2" "T \<inter> U2 \<subseteq> V2 \<and> V2 \<subseteq> T \<inter> A"
+    using S T unfolding locally_iff by (meson IntI)
+  then have "S \<inter> T \<inter> (U1 \<inter> U2) \<subseteq> V1 \<inter> V2" "V1 \<inter> V2 \<subseteq> S \<inter> T \<inter> A" "x \<in> S \<inter> T \<inter> (U1 \<inter> U2)"
+    by blast+
+  moreover have "P (V1 \<inter> V2)"
+    by (simp add: P \<open>P V1\<close> \<open>P V2\<close>)
+  ultimately show "\<exists>U. open U \<and> (\<exists>V. P V \<and> x \<in> S \<inter> T \<inter> U \<and> S \<inter> T \<inter> U \<subseteq> V \<and> V \<subseteq> S \<inter> T \<inter> A)"
+    using \<open>open U1\<close> \<open>open U2\<close> by blast
+qed
+
 
 lemma locally_Times:
   fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set"
@@ -1728,73 +1709,73 @@
 
 
 proposition homeomorphism_locally_imp:
-  fixes S :: "'a::metric_space set" and t :: "'b::t2_space set"
-  assumes S: "locally P S" and hom: "homeomorphism S t f g"
+  fixes S :: "'a::metric_space set" and T :: "'b::t2_space set"
+  assumes S: "locally P S" and hom: "homeomorphism S T f g"
       and Q: "\<And>S S'. \<lbrakk>P S; homeomorphism S S' f g\<rbrakk> \<Longrightarrow> Q S'"
-    shows "locally Q t"
+    shows "locally Q T"
 proof (clarsimp simp: locally_def)
   fix W y
-  assume "y \<in> W" and "openin (top_of_set t) W"
-  then obtain T where T: "open T" "W = t \<inter> T"
+  assume "y \<in> W" and "openin (top_of_set T) W"
+  then obtain A where T: "open A" "W = T \<inter> A"
     by (force simp: openin_open)
-  then have "W \<subseteq> t" by auto
-  have f: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "f ` S = t" "continuous_on S f"
-   and g: "\<And>y. y \<in> t \<Longrightarrow> f(g y) = y" "g ` t = S" "continuous_on t g"
+  then have "W \<subseteq> T" by auto
+  have f: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "f ` S = T" "continuous_on S f"
+   and g: "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y" "g ` T = S" "continuous_on T g"
     using hom by (auto simp: homeomorphism_def)
   have gw: "g ` W = S \<inter> f -` W"
-    using \<open>W \<subseteq> t\<close>
-    apply auto
-    using \<open>g ` t = S\<close> \<open>W \<subseteq> t\<close> apply blast
-    using g \<open>W \<subseteq> t\<close> apply auto[1]
-    by (simp add: f rev_image_eqI)
+    using \<open>W \<subseteq> T\<close> g by force
   have \<circ>: "openin (top_of_set S) (g ` W)"
   proof -
     have "continuous_on S f"
       using f(3) by blast
     then show "openin (top_of_set S) (g ` W)"
-      by (simp add: gw Collect_conj_eq \<open>openin (top_of_set t) W\<close> continuous_on_open f(2))
+      by (simp add: gw Collect_conj_eq \<open>openin (top_of_set T) W\<close> continuous_on_open f(2))
   qed
-  then obtain u v
-    where osu: "openin (top_of_set S) u" and uv: "P v" "g y \<in> u" "u \<subseteq> v" "v \<subseteq> g ` W"
+  then obtain U V
+    where osu: "openin (top_of_set S) U" and uv: "P V" "g y \<in> U" "U \<subseteq> V" "V \<subseteq> g ` W"
     using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \<open>y \<in> W\<close> by force
-  have "v \<subseteq> S" using uv by (simp add: gw)
-  have fv: "f ` v = t \<inter> {x. g x \<in> v}"
-    using \<open>f ` S = t\<close> f \<open>v \<subseteq> S\<close> by auto
-  have "f ` v \<subseteq> W"
+  have "V \<subseteq> S" using uv by (simp add: gw)
+  have fv: "f ` V = T \<inter> {x. g x \<in> V}"
+    using \<open>f ` S = T\<close> f \<open>V \<subseteq> S\<close> by auto
+  have "f ` V \<subseteq> W"
     using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto
-  have contvf: "continuous_on v f"
-    using \<open>v \<subseteq> S\<close> continuous_on_subset f(3) by blast
-  have contvg: "continuous_on (f ` v) g"
-    using \<open>f ` v \<subseteq> W\<close> \<open>W \<subseteq> t\<close> continuous_on_subset [OF g(3)] by blast
-  have homv: "homeomorphism v (f ` v) f g"
-    using \<open>v \<subseteq> S\<close> \<open>W \<subseteq> t\<close> f
-    apply (simp add: homeomorphism_def contvf contvg, auto)
-    by (metis f(1) rev_image_eqI rev_subsetD)
-  have 1: "openin (top_of_set t) (t \<inter> g -` u)"
-    apply (rule continuous_on_open [THEN iffD1, rule_format])
-    apply (rule \<open>continuous_on t g\<close>)
-    using \<open>g ` t = S\<close> apply (simp add: osu)
+  have contvf: "continuous_on V f"
+    using \<open>V \<subseteq> S\<close> continuous_on_subset f(3) by blast
+  have contvg: "continuous_on (f ` V) g"
+    using \<open>f ` V \<subseteq> W\<close> \<open>W \<subseteq> T\<close> continuous_on_subset [OF g(3)] by blast
+  have "V \<subseteq> g ` f ` V"
+    by (metis hom homeomorphism_def homeomorphism_of_subsets set_eq_subset \<open>V \<subseteq> S\<close>)
+  then have homv: "homeomorphism V (f ` V) f g"
+    using \<open>V \<subseteq> S\<close> f by (auto simp add: homeomorphism_def contvf contvg)
+  have "openin (top_of_set (g ` T)) U"
+    using \<open>g ` T = S\<close> by (simp add: osu)
+  then have 1: "openin (top_of_set T) (T \<inter> g -` U)"
+    using \<open>continuous_on T g\<close> continuous_on_open [THEN iffD1] by blast
+  have 2: "\<exists>V. Q V \<and> y \<in> (T \<inter> g -` U) \<and> (T \<inter> g -` U) \<subseteq> V \<and> V \<subseteq> W"
+    apply (rule_tac x="f ` V" in exI)
+    apply (intro conjI Q [OF \<open>P V\<close> homv])
+    using \<open>W \<subseteq> T\<close> \<open>y \<in> W\<close>  \<open>f ` V \<subseteq> W\<close>  uv  apply (auto simp: fv)
     done
-  have 2: "\<exists>V. Q V \<and> y \<in> (t \<inter> g -` u) \<and> (t \<inter> g -` u) \<subseteq> V \<and> V \<subseteq> W"
-    apply (rule_tac x="f ` v" in exI)
-    apply (intro conjI Q [OF \<open>P v\<close> homv])
-    using \<open>W \<subseteq> t\<close> \<open>y \<in> W\<close>  \<open>f ` v \<subseteq> W\<close>  uv  apply (auto simp: fv)
-    done
-  show "\<exists>U. openin (top_of_set t) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)"
+  show "\<exists>U. openin (top_of_set T) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)"
     by (meson 1 2)
 qed
 
 lemma homeomorphism_locally:
   fixes f:: "'a::metric_space \<Rightarrow> 'b::metric_space"
-  assumes hom: "homeomorphism S t f g"
-      and eq: "\<And>S t. homeomorphism S t f g \<Longrightarrow> (P S \<longleftrightarrow> Q t)"
-    shows "locally P S \<longleftrightarrow> locally Q t"
-apply (rule iffI)
-apply (erule homeomorphism_locally_imp [OF _ hom])
-apply (simp add: eq)
-apply (erule homeomorphism_locally_imp)
-using eq homeomorphism_sym homeomorphism_symD [OF hom] apply blast+
-done
+  assumes hom: "homeomorphism S T f g"
+      and eq: "\<And>S T. homeomorphism S T f g \<Longrightarrow> (P S \<longleftrightarrow> Q T)"
+    shows "locally P S \<longleftrightarrow> locally Q T"
+     (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    using eq hom homeomorphism_locally_imp by blast
+next
+  assume ?rhs
+  then show ?lhs
+    using eq homeomorphism_sym homeomorphism_symD [OF hom] 
+    by (blast intro: homeomorphism_locally_imp) 
+qed
 
 lemma homeomorphic_locally:
   fixes S:: "'a::metric_space set" and T:: "'b::metric_space set"
@@ -1816,28 +1797,23 @@
 
 lemma locally_translation:
   fixes P :: "'a :: real_normed_vector set \<Rightarrow> bool"
-  shows
-   "(\<And>S. P (image (\<lambda>x. a + x) S) \<longleftrightarrow> P S)
-        \<Longrightarrow> locally P (image (\<lambda>x. a + x) S) \<longleftrightarrow> locally P S"
+  shows "(\<And>S. P ((+) a ` S) = P S) \<Longrightarrow> locally P ((+) a ` S) = locally P S"
 apply (rule homeomorphism_locally [OF homeomorphism_translation])
-apply (simp add: homeomorphism_def)
-by metis
+  by (metis (no_types) homeomorphism_def)
 
 lemma locally_injective_linear_image:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes f: "linear f" "inj f" and iff: "\<And>S. P (f ` S) \<longleftrightarrow> Q S"
-    shows "locally P (f ` S) \<longleftrightarrow> locally Q S"
-apply (rule linear_homeomorphism_image [OF f])
-apply (rule_tac f=g and g = f in homeomorphism_locally, assumption)
-by (metis iff homeomorphism_def)
+  shows "locally P (f ` S) \<longleftrightarrow> locally Q S"
+  using homeomorphism_locally [of "f`S" _ _ f] linear_homeomorphism_image [OF f]
+  by (metis (no_types, lifting) homeomorphism_image2 iff)
 
 lemma locally_open_map_image:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes P: "locally P S"
       and f: "continuous_on S f"
-      and oo: "\<And>t. openin (top_of_set S) t
-                   \<Longrightarrow> openin (top_of_set (f ` S)) (f ` t)"
-      and Q: "\<And>t. \<lbrakk>t \<subseteq> S; P t\<rbrakk> \<Longrightarrow> Q(f ` t)"
+      and oo: "\<And>T. openin (top_of_set S) T \<Longrightarrow> openin (top_of_set (f ` S)) (f ` T)"
+      and Q: "\<And>T. \<lbrakk>T \<subseteq> S; P T\<rbrakk> \<Longrightarrow> Q(f ` T)"
     shows "locally Q (f ` S)"
 proof (clarsimp simp add: locally_def)
   fix W y
@@ -1851,12 +1827,10 @@
     where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> S \<inter> f -` W"
     using P [unfolded locally_def, rule_format, of "(S \<inter> f -` W)" x] oivf \<open>y \<in> W\<close>
     by auto
+  then have "openin (top_of_set (f ` S)) (f ` U)"
+    by (simp add: oo)
   then show "\<exists>X. openin (top_of_set (f ` S)) X \<and> (\<exists>Y. Q Y \<and> y \<in> X \<and> X \<subseteq> Y \<and> Y \<subseteq> W)"
-    apply (rule_tac x="f ` U" in exI)
-    apply (rule conjI, blast intro!: oo)
-    apply (rule_tac x="f ` V" in exI)
-    apply (force simp: \<open>f x = y\<close> rev_image_eqI intro: Q)
-    done
+    using Q \<open>P V\<close> \<open>U \<subseteq> V\<close> \<open>V \<subseteq> S \<inter> f -` W\<close> \<open>f x = y\<close> \<open>x \<in> U\<close> by blast
 qed
 
 
@@ -1871,28 +1845,25 @@
       and etc: "a \<in> S" "b \<in> S" "P a" "P b" "Q a"
     shows "Q b"
 proof -
-  have 1: "openin (top_of_set S)
-             {b. \<exists>T. openin (top_of_set S) T \<and>
-                     b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> Q x)}"
+  let ?A = "{b. \<exists>T. openin (top_of_set S) T \<and> b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> Q x)}"
+  let ?B = "{b. \<exists>T. openin (top_of_set S) T \<and> b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> Q x)}"
+  have 1: "openin (top_of_set S) ?A"
+    apply (subst openin_subopen, clarify)
+    apply (rule_tac x=T in exI, auto)
+    done
+  have 2: "openin (top_of_set S) ?B"
     apply (subst openin_subopen, clarify)
     apply (rule_tac x=T in exI, auto)
     done
-  have 2: "openin (top_of_set S)
-             {b. \<exists>T. openin (top_of_set S) T \<and>
-                     b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> Q x)}"
-    apply (subst openin_subopen, clarify)
-    apply (rule_tac x=T in exI, auto)
-    done
-  show ?thesis
-    using \<open>connected S\<close>
-    apply (simp only: connected_openin HOL.not_ex HOL.de_Morgan_conj)
-    apply (elim disjE allE)
-         apply (blast intro: 1)
-        apply (blast intro: 2, simp_all)
-       apply clarify apply (metis opI)
-      using opD apply (blast intro: etc elim: dest:)
-     using opI etc apply meson+
-    done
+  have \<section>: "?A \<inter> ?B = {}"
+    by (clarsimp simp: set_eq_iff) (metis (no_types, hide_lams) Int_iff opD openin_Int)
+  have *: "S \<subseteq> ?A \<union> ?B"
+    by clarsimp (meson opI)
+  have "?A = {} \<or> ?B = {}"
+    using \<open>connected S\<close> [unfolded connected_openin, simplified, rule_format, OF 1 \<section> * 2] 
+    by blast
+  then show ?thesis
+    by clarsimp (meson opI etc)
 qed
 
 lemma connected_equivalence_relation_gen:
@@ -2000,26 +1971,25 @@
 qed
 
 lemma locally_compactE:
-  fixes s :: "'a :: metric_space set"
-  assumes "locally compact s"
-  obtains u v where "\<And>x. x \<in> s \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> s \<and>
-                             openin (top_of_set s) (u x) \<and> compact (v x)"
-using assms
-unfolding locally_compact by metis
+  fixes S :: "'a :: metric_space set"
+  assumes "locally compact S"
+  obtains u v where "\<And>x. x \<in> S \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> S \<and>
+                             openin (top_of_set S) (u x) \<and> compact (v x)"
+  using assms unfolding locally_compact by metis
 
 lemma locally_compact_alt:
-  fixes s :: "'a :: heine_borel set"
-  shows "locally compact s \<longleftrightarrow>
-         (\<forall>x \<in> s. \<exists>u. x \<in> u \<and>
-                    openin (top_of_set s) u \<and> compact(closure u) \<and> closure u \<subseteq> s)"
-apply (simp add: locally_compact)
-apply (intro ball_cong ex_cong refl iffI)
-apply (metis bounded_subset closure_eq closure_mono compact_eq_bounded_closed dual_order.trans)
-by (meson closure_subset compact_closure)
+  fixes S :: "'a :: heine_borel set"
+  shows "locally compact S \<longleftrightarrow>
+         (\<forall>x \<in> S. \<exists>u. x \<in> u \<and>
+                    openin (top_of_set S) u \<and> compact(closure u) \<and> closure u \<subseteq> S)"
+  apply (simp add: locally_compact)
+  apply (intro ball_cong ex_cong refl iffI)
+   apply (meson bounded_subset closure_minimal compact_eq_bounded_closed dual_order.trans)
+  by (meson closure_subset compact_closure)
 
 lemma locally_compact_Int_cball:
-  fixes s :: "'a :: heine_borel set"
-  shows "locally compact s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>e. 0 < e \<and> closed(cball x e \<inter> s))"
+  fixes S :: "'a :: heine_borel set"
+  shows "locally compact S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>e. 0 < e \<and> closed(cball x e \<inter> S))"
         (is "?lhs = ?rhs")
 proof
   assume ?lhs
@@ -2032,8 +2002,8 @@
   then show ?lhs
     apply (simp add: locally_compact openin_contains_cball)
     apply (clarify | assumption | drule bspec)+
-    apply (rule_tac x="ball x e \<inter> s" in exI, simp)
-    apply (rule_tac x="cball x e \<inter> s" in exI)
+    apply (rule_tac x="ball x e \<inter> S" in exI, simp)
+    apply (rule_tac x="cball x e \<inter> S" in exI)
     using compact_eq_bounded_closed
     apply auto
     apply (metis open_ball le_infI1 mem_ball open_contains_cball_eq)
@@ -2041,20 +2011,20 @@
 qed
 
 lemma locally_compact_compact:
-  fixes s :: "'a :: heine_borel set"
-  shows "locally compact s \<longleftrightarrow>
-         (\<forall>k. k \<subseteq> s \<and> compact k
-              \<longrightarrow> (\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and>
-                         openin (top_of_set s) u \<and> compact v))"
+  fixes S :: "'a :: heine_borel set"
+  shows "locally compact S \<longleftrightarrow>
+         (\<forall>k. k \<subseteq> S \<and> compact k
+              \<longrightarrow> (\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<and>
+                         openin (top_of_set S) u \<and> compact v))"
         (is "?lhs = ?rhs")
 proof
   assume ?lhs
   then obtain u v where
-    uv: "\<And>x. x \<in> s \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> s \<and>
-                             openin (top_of_set s) (u x) \<and> compact (v x)"
+    uv: "\<And>x. x \<in> S \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> S \<and>
+                             openin (top_of_set S) (u x) \<and> compact (v x)"
     by (metis locally_compactE)
-  have *: "\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (top_of_set s) u \<and> compact v"
-          if "k \<subseteq> s" "compact k" for k
+  have *: "\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<and> openin (top_of_set S) u \<and> compact v"
+          if "k \<subseteq> S" "compact k" for k
   proof -
     have "\<And>C. (\<forall>c\<in>C. openin (top_of_set k) c) \<and> k \<subseteq> \<Union>C \<Longrightarrow>
                     \<exists>D\<subseteq>C. finite D \<and> k \<subseteq> \<Union>D"
@@ -2072,9 +2042,7 @@
     show ?thesis
       apply (rule_tac x="\<Union>(u ` T)" in exI)
       apply (rule_tac x="\<Union>(v ` T)" in exI)
-      apply (simp add: Tuv)
-      using T that
-      apply (auto simp: dest!: uv)
+      using T that apply (auto simp: Tuv dest!: uv)
       done
   qed
   show ?rhs
@@ -2088,43 +2056,42 @@
 qed
 
 lemma open_imp_locally_compact:
-  fixes s :: "'a :: heine_borel set"
-  assumes "open s"
-    shows "locally compact s"
+  fixes S :: "'a :: heine_borel set"
+  assumes "open S"
+    shows "locally compact S"
 proof -
-  have *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (top_of_set s) u \<and> compact v"
-          if "x \<in> s" for x
+  have *: "\<exists>U V. x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V"
+          if "x \<in> S" for x
   proof -
-    obtain e where "e>0" and e: "cball x e \<subseteq> s"
-      using open_contains_cball assms \<open>x \<in> s\<close> by blast
-    have ope: "openin (top_of_set s) (ball x e)"
+    obtain e where "e>0" and e: "cball x e \<subseteq> S"
+      using open_contains_cball assms \<open>x \<in> S\<close> by blast
+    have ope: "openin (top_of_set S) (ball x e)"
       by (meson e open_ball ball_subset_cball dual_order.trans open_subset)
     show ?thesis
-      apply (rule_tac x="ball x e" in exI)
-      apply (rule_tac x="cball x e" in exI)
-      using \<open>e > 0\<close> e apply (auto simp: ope)
-      done
+    proof (intro exI conjI)
+      let ?U = "ball x e"
+      let ?V = "cball x e"
+      show "x \<in> ?U" "?U \<subseteq> ?V" "?V \<subseteq> S" "compact ?V"
+        using \<open>e > 0\<close> e by auto
+      show "openin (top_of_set S) ?U"
+        using ope by blast
+    qed
   qed
   show ?thesis
-    unfolding locally_compact
-    by (blast intro: *)
+    unfolding locally_compact by (blast intro: *)
 qed
 
 lemma closed_imp_locally_compact:
-  fixes s :: "'a :: heine_borel set"
-  assumes "closed s"
-    shows "locally compact s"
+  fixes S :: "'a :: heine_borel set"
+  assumes "closed S"
+    shows "locally compact S"
 proof -
-  have *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and>
-                 openin (top_of_set s) u \<and> compact v"
-          if "x \<in> s" for x
-  proof -
-    show ?thesis
-      apply (rule_tac x = "s \<inter> ball x 1" in exI)
-      apply (rule_tac x = "s \<inter> cball x 1" in exI)
-      using \<open>x \<in> s\<close> assms apply auto
-      done
-  qed
+  have *: "\<exists>U V. x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V"
+          if "x \<in> S" for x
+    apply (rule_tac x = "S \<inter> ball x 1" in exI)
+    apply (rule_tac x = "S \<inter> cball x 1" in exI)
+    using \<open>x \<in> S\<close> assms apply auto
+    done
   show ?thesis
     unfolding locally_compact
     by (blast intro: *)
@@ -2134,25 +2101,25 @@
   by (simp add: closed_imp_locally_compact)
 
 lemma locally_compact_Int:
-  fixes s :: "'a :: t2_space set"
-  shows "\<lbrakk>locally compact s; locally compact t\<rbrakk> \<Longrightarrow> locally compact (s \<inter> t)"
+  fixes S :: "'a :: t2_space set"
+  shows "\<lbrakk>locally compact S; locally compact t\<rbrakk> \<Longrightarrow> locally compact (S \<inter> t)"
 by (simp add: compact_Int locally_Int)
 
 lemma locally_compact_closedin:
-  fixes s :: "'a :: heine_borel set"
-  shows "\<lbrakk>closedin (top_of_set s) t; locally compact s\<rbrakk>
+  fixes S :: "'a :: heine_borel set"
+  shows "\<lbrakk>closedin (top_of_set S) t; locally compact S\<rbrakk>
         \<Longrightarrow> locally compact t"
-unfolding closedin_closed
-using closed_imp_locally_compact locally_compact_Int by blast
+  unfolding closedin_closed
+  using closed_imp_locally_compact locally_compact_Int by blast
 
 lemma locally_compact_delete:
-     fixes s :: "'a :: t1_space set"
-     shows "locally compact s \<Longrightarrow> locally compact (s - {a})"
+     fixes S :: "'a :: t1_space set"
+     shows "locally compact S \<Longrightarrow> locally compact (S - {a})"
   by (auto simp: openin_delete locally_open_subset)
 
 lemma locally_closed:
-  fixes s :: "'a :: heine_borel set"
-  shows "locally closed s \<longleftrightarrow> locally compact s"
+  fixes S :: "'a :: heine_borel set"
+  shows "locally closed S \<longleftrightarrow> locally compact S"
         (is "?lhs = ?rhs")
 proof
   assume ?lhs
@@ -2376,10 +2343,8 @@
             and cloV2: "closedin (top_of_set D) (D \<inter> closure V2)"
             by (simp_all add: closedin_closed_Int)
           moreover have "D \<inter> closure V1 = D \<inter> V1" "D \<inter> closure V2 = D \<inter> V2"
-            apply safe
             using \<open>D \<subseteq> V1 \<union> V2\<close> \<open>open V1\<close> \<open>open V2\<close> V12
-               apply (simp_all add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+)
-            done
+            by (auto simp add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+)
           ultimately have cloDV1: "closedin (top_of_set D) (D \<inter> V1)"
                       and cloDV2:  "closedin (top_of_set D) (D \<inter> V2)"
             by metis+
@@ -2391,10 +2356,12 @@
             assume "D \<inter> U1 \<inter> C = {}"
             then have *: "C \<subseteq> D \<inter> V2"
               using D1 DV12 \<open>C \<subseteq> D\<close> by auto
-            have "\<Inter>?\<T> \<subseteq> D \<inter> V2"
-              apply (rule Inter_lower)
-              using * apply simp
-              by (meson cloDV2 \<open>open V2\<close> cloD closedin_trans le_inf_iff opeD openin_Int_open)
+            have 1: "openin (top_of_set S) (D \<inter> V2)"
+              by (simp add: \<open>open V2\<close> opeD openin_Int_open)
+            have 2: "closedin (top_of_set S) (D \<inter> V2)"
+              using cloD cloDV2 closedin_trans by blast
+            have "\<Inter> ?\<T> \<subseteq> D \<inter> V2"
+              by (rule Inter_lower) (use * 1 2 in simp)
             then show False
               using K1 V12 \<open>K1 \<noteq> {}\<close> \<open>K1 \<subseteq> V1\<close> closedin_imp_subset by blast
           qed
@@ -2403,18 +2370,17 @@
             assume "D \<inter> U2 \<inter> C = {}"
             then have *: "C \<subseteq> D \<inter> V1"
               using D2 DV12 \<open>C \<subseteq> D\<close> by auto
+            have 1: "openin (top_of_set S) (D \<inter> V1)"
+              by (simp add: \<open>open V1\<close> opeD openin_Int_open)
+            have 2: "closedin (top_of_set S) (D \<inter> V1)"
+              using cloD cloDV1 closedin_trans by blast
             have "\<Inter>?\<T> \<subseteq> D \<inter> V1"
-              apply (rule Inter_lower)
-              using * apply simp
-              by (meson cloDV1 \<open>open V1\<close> cloD closedin_trans le_inf_iff opeD openin_Int_open)
+              by (rule Inter_lower) (use * 1 2 in simp)
             then show False
               using K2 V12 \<open>K2 \<noteq> {}\<close> \<open>K2 \<subseteq> V2\<close> closedin_imp_subset by blast
           qed
           ultimately show False
-            using \<open>connected C\<close> unfolding connected_closed
-            apply (simp only: not_ex)
-            apply (drule_tac x="D \<inter> U1" in spec)
-            apply (drule_tac x="D \<inter> U2" in spec)
+            using \<open>connected C\<close> [unfolded connected_closed, simplified, rule_format, of concl: "D \<inter> U1" "D \<inter> U2"]
             using \<open>C \<subseteq> D\<close> D1 D2 V12 DV12 \<open>closed U1\<close> \<open>closed U2\<close> \<open>closed D\<close>
             by blast
         qed
@@ -2648,7 +2614,7 @@
   have "path_component v x x"
     by (meson assms(3) path_component_refl)
   then show ?thesis
-    by (metis assms(1) assms(2) assms(3) mem_Collect_eq path_component_subset path_connected_path_component)
+    by (metis assms mem_Collect_eq path_component_subset path_connected_path_component)
 qed
 
 proposition locally_path_connected:
@@ -3446,10 +3412,11 @@
     apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono])
     using Q by (auto simp: contk imk)
   then show ?thesis
-    apply (rule homotopic_with_eq)
-    using feq apply auto[1]
-    using geq apply auto[1]
-    using Qeq topspace_euclidean_subtopology by blast
+  proof (rule homotopic_with_eq)
+    show "f x = (f \<circ> h \<circ> k) x" "g x = (g \<circ> h \<circ> k) x" 
+      if "x \<in> topspace (top_of_set t)" for x
+      using feq geq that by force+
+  qed (use Qeq topspace_euclidean_subtopology in blast)
 qed
 
 lemma cohomotopically_trivial_retraction_null_gen:
@@ -3470,15 +3437,15 @@
     by (simp add: P Qf contf imf)
   ultimately obtain c where "homotopic_with_canon P s u (f \<circ> h) (\<lambda>x. c)"
     by (metis hom)
-  then have "homotopic_with_canon Q t u (f \<circ> h \<circ> k) ((\<lambda>x. c) \<circ> k)"
-    apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono])
-    using Q by (auto simp: contk imk)
-  then show ?thesis
-    apply (rule_tac c = c in that)
-    apply (erule homotopic_with_eq)
-    using feq apply auto[1]
-    apply simp
-    using Qeq topspace_euclidean_subtopology by blast
+  then have \<section>: "homotopic_with_canon Q t u (f \<circ> h \<circ> k) ((\<lambda>x. c) \<circ> k)"
+  proof (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono])
+    show "\<And>h. \<lbrakk>continuous_map (top_of_set s) (top_of_set u) h; P h\<rbrakk> \<Longrightarrow> Q (h \<circ> k)"
+      using Q by auto
+  qed (auto simp: contk imk)
+  moreover have "homotopic_with_canon Q t u f (\<lambda>x. c)"
+    using homotopic_with_eq [OF \<section>] feq Qeq by fastforce
+  ultimately show ?thesis 
+    using that by blast
 qed
 
 end
@@ -3543,12 +3510,12 @@
                            "homotopic_with (\<lambda>x. True) U U (f2 \<circ> g2) id"
     using 2 by (auto simp: homotopy_equivalent_space_def)
   have "homotopic_with (\<lambda>f. True) X Y (g2 \<circ> f2 \<circ> f1) (id \<circ> f1)"
-    using f1 hom2(1) homotopic_compose_continuous_map_right by blast
+    using f1 hom2(1) homotopic_with_compose_continuous_map_right by metis
   then have "homotopic_with (\<lambda>f. True) X Y (g2 \<circ> (f2 \<circ> f1)) (id \<circ> f1)"
     by (simp add: o_assoc)
   then have "homotopic_with (\<lambda>x. True) X X
          (g1 \<circ> (g2 \<circ> (f2 \<circ> f1))) (g1 \<circ> (id \<circ> f1))"
-    by (simp add: g1 homotopic_compose_continuous_map_left)
+    by (simp add: g1 homotopic_with_compose_continuous_map_left)
   moreover have "homotopic_with (\<lambda>x. True) X X (g1 \<circ> id \<circ> f1) id"
     using hom1 by simp
   ultimately have SS: "homotopic_with (\<lambda>x. True) X X (g1 \<circ> g2 \<circ> (f2 \<circ> f1)) id"
@@ -3607,7 +3574,7 @@
       have "homotopic_with (\<lambda>x. True) X X f r"
         proof (rule homotopic_with_eq)
           show "homotopic_with (\<lambda>x. True) X X (r \<circ> f) (r \<circ> id)"
-            using homotopic_with_symD continuous_map_into_fulltopology f homotopic_compose_continuous_map_left r by blast
+            by (metis continuous_map_into_fulltopology f homotopic_with_compose_continuous_map_left homotopic_with_symD r)
           show "f x = (r \<circ> f) x" if "x \<in> topspace X" for x
             using that fS req by auto
         qed auto
@@ -3637,7 +3604,7 @@
 
 lemma contractible_space_empty:
    "topspace X = {} \<Longrightarrow> contractible_space X"
-  apply (simp add: contractible_space_def homotopic_with_def)
+  unfolding contractible_space_def homotopic_with_def
   apply (rule_tac x=undefined in exI)
   apply (rule_tac x="\<lambda>(t,x). if t = 0 then x else undefined" in exI)
   apply (auto simp: continuous_map_on_empty)
@@ -3645,7 +3612,7 @@
 
 lemma contractible_space_singleton:
   "topspace X = {a} \<Longrightarrow> contractible_space X"
-  apply (simp add: contractible_space_def homotopic_with_def)
+  unfolding contractible_space_def homotopic_with_def
   apply (rule_tac x=a in exI)
   apply (rule_tac x="\<lambda>(t,x). if t = 0 then x else a" in exI)
   apply (auto intro: continuous_map_eq [where f = "\<lambda>z. a"])
@@ -3666,8 +3633,7 @@
 proof (cases "topspace X = {}")
   case False
   then show ?thesis
-    apply (auto simp: contractible_space_def)
-    using homotopic_with_imp_continuous_maps by fastforce
+    using homotopic_with_imp_continuous_maps  by (fastforce simp: contractible_space_def)
 qed (simp add: contractible_space_empty)
 
 lemma contractible_imp_path_connected_space:
@@ -3680,12 +3646,16 @@
     for a and h :: "real \<times> 'a \<Rightarrow> 'a"
   proof -
     have "path_component_of X b a" if "b \<in> topspace X" for b
-      using that unfolding path_component_of_def
-      apply (rule_tac x="h \<circ> (\<lambda>x. (x,b))" in exI)
-      apply (simp add: h pathin_def)
-      apply (rule continuous_map_compose [OF _ conth])
-      apply (auto simp: continuous_map_pairwise intro!: continuous_intros continuous_map_compose continuous_map_id [unfolded id_def])
-      done
+      unfolding path_component_of_def
+    proof (intro exI conjI)
+      let ?g = "h \<circ> (\<lambda>x. (x,b))"
+      show "pathin X ?g"
+        unfolding pathin_def
+        apply (rule continuous_map_compose [OF _ conth])
+        using that
+        apply (auto simp: continuous_map_pairwise intro!: continuous_intros continuous_map_compose continuous_map_id [unfolded id_def])
+        done
+    qed (use h in auto)
   then show ?thesis
     by (metis path_component_of_equiv path_connected_space_iff_path_component)
   qed
@@ -3706,15 +3676,16 @@
   show ?rhs
   proof
     show "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. b)" if "b \<in> topspace X" for b
-      apply (rule homotopic_with_trans [OF a])
-      using homotopic_constant_maps path_connected_space_imp_path_component_of
-      by (metis (full_types) X a continuous_map_const contractible_imp_path_connected_space homotopic_with_imp_continuous_maps that)
+    proof (rule homotopic_with_trans [OF a])
+      show "homotopic_with (\<lambda>x. True) X X (\<lambda>x. a) (\<lambda>x. b)"
+        using homotopic_constant_maps path_connected_space_imp_path_component_of
+        by (metis (full_types) X a continuous_map_const contractible_imp_path_connected_space homotopic_with_imp_continuous_maps that)
+    qed
   qed
 next
   assume R: ?rhs
   then show ?lhs
-    apply (simp add: contractible_space_def)
-    by (metis equals0I homotopic_on_emptyI)
+    unfolding contractible_space_def by (metis equals0I homotopic_on_emptyI)
 qed
 
 
@@ -3728,9 +3699,9 @@
   obtain b where b: "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. b)"
     using Y by (auto simp: contractible_space_def)
   show thesis
-    using homotopic_compose_continuous_map_right
-           [OF homotopic_compose_continuous_map_left [OF b g] f]
-    by (simp add: that)
+    using homotopic_with_compose_continuous_map_right
+           [OF homotopic_with_compose_continuous_map_left [OF b g] f]
+    by (force simp add: that)
 qed
 
 lemma nullhomotopic_into_contractible_space:
@@ -3753,7 +3724,7 @@
   obtain c where c: "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)"
     using nullhomotopic_from_contractible_space [OF f X] .
   have "homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) (\<lambda>x. c)"
-    using homotopic_compose_continuous_map_right [OF c g] by fastforce
+    using homotopic_with_compose_continuous_map_right [OF c g] by fastforce
   then have "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. c)"
     using homotopic_with_trans [OF _ hom] homotopic_with_symD by blast
   then show ?thesis
@@ -3808,7 +3779,7 @@
   have "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. f a)"
   proof (rule homotopic_with_eq)
     show "homotopic_with (\<lambda>x. True) Y Y (f \<circ> id \<circ> g) (f \<circ> (\<lambda>x. a) \<circ> g)"
-      using f g a homotopic_compose_continuous_map_left homotopic_compose_continuous_map_right by metis
+      using f g a homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right by metis
   qed (use fg in auto)
   then show ?thesis
     unfolding contractible_space_def by blast
@@ -3907,12 +3878,11 @@
     show ?thesis
       unfolding contractible_space_def homotopic_with_def
     proof (intro exI conjI allI)
+      note \<section> = convexD [OF \<open>convex S\<close>, simplified]
       show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S)
                            (\<lambda>(t,x). (1 - t) * x + t * z)"
-        apply (auto simp: case_prod_unfold)
-         apply (intro continuous_intros)
-        using  \<open>z \<in> S\<close> apply (auto intro: convexD [OF \<open>convex S\<close>, simplified])
-        done
+        using  \<open>z \<in> S\<close> 
+        by (auto simp add: case_prod_unfold intro!: continuous_intros \<section>)
     qed auto
   qed (simp add: contractible_space_empty)
 qed
@@ -3947,14 +3917,12 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "linear f" "inj f"
     shows "(f ` S) homotopy_eqv S"
-apply (rule homeomorphic_imp_homotopy_eqv)
-using assms homeomorphic_sym linear_homeomorphic_image by auto
+  by (metis assms homeomorphic_sym homeomorphic_imp_homotopy_eqv linear_homeomorphic_image)
 
 lemma homotopy_eqv_translation:
     fixes S :: "'a::real_normed_vector set"
     shows "(+) a ` S homotopy_eqv S"
-  apply (rule homeomorphic_imp_homotopy_eqv)
-  using homeomorphic_translation homeomorphic_sym by blast
+  using homeomorphic_imp_homotopy_eqv homeomorphic_translation homeomorphic_sym by blast
 
 lemma homotopy_eqv_homotopic_triviality_imp:
   fixes S :: "'a::real_normed_vector set"
@@ -3974,26 +3942,19 @@
                         "homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id"
     using assms by (auto simp: homotopy_equivalent_space_def)
   have "homotopic_with_canon (\<lambda>f. True) U S (k \<circ> f) (k \<circ> g)"
-    apply (rule homUS)
-    using f g k
-    apply (safe intro!: continuous_on_compose h k f elim!: continuous_on_subset)
-    apply (force simp: o_def)+
-    done
+  proof (rule homUS)
+    show "continuous_on U (k \<circ> f)" "continuous_on U (k \<circ> g)"
+      using continuous_on_compose continuous_on_subset f g k by blast+
+  qed (use f g k in \<open>(force simp: o_def)+\<close> )
   then have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))"
-    apply (rule homotopic_with_compose_continuous_left)
-    apply (simp_all add: h)
-    done
+    by (rule homotopic_with_compose_continuous_left; simp add: h)
   moreover have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> k \<circ> f) (id \<circ> f)"
-    apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T])
-    apply (auto simp: hom f)
-    done
+    by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom f)
   moreover have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> k \<circ> g) (id \<circ> g)"
-    apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T])
-    apply (auto simp: hom g)
-    done
+    by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom g)
   ultimately show "homotopic_with_canon (\<lambda>x. True) U T f g"
-    apply (simp add: o_assoc)
-    using homotopic_with_trans homotopic_with_sym by blast
+    unfolding o_assoc
+    by (metis homotopic_with_trans homotopic_with_sym id_comp) 
 qed
 
 lemma homotopy_eqv_homotopic_triviality:
@@ -4038,18 +3999,15 @@
                         "homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id"
     using assms by (auto simp: homotopy_equivalent_space_def)
   obtain c where "homotopic_with_canon (\<lambda>x. True) S U (f \<circ> h) (\<lambda>x. c)"
-    apply (rule exE [OF homSU [of "f \<circ> h"]])
-    apply (intro continuous_on_compose h)
-    using h f  apply (force elim!: continuous_on_subset)+
-    done
+  proof (rule exE [OF homSU])
+    show "continuous_on S (f \<circ> h)"
+      using continuous_on_compose continuous_on_subset f h by blast
+  qed (use f h in force)
   then have "homotopic_with_canon (\<lambda>x. True) T U ((f \<circ> h) \<circ> k) ((\<lambda>x. c) \<circ> k)"
-    apply (rule homotopic_with_compose_continuous_right [where X=S])
-    using k by auto
+    by (rule homotopic_with_compose_continuous_right [where X=S]) (use k in auto)
   moreover have "homotopic_with_canon (\<lambda>x. True) T U (f \<circ> id) (f \<circ> (h \<circ> k))"
-    apply (rule homotopic_with_compose_continuous_left [where Y=T])
-      apply (simp add: hom homotopic_with_symD)
-     using f apply auto
-    done
+    by (rule homotopic_with_compose_continuous_left [where Y=T])
+       (use f in \<open>auto simp add: hom homotopic_with_symD\<close>)
   ultimately show ?thesis
     apply (rule_tac c=c in that)
     apply (simp add: o_def)
@@ -4065,10 +4023,9 @@
                 \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S U f (\<lambda>x. c))) \<longleftrightarrow>
            (\<forall>f. continuous_on T f \<and> f ` T \<subseteq> U
                 \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) T U f (\<lambda>x. c)))"
-apply (rule iffI)
-apply (metis assms homotopy_eqv_cohomotopic_triviality_null_imp)
-by (metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_equivalent_space_sym)
-
+by (rule iffI; metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_equivalent_space_sym)
+
+text \<open>Similar to the proof above\<close>
 lemma homotopy_eqv_homotopic_triviality_null_imp:
   fixes S :: "'a::real_normed_vector set"
     and T :: "'b::real_normed_vector set"
@@ -4085,18 +4042,15 @@
                         "homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id"
     using assms by (auto simp: homotopy_equivalent_space_def)
   obtain c::'a where "homotopic_with_canon (\<lambda>x. True) U S (k \<circ> f) (\<lambda>x. c)"
-    apply (rule exE [OF homSU [of "k \<circ> f"]])
-    apply (intro continuous_on_compose h)
-    using k f  apply (force elim!: continuous_on_subset)+
-    done
+  proof (rule exE [OF homSU [of "k \<circ> f"]])
+    show "continuous_on U (k \<circ> f)"
+      using continuous_on_compose continuous_on_subset f k by blast
+  qed (use f k in force)
   then have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))"
-    apply (rule homotopic_with_compose_continuous_left [where Y=S])
-    using h by auto
+    by (rule homotopic_with_compose_continuous_left [where Y=S]) (use h in auto)
   moreover have "homotopic_with_canon (\<lambda>x. True) U T (id \<circ> f) ((h \<circ> k) \<circ> f)"
-    apply (rule homotopic_with_compose_continuous_right [where X=T])
-      apply (simp add: hom homotopic_with_symD)
-     using f apply auto
-    done
+    by (rule homotopic_with_compose_continuous_right [where X=T])
+       (use f in \<open>auto simp add: hom homotopic_with_symD\<close>)
   ultimately show ?thesis
     using homotopic_with_trans by (fastforce simp add: o_def)
 qed
@@ -4110,9 +4064,7 @@
                   \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. c))) \<longleftrightarrow>
            (\<forall>f. continuous_on U f \<and> f ` U \<subseteq> T
                   \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) U T f (\<lambda>x. c)))"
-apply (rule iffI)
-apply (metis assms homotopy_eqv_homotopic_triviality_null_imp)
-by (metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_equivalent_space_sym)
+by (rule iffI; metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_equivalent_space_sym)
 
 lemma homotopy_eqv_contractible_sets:
   fixes S :: "'a::real_normed_vector set"
--- a/src/HOL/Analysis/Improper_Integral.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Improper_Integral.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -445,7 +445,7 @@
     obtain u v where K: "K = cbox u v" "K \<noteq> {}" "K \<subseteq> cbox a b"
       using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
     with i show "extend K \<noteq> {}" "extend K \<subseteq> cbox a b"
-      apply (auto simp: extend_def subset_box box_ne_empty sum_if_inner)
+      apply (auto simp: extend_def subset_box box_ne_empty)
       by fastforce
   qed
   have int_extend_disjoint:
@@ -468,7 +468,7 @@
        and xv: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> x \<bullet> k < v \<bullet> k"
        and wx: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> w \<bullet> k < x \<bullet> k"
        and xz: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> x \<bullet> k < z \<bullet> k"
-        using that K1 K2 i by (auto simp: extend_def box_ne_empty sum_if_inner mem_box)
+        using that K1 K2 i by (auto simp: extend_def box_ne_empty mem_box)
       have "box u v \<noteq> {}" "box w z \<noteq> {}"
         using cboxes interior_cbox by (auto simp: content_eq_0_interior dest: mt)
       then obtain q s
@@ -630,7 +630,7 @@
     show "\<Union>Dlec \<subseteq> cbox a b'"
       using Dlec_def div S by (auto simp: b'_def division_of_def mem_box)
     show "(\<forall>K\<in>Dlec. K \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}) \<or> (\<forall>K\<in>Dlec. K \<inter> {x. x \<bullet> i = b' \<bullet> i} \<noteq> {})"
-      using nonmt by (fastforce simp: Dlec_def b'_def sum_if_inner i)
+      using nonmt by (fastforce simp: Dlec_def b'_def i)
   qed (use i Dlec_def in auto)
   moreover
   have "(\<Sum>K\<in>Dlec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) =
@@ -640,7 +640,7 @@
     unfolding Dlec_def using \<open>finite \<D>\<close> apply (auto simp: sum.mono_neutral_left)
     done
   moreover have "(b' \<bullet> i - a \<bullet> i) = (c - a \<bullet> i)"
-    by (simp add: b'_def sum_if_inner i)
+    by (simp add: b'_def i)
   ultimately
   have lec: "(c - a \<bullet> i) * (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)
              \<le> content(cbox a b')"
@@ -657,7 +657,7 @@
     show "\<Union>Dgec \<subseteq> cbox a' b"
       using Dgec_def div S by (auto simp: a'_def division_of_def mem_box)
     show "(\<forall>K\<in>Dgec. K \<inter> {x. x \<bullet> i = a' \<bullet> i} \<noteq> {}) \<or> (\<forall>K\<in>Dgec. K \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {})"
-      using nonmt by (fastforce simp: Dgec_def a'_def sum_if_inner i)
+      using nonmt by (fastforce simp: Dgec_def a'_def i)
   qed (use i Dgec_def in auto)
   moreover
   have "(\<Sum>K\<in>Dgec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) =
@@ -667,7 +667,7 @@
     unfolding Dgec_def using \<open>finite \<D>\<close> apply (auto simp: sum.mono_neutral_left)
     done
   moreover have "(b \<bullet> i - a' \<bullet> i) = (b \<bullet> i - c)"
-    by (simp add: a'_def sum_if_inner i)
+    by (simp add: a'_def i)
   ultimately
   have gec: "(b \<bullet> i - c) * (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)
              \<le> content(cbox a' b)"
@@ -679,7 +679,7 @@
     proof
       assume c: "c = a \<bullet> i"
       then have "a' = a"
-        apply (simp add: sum_if_inner i a'_def cong: if_cong)
+        apply (simp add: i a'_def cong: if_cong)
         using euclidean_representation [of a] sum.cong [OF refl, of Basis "\<lambda>i. (a \<bullet> i) *\<^sub>R i"] by presburger
       then have "content (cbox a' b) \<le> 2 * content (cbox a b)"  by simp
       moreover
@@ -701,7 +701,7 @@
     next
       assume c: "c = b \<bullet> i"
       then have "b' = b"
-        apply (simp add: sum_if_inner i b'_def cong: if_cong)
+        apply (simp add: i b'_def cong: if_cong)
         using euclidean_representation [of b] sum.cong [OF refl, of Basis "\<lambda>i. (b \<bullet> i) *\<^sub>R i"] by presburger
       then have "content (cbox a b') \<le> 2 * content (cbox a b)"  by simp
       moreover
@@ -726,14 +726,14 @@
     have prod_if: "(\<Prod>k\<in>Basis \<inter> - {i}. f k) = (\<Prod>k\<in>Basis. f k) / f i" if "f i \<noteq> (0::real)" for f
       using that mk_disjoint_insert [OF i]
       apply (clarsimp simp add: field_split_simps)
-      by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf mult.commute order_refl prod.insert subset_Compl_singleton)
+      by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf order_refl prod.insert subset_Compl_singleton)
     have abc: "a \<bullet> i < c" "c < b \<bullet> i"
       using False assms by auto
     then have "(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)
                   \<le> content(cbox a b') / (c - a \<bullet> i)"
               "(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)
                  \<le> content(cbox a' b) / (b \<bullet> i - c)"
-      using lec gec by (simp_all add: field_split_simps mult.commute)
+      using lec gec by (simp_all add: field_split_simps)
     moreover
     have "(\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
           \<le> (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K) +
@@ -780,7 +780,7 @@
           \<le> 2 * content (cbox a b) / (b \<bullet> i - a \<bullet> i)"
       by linarith
     then show ?thesis
-      using abc by (simp add: field_split_simps mult.commute)
+      using abc by (simp add: field_split_simps)
   qed
 qed
 
@@ -856,7 +856,7 @@
   have "gauge (\<lambda>x. ball x
                     (\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b))))"
     using \<open>0 < content (cbox a b)\<close> \<open>0 < \<epsilon>\<close> a_less_b
-    apply (auto simp: gauge_def field_split_simps mult_less_0_iff zero_less_mult_iff add_nonneg_eq_0_iff finite_less_Inf_iff)
+    apply (auto simp: gauge_def field_split_simps add_nonneg_eq_0_iff finite_less_Inf_iff)
     apply (meson add_increasing measure_nonneg mult_nonneg_nonneg norm_ge_zero not_le zero_le_numeral)
     done
   then have "gauge \<gamma>"
@@ -936,7 +936,7 @@
               by simp
             also have "... < \<epsilon> * (b \<bullet> i - a \<bullet> i) / dist u v / (4 * content (cbox a b))"
               using duv dist_uv contab_gt0
-              apply (simp add: field_split_simps algebra_simps mult_less_0_iff zero_less_mult_iff split: if_split_asm)
+              apply (simp add: field_split_simps split: if_split_asm)
               by (meson add_nonneg_nonneg linorder_not_le measure_nonneg mult_nonneg_nonneg norm_ge_zero zero_le_numeral)
             also have "... = \<epsilon> * (b \<bullet> i - a \<bullet> i) / norm (v - u) / (4 * content (cbox a b))"
               by (simp add: dist_norm norm_minus_commute)
@@ -1239,7 +1239,7 @@
                 obtain u v where uv: "K = cbox u v"
                   using T''_tagged \<open>(x,K) \<in> T''\<close> by blast
                 then have "connected K"
-                  by (simp add: is_interval_cbox is_interval_connected)
+                  by (simp add: is_interval_connected)
                 then have "(\<exists>z \<in> K. z \<bullet> i = c)"
                   using y connected_ivt_component by fastforce
                 then show ?thesis
@@ -1883,7 +1883,7 @@
   then obtain m::nat where m: "floor(n * f x) = int m"
     using nonneg_int_cases zero_le_floor by blast
   then have kn: "real k / real n \<le> f x \<longleftrightarrow> k \<le> m" for k
-    using \<open>n \<noteq> 0\<close> by (simp add: field_split_simps mult.commute) linarith
+    using \<open>n \<noteq> 0\<close> by (simp add: field_split_simps) linarith
   then have "Suc n / real n \<le> f x \<longleftrightarrow> Suc n \<le> m"
     by blast
   have "real n * f x \<le> real n"
@@ -1894,7 +1894,7 @@
     by (subst sum.inter_restrict) (auto simp: kn)
   also have "\<dots> < inverse n"
     using \<open>m \<le> n\<close> \<open>n \<noteq> 0\<close> m
-    by (simp add: min_absorb2 field_split_simps mult.commute) linarith
+    by (simp add: min_absorb2 field_split_simps) linarith
   finally show ?thesis .
 qed
 
@@ -2222,7 +2222,7 @@
       proof (subst has_integral_cong)
         show "g x * f x - g a * f x = (g b - g a) * h x *\<^sub>R f x"
           if "x \<in> {a..b}" for x
-          using 1 that by (simp add: h_def field_split_simps algebra_simps)
+          using 1 that by (simp add: h_def field_split_simps)
         show "((\<lambda>x. (g b - g a) * h x *\<^sub>R f x) has_integral (g b - g a) * integral {c..b} f) {a..b}"
           using has_integral_mult_right [OF c, of "g b - g a"] .
       qed
--- a/src/HOL/Analysis/Infinite_Set_Sum.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Infinite_Set_Sum.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -154,7 +154,7 @@
 lemma abs_summable_on_altdef':
   "A \<subseteq> B \<Longrightarrow> f abs_summable_on A \<longleftrightarrow> set_integrable (count_space B) A f"
   unfolding abs_summable_on_def set_integrable_def
-  by (metis (no_types) Pow_iff abs_summable_on_def inf.orderE integrable_restrict_space restrict_count_space_subset set_integrable_def sets_count_space space_count_space)
+  by (metis (no_types) Pow_iff abs_summable_on_def inf.orderE integrable_restrict_space restrict_count_space_subset sets_count_space space_count_space)
 
 lemma abs_summable_on_norm_iff [simp]:
   "(\<lambda>x. norm (f x)) abs_summable_on A \<longleftrightarrow> f abs_summable_on A"
--- a/src/HOL/Analysis/Jordan_Curve.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Jordan_Curve.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -624,7 +624,7 @@
         then show "connected (- closure (inside (?\<Theta>1 \<union> ?\<Theta>)))"
           by (metis Compl_Un outside_inside co_out1c closure_Un_frontier)
         have if2: "- (inside (?\<Theta>2 \<union> ?\<Theta>) \<union> frontier (inside (?\<Theta>2 \<union> ?\<Theta>))) = - ?\<Theta>2 \<inter> - ?\<Theta> \<inter> - inside (?\<Theta>2 \<union> ?\<Theta>)"
-          by (metis (no_types, lifting) Int_commute Jordan_inside_outside c c2 compl_sup path_image_join path_image_reversepath pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath sp closure_Un_frontier fr_out(2))
+          by (metis (no_types, lifting) Int_commute Jordan_inside_outside c c2 compl_sup path_image_join path_image_reversepath pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath sp(3) closure_Un_frontier fr_out(2))
         then show "connected (- closure (inside (?\<Theta>2 \<union> ?\<Theta>)))"
           by (metis Compl_Un outside_inside co_out2c closure_Un_frontier)
         have "connected(?\<Theta>)"
--- a/src/HOL/Analysis/Lipschitz.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Lipschitz.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -684,13 +684,16 @@
       from this elim d[of "rx (ry (rt n))"]
       have "\<dots> < dist (f (?t n) (?y n)) (f (?t n) (?x n)) / rx (ry (rt (n)))"
         using lx'(2) ly'(2) lt'(2) \<open>0 < rx _\<close>
-        by (auto simp add: field_split_simps algebra_simps strict_mono_def)
+        by (auto simp add: field_split_simps strict_mono_def)
       also have "\<dots> \<le> diameter ?S / n"
-        by (force intro!: \<open>0 < n\<close> strict_mono_def xy diameter_bounded_bound frac_le
-          compact_imp_bounded compact t
-          intro: le_trans[OF seq_suble[OF lt'(2)]]
-            le_trans[OF seq_suble[OF ly'(2)]]
-            le_trans[OF seq_suble[OF lx'(2)]])
+      proof (rule frac_le)
+        show "diameter ?S \<ge> 0"
+          using compact compact_imp_bounded diameter_ge_0 by blast
+        show "dist (f (?t n) (?y n)) (f (?t n) (?x n)) \<le> diameter ((\<lambda>(t,x). f t x) ` (T \<times> X))"
+          by (metis (no_types) compact compact_imp_bounded diameter_bounded_bound image_eqI mem_Sigma_iff o_apply split_conv t xy(1) xy(2))
+        show "real n \<le> real (rx (ry (rt n)))"
+          by (meson le_trans lt'(2) lx'(2) ly'(2) of_nat_mono strict_mono_imp_increasing)
+      qed (use \<open>n > 0\<close> in auto)
       also have "\<dots> \<le> abs (diameter ?S) / n"
         by (auto intro!: divide_right_mono)
       finally show ?case by simp
--- a/src/HOL/Analysis/Measure_Space.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Measure_Space.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1574,7 +1574,7 @@
   by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def)
 
 lemma measure_zero_top: "emeasure M A = top \<Longrightarrow> measure M A = 0"
-  by (simp add: measure_def enn2ereal_top)
+  by (simp add: measure_def)
 
 lemma measure_eq_emeasure_eq_ennreal: "0 \<le> x \<Longrightarrow> emeasure M A = ennreal x \<Longrightarrow> measure M A = x"
   using emeasure_eq_ennreal_measure[of M A]
--- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -166,7 +166,7 @@
     (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
     by (auto intro!: sum.cong)
   also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
-    using assms by (auto dest: simple_functionD simp: sum.delta)
+    using assms by (auto dest: simple_functionD)
   also have "... = f x" using x by (auto simp: indicator_def)
   finally show ?thesis by auto
 qed
@@ -327,10 +327,10 @@
     { fix d :: nat
       have "\<lfloor>2^d::real\<rfloor> * \<lfloor>2^m * enn2real (min (of_nat m) (u x))\<rfloor> \<le>
         \<lfloor>2^d * (2^m * enn2real (min (of_nat m) (u x)))\<rfloor>"
-        by (rule le_mult_floor) (auto simp: enn2real_nonneg)
+        by (rule le_mult_floor) (auto)
       also have "\<dots> \<le> \<lfloor>2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))\<rfloor>"
         by (intro floor_mono mult_mono enn2real_mono min.mono)
-           (auto simp: enn2real_nonneg min_less_iff_disj of_nat_less_top)
+           (auto simp: min_less_iff_disj of_nat_less_top)
       finally have "f m x \<le> f (m + d) x"
         unfolding f_def
         by (auto simp: field_simps power_add * simp del: of_int_mult) }
@@ -348,7 +348,7 @@
       by (cases "u x" rule: ennreal_cases)
          (auto split: split_min intro!: floor_mono)
     then have "f i ` space M \<subseteq> (\<lambda>n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
-      unfolding floor_of_int by (auto simp: f_def enn2real_nonneg intro!: imageI)
+      unfolding floor_of_int by (auto simp: f_def intro!: imageI)
     then show "finite (f i ` space M)"
       by (rule finite_subset) auto
     show "f i \<in> borel_measurable M"
@@ -680,7 +680,7 @@
 proof cases
   assume "finite P"
   from this assms show ?thesis
-    by induct (auto simp: simple_function_sum simple_integral_add sum_nonneg)
+    by induct (auto)
 qed auto
 
 lemma simple_integral_mult[simp]:
@@ -1090,8 +1090,7 @@
   by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator)
 
 lemma nn_integral_cmult_indicator: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * emeasure M A"
-  by (subst nn_integral_eq_simple_integral)
-     (auto simp: simple_function_indicator simple_integral_indicator)
+  by (subst nn_integral_eq_simple_integral) (auto)
 
 lemma nn_integral_indicator':
   assumes [measurable]: "A \<inter> space M \<in> sets M"
@@ -1120,7 +1119,7 @@
 lemma nn_integral_indicator_singleton'[simp]:
   assumes [measurable]: "{y} \<in> sets M"
   shows "(\<integral>\<^sup>+x. ennreal (f x * indicator {y} x) \<partial>M) = f y * emeasure M {y}"
-  by (subst nn_integral_set_ennreal[symmetric]) (simp add: nn_integral_indicator_singleton)
+  by (subst nn_integral_set_ennreal[symmetric]) (simp)
 
 lemma nn_integral_add:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g"
--- a/src/HOL/Analysis/Ordered_Euclidean_Space.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -48,7 +48,7 @@
 
 lemma inner_Basis_inf_left: "i \<in> Basis \<Longrightarrow> inf x y \<bullet> i = inf (x \<bullet> i) (y \<bullet> i)"
   and inner_Basis_sup_left: "i \<in> Basis \<Longrightarrow> sup x y \<bullet> i = sup (x \<bullet> i) (y \<bullet> i)"
-  by (simp_all add: eucl_inf eucl_sup inner_sum_left inner_Basis if_distrib comm_monoid_add_class.sum.delta
+  by (simp_all add: eucl_inf eucl_sup inner_sum_left inner_Basis if_distrib
       cong: if_cong)
 
 lemma inner_Basis_INF_left: "i \<in> Basis \<Longrightarrow> (INF x\<in>X. f x) \<bullet> i = (INF x\<in>X. f x \<bullet> i)"
@@ -118,7 +118,7 @@
   hence "Inf ?proj = x \<bullet> b"
     by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)
   hence "x \<bullet> b = Inf X \<bullet> b"
-    by (auto simp: eucl_Inf inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta
+    by (auto simp: eucl_Inf inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close>
       cong: if_cong)
   with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Inf X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> x \<bullet> b \<le> y \<bullet> b)" by blast
 qed
@@ -140,7 +140,7 @@
   hence "Sup ?proj = x \<bullet> b"
     by (auto intro!: cSup_eq_maximum)
   hence "x \<bullet> b = Sup X \<bullet> b"
-    by (auto simp: eucl_Sup[where 'a='a] inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta
+    by (auto simp: eucl_Sup[where 'a='a] inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close>
       cong: if_cong)
   with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Sup X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> y \<bullet> b \<le> x \<bullet> b)" by blast
 qed
--- a/src/HOL/Analysis/Path_Connected.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Path_Connected.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -511,7 +511,7 @@
    } note ** = this
   show ?thesis
     using assms
-    apply (simp add: arc_def simple_path_def path_join, clarify)
+    apply (simp add: arc_def simple_path_def, clarify)
     apply (simp add: joinpaths_def split: if_split_asm)
     apply (force dest: inj_onD [OF injg1])
     apply (metis *)
--- a/src/HOL/Analysis/Polytope.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Polytope.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -994,13 +994,13 @@
    "a \<noteq> 0 \<Longrightarrow> {x. a \<bullet> x = b} facet_of {x. a \<bullet> x \<le> b}"
 unfolding facet_of_def hyperplane_eq_empty
 by (auto simp: hyperplane_face_of_halfspace_ge hyperplane_face_of_halfspace_le
-           DIM_positive Suc_leI of_nat_diff aff_dim_halfspace_le)
+           Suc_leI of_nat_diff aff_dim_halfspace_le)
 
 lemma hyperplane_facet_of_halfspace_ge:
     "a \<noteq> 0 \<Longrightarrow> {x. a \<bullet> x = b} facet_of {x. a \<bullet> x \<ge> b}"
 unfolding facet_of_def hyperplane_eq_empty
 by (auto simp: hyperplane_face_of_halfspace_le hyperplane_face_of_halfspace_ge
-           DIM_positive Suc_leI of_nat_diff aff_dim_halfspace_ge)
+           Suc_leI of_nat_diff aff_dim_halfspace_ge)
 
 lemma facet_of_halfspace_le:
     "F facet_of {x. a \<bullet> x \<le> b} \<longleftrightarrow> a \<noteq> 0 \<and> F = {x. a \<bullet> x = b}"
@@ -1316,7 +1316,7 @@
     apply (rule_tac x="u/x" in exI)
     apply (rule_tac x="v/x" in exI)
     apply (rule_tac x="w/x" in exI)
-    using x apply (auto simp: algebra_simps field_split_simps)
+    using x apply (auto simp: field_split_simps)
     done
   ultimately show ?thesis by force
 qed
@@ -1445,14 +1445,14 @@
               have "x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c) /\<^sub>R ux"
                 by (metis \<open>ux \<noteq> 0\<close> uxx mult.commute right_inverse scaleR_one scaleR_scaleR)
               also have "... = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c"
-                using \<open>ux \<noteq> 0\<close> equx apply (auto simp: algebra_simps field_split_simps)
+                using \<open>ux \<noteq> 0\<close> equx apply (auto simp: field_split_simps)
                 by (metis add.commute add_diff_eq add_divide_distrib diff_add_cancel scaleR_add_left)
               finally have "x = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" .
               then have "x \<in> open_segment b c"
                 apply (simp add: in_segment \<open>b \<noteq> c\<close>)
                 apply (rule_tac x="(v * uc) / ux" in exI)
                 using \<open>0 \<le> ux\<close> \<open>ux \<noteq> 0\<close> \<open>0 < uc\<close> \<open>0 < v\<close> \<open>0 < ub\<close> \<open>v < 1\<close> equx
-                apply (force simp: algebra_simps field_split_simps)
+                apply (force simp: field_split_simps)
                 done
               then show ?thesis
                 by (rule face_ofD [OF F _ b c \<open>x \<in> F\<close>])
@@ -2116,7 +2116,7 @@
               done
             then show ?thesis
               using \<open>0 < inff\<close> awlt [OF that] mult_strict_left_mono
-              by (fastforce simp add: algebra_simps field_split_simps split: if_split_asm)
+              by (fastforce simp add: field_split_simps split: if_split_asm)
           next
             case False
             with \<open>0 < inff\<close> have "inff * (a j \<bullet> y - a j \<bullet> w) \<le> 0"
--- a/src/HOL/Analysis/Product_Topology.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Product_Topology.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -290,7 +290,7 @@
 next
   case False
   then show ?thesis
-    by (simp add: continuous_compose_quotient_map_eq quotient_map_fst)
+    by (simp add: continuous_compose_quotient_map_eq)
 qed
 
 lemma continuous_map_of_snd:
@@ -302,7 +302,7 @@
 next
   case False
   then show ?thesis
-    by (simp add: continuous_compose_quotient_map_eq quotient_map_snd)
+    by (simp add: continuous_compose_quotient_map_eq)
 qed
 
 lemma continuous_map_prod_top:
--- a/src/HOL/Analysis/Regularity.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Regularity.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -53,7 +53,7 @@
     hence "1/n > 0" "e * 2 powr - n > 0" by (auto)
     from M_space[OF \<open>1/n>0\<close>]
     have "(\<lambda>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) \<longlonglongrightarrow> measure M (space M)"
-      unfolding emeasure_eq_measure by (auto simp: measure_nonneg)
+      unfolding emeasure_eq_measure by (auto)
     from metric_LIMSEQ_D[OF this \<open>0 < e * 2 powr -n\<close>]
     obtain k where "dist (measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) (measure M (space M)) <
       e * 2 powr -n"
@@ -95,7 +95,7 @@
     also have "\<dots> \<le> (\<Sum>n. emeasure M (space M - B n))"
       by (rule emeasure_subadditive_countably) (auto simp: summable_def)
     also have "\<dots> \<le> (\<Sum>n. ennreal (e*2 powr - real (Suc n)))"
-      using B_compl_le by (intro suminf_le) (simp_all add: measure_nonneg emeasure_eq_measure ennreal_leI)
+      using B_compl_le by (intro suminf_le) (simp_all add: emeasure_eq_measure ennreal_leI)
     also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
       by (simp add: powr_minus powr_realpow field_simps del: of_nat_Suc)
     also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^ Suc n))"
@@ -268,7 +268,7 @@
     also have "(\<lambda>n. \<Sum>i<n. M (D i)) \<longlonglongrightarrow> (\<Sum>i. M (D i))"
       by (intro summable_LIMSEQ) auto
     finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i<n. measure M (D i)) \<longlonglongrightarrow> measure M (\<Union>i. D i)"
-      by (simp add: emeasure_eq_measure measure_nonneg sum_nonneg)
+      by (simp add: emeasure_eq_measure sum_nonneg)
     have "(\<Union>i. D i) \<in> sets M" using \<open>range D \<subseteq> sets M\<close> by auto
 
     case 1
@@ -282,12 +282,12 @@
       then obtain n0 where n0: "\<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>i. D i)\<bar> < e/2"
         unfolding choice_iff by blast
       have "ennreal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))"
-        by (auto simp add: emeasure_eq_measure sum_nonneg measure_nonneg)
+        by (auto simp add: emeasure_eq_measure)
       also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule sum_le_suminf) auto
       also have "\<dots> = M (\<Union>i. D i)" by (simp add: M)
       also have "\<dots> = measure M (\<Union>i. D i)" by (simp add: emeasure_eq_measure)
       finally have n0: "measure M (\<Union>i. D i) - (\<Sum>i<n0. measure M (D i)) < e/2"
-        using n0 by (auto simp: measure_nonneg sum_nonneg)
+        using n0 by (auto simp: sum_nonneg)
       have "\<forall>i. \<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
       proof
         fix i
@@ -351,7 +351,7 @@
       have "ennreal (measure M ?U - measure M (\<Union>i. D i)) = M ?U - M (\<Union>i. D i)"
         using U(1,2)
         by (subst ennreal_minus[symmetric])
-           (auto intro!: finite_measure_mono simp: sb measure_nonneg emeasure_eq_measure)
+           (auto intro!: finite_measure_mono simp: sb emeasure_eq_measure)
       also have "\<dots> = M (?U - (\<Union>i. D i))" using U  \<open>(\<Union>i. D i) \<in> sets M\<close>
         by (subst emeasure_Diff) (auto simp: sb)
       also have "\<dots> \<le> M (\<Union>i. U i - D i)" using U  \<open>range D \<subseteq> sets M\<close>
@@ -361,7 +361,7 @@
       also have "\<dots> \<le> (\<Sum>i. ennreal e/(2 powr Suc i))" using U \<open>range D \<subseteq> sets M\<close>
         using \<open>0<e\<close>
         by (intro suminf_le, subst emeasure_Diff)
-           (auto simp: emeasure_Diff emeasure_eq_measure sb measure_nonneg ennreal_minus
+           (auto simp: emeasure_Diff emeasure_eq_measure sb ennreal_minus
                        finite_measure_mono divide_ennreal ennreal_less_iff
                  intro: less_imp_le)
       also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
--- a/src/HOL/Analysis/Simplex_Content.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Simplex_Content.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -252,7 +252,7 @@
                     ((a + b + c) - 2 * c)"
     unfolding power2_eq_square by (simp add: s_def a_def b_def c_def algebra_simps)
   also have "\<dots> = 16 * s * (s - a) * (s - b) * (s - c)"
-    by (simp add: s_def field_split_simps mult_ac)
+    by (simp add: s_def field_split_simps)
   finally have "content (convex hull {A, B, C}) ^ 2 = s * (s - a) * (s - b) * (s - c)"
     by simp
   also have "\<dots> = sqrt (s * (s - a) * (s - b) * (s - c)) ^ 2"
--- a/src/HOL/Analysis/Starlike.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Starlike.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -5199,7 +5199,7 @@
         using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
         done
       also have "... = c a + (1 - c a)"
-        by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS')
+        by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS'(1))
       finally show "sum (cc(a := c a)) (insert a (T \<inter> T')) = 1"
         by simp
       have "(\<Sum>x\<in>insert a (T \<inter> T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\<Sum>x \<in> T \<inter> T'. (cc(a := c a)) x *\<^sub>R x)"
--- a/src/HOL/Analysis/Summation_Tests.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Summation_Tests.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -149,7 +149,7 @@
   also have "Discrete.log k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro Discrete.log_eqI) simp_all
   hence "(\<Sum>k = 2^n..<2^Suc n. f (2^Discrete.log k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^n))"
     by (intro sum.cong) simp_all
-  also have "\<dots> = 2^n * f (2^n)" by (simp add: of_nat_power)
+  also have "\<dots> = 2^n * f (2^n)" by (simp)
   finally show ?case by simp
 qed simp
 
@@ -163,7 +163,7 @@
   also have "Discrete.log k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro Discrete.log_eqI) simp_all
   hence "(\<Sum>k = 2^n..<2^Suc n. f (2*2^Discrete.log k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^Suc n))"
     by (intro sum.cong) simp_all
-  also have "\<dots> = 2^n * f (2^Suc n)" by (simp add: of_nat_power)
+  also have "\<dots> = 2^n * f (2^Suc n)" by (simp)
   finally show ?case by simp
 qed simp
 
--- a/src/HOL/Analysis/Tagged_Division.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Tagged_Division.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -2447,7 +2447,7 @@
               finally show "a \<bullet> i + real ?k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i" .
             next
               have "x \<bullet> i \<le> a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
-                using abi_less by (simp add: field_split_simps algebra_simps)
+                using abi_less by (simp add: field_split_simps)
               also have "... \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
                 apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
                 using aibi [OF \<open>i \<in> Basis\<close>] xab
@@ -2488,12 +2488,12 @@
       proof -
         have dd: "w / m \<le> v / n \<and> (v+1) / n \<le> (w+1) / m
                   \<Longrightarrow> inverse n \<le> inverse m" for w m v n::real
-          by (auto simp: field_split_simps algebra_simps)
+          by (auto simp: field_split_simps)
         have bjaj: "b \<bullet> j - a \<bullet> j > 0"
           using \<open>j \<in> Basis\<close> \<open>box a b \<noteq> {}\<close> box_eq_empty(1) by fastforce
         have "((g j) / 2 ^ m) * (b \<bullet> j - a \<bullet> j) \<le> ((f j) / 2 ^ n) * (b \<bullet> j - a \<bullet> j) \<and>
               (((f j) + 1) / 2 ^ n) * (b \<bullet> j - a \<bullet> j) \<le> (((g j) + 1) / 2 ^ m) * (b \<bullet> j - a \<bullet> j)"
-          using that \<open>j \<in> Basis\<close> by (simp add: subset_box algebra_simps field_split_simps aibi)
+          using that \<open>j \<in> Basis\<close> by (simp add: subset_box field_split_simps aibi)
         then have "((g j) / 2 ^ m) \<le> ((f j) / 2 ^ n) \<and>
           ((real(f j) + 1) / 2 ^ n) \<le> ((real(g j) + 1) / 2 ^ m)"
           by (metis bjaj mult.commute of_nat_1 of_nat_add real_mult_le_cancel_iff2)
--- a/src/HOL/Analysis/Weierstrass_Theorems.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Analysis/Weierstrass_Theorems.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -110,7 +110,7 @@
       then have "(\<Sum>k\<le>n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n"
         by (simp add: power2_eq_square)
       then show ?thesis
-        using n by (simp add: sum_divide_distrib field_split_simps mult.commute power2_commute)
+        using n by (simp add: sum_divide_distrib field_split_simps power2_commute)
     qed
     { fix k
       assume k: "k \<le> n"
@@ -158,7 +158,7 @@
     also have "... < e"
       apply (simp add: field_simps)
       using \<open>d>0\<close> nbig e \<open>n>0\<close>
-      apply (simp add: field_split_simps algebra_simps)
+      apply (simp add: field_split_simps)
       using ed0 by linarith
     finally have "\<bar>f x - (\<Sum>k\<le>n. f (real k / real n) * Bernstein n k x)\<bar> < e" .
   }
@@ -576,7 +576,7 @@
   define B where "B j = {x \<in> S. f x \<ge> (j + 1/3)*e}" for j :: nat
   have ngt: "(n-1) * e \<ge> normf f" "n\<ge>1"
     using e
-    apply (simp_all add: n_def field_simps of_nat_Suc)
+    apply (simp_all add: n_def field_simps)
     by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq)
   then have ge_fx: "(n-1) * e \<ge> f x" if "x \<in> S" for x
     using f normf_upper that by fastforce
@@ -646,7 +646,7 @@
       then have "t \<in> B i"
         using Anj e ge_fx [OF t] \<open>1 \<le> n\<close> fpos [OF t] t
         apply (simp add: A_def B_def)
-        apply (clarsimp simp add: field_simps of_nat_diff not_le of_nat_Suc)
+        apply (clarsimp simp add: field_simps of_nat_diff not_le)
         apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"])
         apply auto
         done
@@ -693,7 +693,7 @@
         apply simp
         apply (rule order_trans [OF _ sum_bounded_below [OF less_imp_le [OF ****]]])
         using True
-        apply (simp_all add: of_nat_Suc of_nat_diff)
+        apply (simp_all add: of_nat_diff)
         done
       also have "... \<le> g t"
         using jn e
@@ -744,7 +744,7 @@
       have "\<not> real (Suc n) < inverse e"
         using \<open>N \<le> n\<close> N using less_imp_inverse_less by force
       then have "1 / (1 + real n) \<le> e"
-        using e by (simp add: field_simps of_nat_Suc)
+        using e by (simp add: field_simps)
       then have "\<bar>f x - g x\<bar> < e"
         using n(2) x by auto
     } note * = this
@@ -927,7 +927,7 @@
   show ?case
     apply (rule_tac x="\<lambda>i. c" in exI)
     apply (rule_tac x=0 in exI)
-    apply (auto simp: mult_ac of_nat_Suc)
+    apply (auto simp: mult_ac)
     done
   case (add f1 f2)
   then obtain a1 n1 a2 n2 where
@@ -1328,7 +1328,7 @@
     also have "... = (\<Sum>j\<in>B. if j = i then x \<bullet> i else 0)"
       by (rule sum.cong; simp)
     also have "... = i \<bullet> x"
-      by (simp add: \<open>finite B\<close> that inner_commute sum.delta)
+      by (simp add: \<open>finite B\<close> that inner_commute)
     finally show ?thesis .
   qed
 qed
--- a/src/HOL/Binomial.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Binomial.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -121,21 +121,35 @@
   using binomial_Suc_Suc [of "n - 1" "k - 1"] by simp
 
 lemma Suc_times_binomial_eq: "Suc n * (n choose k) = (Suc n choose Suc k) * Suc k"
-  apply (induct n arbitrary: k)
-   apply simp
-   apply arith
-  apply (case_tac k)
-   apply (auto simp add: add_mult_distrib add_mult_distrib2 le_Suc_eq binomial_eq_0)
-  done
+proof (induction n arbitrary: k)
+  case 0
+  then show ?case
+    by auto
+next
+  case (Suc n)
+  show ?case 
+  proof (cases k)
+    case (Suc k')
+    then show ?thesis
+      using Suc.IH
+      by (auto simp add: add_mult_distrib add_mult_distrib2 le_Suc_eq binomial_eq_0)
+  qed auto
+qed
 
 lemma binomial_le_pow2: "n choose k \<le> 2^n"
-  apply (induct n arbitrary: k)
-   apply (case_tac k)
-    apply simp_all
-  apply (case_tac k)
-   apply auto
-  apply (simp add: add_le_mono mult_2)
-  done
+proof (induction n arbitrary: k)
+  case 0
+  then show ?case
+    using le_less less_le_trans by fastforce
+next
+  case (Suc n)
+  show ?case
+  proof (cases k)
+    case (Suc k')
+    then show ?thesis
+      using Suc.IH by (simp add: add_le_mono mult_2)
+  qed auto
+qed
 
 text \<open>The absorption property.\<close>
 lemma Suc_times_binomial: "Suc k * (Suc n choose Suc k) = Suc n * (n choose k)"
@@ -246,9 +260,7 @@
   assumes kn: "k \<le> n"
   shows "(of_nat (n choose k) :: 'a::field_char_0) = fact n / (fact k * fact (n - k))"
   using binomial_fact_lemma[OF kn]
-  apply (simp add: field_simps)
-  apply (metis mult.commute of_nat_fact of_nat_mult)
-  done
+  by (metis (mono_tags, lifting) fact_nonzero mult_eq_0_iff nonzero_mult_div_cancel_left of_nat_fact of_nat_mult)
 
 lemma fact_binomial:
   assumes "k \<le> n"
@@ -361,11 +373,11 @@
   for a :: "'a::field_char_0"
 proof (cases k)
   case (Suc k')
+  then have "a gchoose k = pochhammer (a - of_nat k') (Suc k') / ((1 + of_nat k') * fact k')"
+    by (simp add: gbinomial_prod_rev pochhammer_prod_rev atLeastLessThanSuc_atLeastAtMost
+        prod.atLeast_Suc_atMost_Suc_shift of_nat_diff flip: power_mult_distrib prod.cl_ivl_Suc)
   then show ?thesis
-    apply (simp add: pochhammer_minus)
-    apply (simp add: gbinomial_prod_rev pochhammer_prod_rev power_mult_distrib [symmetric] atLeastLessThanSuc_atLeastAtMost
-        prod.atLeast_Suc_atMost_Suc_shift of_nat_diff del: prod.cl_ivl_Suc)
-    done
+    by (simp add: pochhammer_minus Suc)
 qed auto
 
 lemma gbinomial_pochhammer': "a gchoose k = pochhammer (a - of_nat k + 1) k / fact k"
@@ -437,10 +449,8 @@
   (is "?l = ?r")
 proof -
   have "?r = ((- 1) ^k * pochhammer (- a) k / fact k) * (of_nat k - (- a + of_nat k))"
-    apply (simp only: gbinomial_pochhammer pochhammer_Suc right_diff_distrib power_Suc)
-    apply (simp del: of_nat_Suc fact_Suc)
-    apply (auto simp add: field_simps simp del: of_nat_Suc)
-    done
+    unfolding gbinomial_pochhammer pochhammer_Suc right_diff_distrib power_Suc
+    by (auto simp add: field_simps simp del: of_nat_Suc)
   also have "\<dots> = ?l"
     by (simp add: field_simps gbinomial_pochhammer)
   finally show ?thesis ..
@@ -459,17 +469,17 @@
 next
   case (Suc h)
   have eq0: "(\<Prod>i\<in>{1..k}. (a + 1) - of_nat i) = (\<Prod>i\<in>{0..h}. a - of_nat i)"
-    apply (rule prod.reindex_cong [where l = Suc])
-      using Suc
-      apply (auto simp add: image_Suc_atMost)
-    done
+  proof (rule prod.reindex_cong)
+    show "{1..k} = Suc ` {0..h}"
+      using Suc by (auto simp add: image_Suc_atMost)
+  qed auto
   have "fact (Suc k) * (a gchoose k + (a gchoose (Suc k))) =
       (a gchoose Suc h) * (fact (Suc (Suc h))) +
       (a gchoose Suc (Suc h)) * (fact (Suc (Suc h)))"
     by (simp add: Suc field_simps del: fact_Suc)
   also have "\<dots> =
     (a gchoose Suc h) * of_nat (Suc (Suc h) * fact (Suc h)) + (\<Prod>i=0..Suc h. a - of_nat i)"
-    apply (simp del: fact_Suc prod.op_ivl_Suc add: gbinomial_mult_fact field_simps mult.left_commute [of _ "2"])
+    apply (simp only: gbinomial_mult_fact field_simps mult.left_commute [of _ "2"])
     apply (simp del: fact_Suc add: fact_Suc [of "Suc h"] field_simps gbinomial_mult_fact
       mult.left_commute [of _ "2"] atLeastLessThanSuc_atLeastAtMost)
     done
@@ -608,10 +618,8 @@
   also have "(\<Sum>i\<le>n. of_nat (n choose Suc i) * pochhammer a (Suc i) * pochhammer b (n - i)) +
         pochhammer b (Suc n) =
       (\<Sum>i=0..Suc n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc n - i))"
-    apply (subst sum.atLeast_Suc_atMost)
-    apply simp
-    apply (subst sum.shift_bounds_cl_Suc_ivl)
-    apply (simp add: atLeast0AtMost)
+    apply (subst sum.atLeast_Suc_atMost, simp)
+    apply (simp add: sum.shift_bounds_cl_Suc_ivl atLeast0AtMost del: sum.cl_ivl_Suc)
     done
   also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc n - i))"
     using Suc by (intro sum.mono_neutral_right) (auto simp: not_le binomial_eq_0)
@@ -641,12 +649,11 @@
     using assms of_nat_0_le_iff order_trans by blast
   have "(a / of_nat k :: 'a) ^ k = (\<Prod>i = 0..<k. a / of_nat k :: 'a)"
     by simp
-  also have "\<dots> \<le> a gchoose k" (* FIXME *)
-    unfolding gbinomial_altdef_of_nat
-    apply (safe intro!: prod_mono)
-    apply simp_all
-    prefer 2
-    subgoal premises for i
+  also have "\<dots> \<le> a gchoose k"
+  proof -
+    have "\<And>i. i < k \<Longrightarrow> 0 \<le> a / of_nat k"
+      by (simp add: x zero_le_divide_iff)
+    moreover have "a / of_nat k \<le> (a - of_nat i) / of_nat (k - i)" if "i < k" for i
     proof -
       from assms have "a * of_nat i \<ge> of_nat (i * k)"
         by (metis mult.commute mult_le_cancel_right of_nat_less_0_iff of_nat_mult)
@@ -655,12 +662,14 @@
       then have "a * of_nat (k - i) \<le> (a - of_nat i) * of_nat k"
         using \<open>i < k\<close> by (simp add: algebra_simps zero_less_mult_iff of_nat_diff)
       then have "a * of_nat (k - i) \<le> (a - of_nat i) * (of_nat k :: 'a)"
-        by (simp only: of_nat_mult[symmetric] of_nat_le_iff)
+        by blast
       with assms show ?thesis
         using \<open>i < k\<close> by (simp add: field_simps)
     qed
-    apply (simp add: x zero_le_divide_iff)
-    done
+    ultimately show ?thesis
+      unfolding gbinomial_altdef_of_nat
+      by (intro prod_mono) auto
+  qed
   finally show ?thesis .
 qed
 
@@ -919,12 +928,16 @@
 
 lemma gbinomial_partial_sum_poly_xpos:
     "(\<Sum>k\<le>m. (of_nat m + a gchoose k) * x^k * y^(m-k)) =
-     (\<Sum>k\<le>m. (of_nat k + a - 1 gchoose k) * x^k * (x + y)^(m-k))"
-  apply (subst gbinomial_partial_sum_poly)
-  apply (subst gbinomial_negated_upper)
-  apply (intro sum.cong, rule refl)
-  apply (simp add: power_mult_distrib [symmetric])
-  done
+     (\<Sum>k\<le>m. (of_nat k + a - 1 gchoose k) * x^k * (x + y)^(m-k))" (is "?lhs = ?rhs")
+proof -
+  have "?lhs = (\<Sum>k\<le>m. (- a gchoose k) * (- x) ^ k * (x + y) ^ (m - k))"
+    by (simp add: gbinomial_partial_sum_poly)
+  also have "... = (\<Sum>k\<le>m. (-1) ^ k * (of_nat k - - a - 1 gchoose k) * (- x) ^ k * (x + y) ^ (m - k))"
+    by (metis (no_types, hide_lams) gbinomial_negated_upper)
+  also have "... = ?rhs"
+    by (intro sum.cong) (auto simp flip: power_mult_distrib)
+  finally show ?thesis .
+qed
 
 lemma binomial_r_part_sum: "(\<Sum>k\<le>m. (2 * m + 1 choose k)) = 2 ^ (2 * m)"
 proof -
@@ -1012,12 +1025,12 @@
   by (subst binomial_fact_lemma [symmetric]) auto
 
 lemma choose_dvd:
-  "k \<le> n \<Longrightarrow> fact k * fact (n - k) dvd (fact n :: 'a::linordered_semidom)"
+  assumes "k \<le> n" shows "fact k * fact (n - k) dvd (fact n :: 'a::linordered_semidom)"
   unfolding dvd_def
-  apply (rule exI [where x="of_nat (n choose k)"])
-  using binomial_fact_lemma [of k n, THEN arg_cong [where f = of_nat and 'b='a]]
-  apply auto
-  done
+proof
+  show "fact n = fact k * fact (n - k) * of_nat (n choose k)"
+    by (metis assms binomial_fact_lemma of_nat_fact of_nat_mult) 
+qed
 
 lemma fact_fact_dvd_fact:
   "fact k * fact n dvd (fact (k + n) :: 'a::linordered_semidom)"
@@ -1030,11 +1043,14 @@
   have "?lhs =
       fact (m + r + k) div (fact (m + k) * fact (m + r - m)) * (fact (m + k) div (fact k * fact m))"
     by (simp add: binomial_altdef_nat)
-  also have "\<dots> = fact (m + r + k) div (fact r * (fact k * fact m))"
+  also have "... = fact (m + r + k) * fact (m + k) div
+                 (fact (m + k) * fact (m + r - m) * (fact k * fact m))"
     apply (subst div_mult_div_if_dvd)
-    apply (auto simp: algebra_simps fact_fact_dvd_fact)
+      apply (auto simp: algebra_simps fact_fact_dvd_fact)
     apply (metis add.assoc add.commute fact_fact_dvd_fact)
     done
+  also have "\<dots> = fact (m + r + k) div (fact r * (fact k * fact m))"
+    by (auto simp: algebra_simps fact_fact_dvd_fact)
   also have "\<dots> = (fact (m + r + k) * fact (m + r)) div (fact r * (fact k * fact m) * fact (m + r))"
     apply (subst div_mult_div_if_dvd [symmetric])
     apply (auto simp add: algebra_simps)
@@ -1042,9 +1058,7 @@
     done
   also have "\<dots> =
       (fact (m + r + k) div (fact k * fact (m + r)) * (fact (m + r) div (fact r * fact m)))"
-    apply (subst div_mult_div_if_dvd)
-    apply (auto simp: fact_fact_dvd_fact algebra_simps)
-    done
+    by (auto simp: div_mult_div_if_dvd fact_fact_dvd_fact algebra_simps)
   finally show ?thesis
     by (simp add: binomial_altdef_nat mult.commute)
 qed
@@ -1163,17 +1177,13 @@
   have 2: "xs \<noteq> [] \<Longrightarrow> sum_list(tl xs) = sum_list xs - hd xs" for xs :: "nat list"
     by (auto simp add: neq_Nil_conv)
   have f: "bij_betw ?f ?A ?A'"
-    apply (rule bij_betw_byWitness[where f' = tl])
-    using assms
-    apply (auto simp: 2 length_0_conv[symmetric] 1 simp del: length_0_conv)
-    done
+    by (rule bij_betw_byWitness[where f' = tl]) (use assms in \<open>auto simp: 2 1 simp flip: length_0_conv\<close>)
   have 3: "xs \<noteq> [] \<Longrightarrow> hd xs + (sum_list xs - hd xs) = sum_list xs" for xs :: "nat list"
     by (metis 1 sum_list_simps(2) 2)
   have g: "bij_betw ?g ?B ?B'"
     apply (rule bij_betw_byWitness[where f' = "\<lambda>l. (hd l - 1) # tl l"])
     using assms
-    by (auto simp: 2 length_0_conv[symmetric] intro!: 3
-        simp del: length_greater_0_conv length_0_conv)
+    by (auto simp: 2 simp flip: length_0_conv intro!: 3)
   have fin: "finite {xs. size xs = M \<and> set xs \<subseteq> {0..<N}}" for M N :: nat
     using finite_lists_length_eq[OF finite_atLeastLessThan] conj_commute by auto
   have fin_A: "finite ?A" using fin[of _ "N+1"]
--- a/src/HOL/Computational_Algebra/Polynomial.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Computational_Algebra/Polynomial.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -4367,11 +4367,11 @@
     define m where "m = (GREATEST m. \<not>c dvd coeff b m)"
     assume "\<not>[:c:] dvd b"
     hence A: "\<exists>i. \<not>c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast
-    have B: "\<forall>i. \<not>c dvd coeff b i \<longrightarrow> i \<le> degree b"
+    have B: "\<And>i. \<not>c dvd coeff b i \<Longrightarrow> i \<le> degree b"
       by (auto intro: le_degree)
     have coeff_m: "\<not>c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B])
     have "i \<le> m" if "\<not>c dvd coeff b i" for i
-      unfolding m_def by (rule Greatest_le_nat[OF that B])
+      unfolding m_def by (metis (mono_tags, lifting) B Greatest_le_nat that)
     hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force
 
     have "c dvd coeff a i" for i
--- a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -286,11 +286,11 @@
     define m where "m = (GREATEST m. \<not>c dvd coeff b m)"
     assume "\<not>[:c:] dvd b"
     hence A: "\<exists>i. \<not>c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast
-    have B: "\<forall>i. \<not>c dvd coeff b i \<longrightarrow> i \<le> degree b"
+    have B: "\<And>i. \<not>c dvd coeff b i \<Longrightarrow> i \<le> degree b"
       by (auto intro: le_degree)
     have coeff_m: "\<not>c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B])
     have "i \<le> m" if "\<not>c dvd coeff b i" for i
-      unfolding m_def by (rule Greatest_le_nat[OF that B])
+      unfolding m_def by (blast intro!: Greatest_le_nat that B)
     hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force
 
     have "c dvd coeff a i" for i
--- a/src/HOL/Data_Structures/AVL_Map.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Data_Structures/AVL_Map.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -18,7 +18,8 @@
 fun delete :: "'a::linorder \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
 "delete _ Leaf = Leaf" |
 "delete x (Node l ((a,b), h) r) = (case cmp x a of
-   EQ \<Rightarrow> del_root (Node l ((a,b), h) r) |
+   EQ \<Rightarrow> if l = Leaf then r
+         else let (l', ab') = split_max l in balR l' ab' r |
    LT \<Rightarrow> balR (delete x l) (a,b) r |
    GT \<Rightarrow> balL l (a,b) (delete x r))"
 
@@ -34,7 +35,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_split_maxD split: prod.splits)
+      inorder_split_maxD split: prod.splits)
 
 
 subsection \<open>AVL invariants\<close>
@@ -120,7 +121,8 @@
   case 1
   show ?case
   proof(cases "x = a")
-    case True with Node 1 show ?thesis by (auto simp:avl_del_root)
+    case True with Node 1 show ?thesis
+      using avl_split_max[of l] by (auto simp: avl_balR split: prod.split)
   next
     case False
     show ?thesis 
@@ -133,11 +135,8 @@
   case 2
   show ?case
   proof(cases "x = a")
-    case True
-    with 1 have "height (Node l (n, h) r) = height(del_root (Node l (n, h) r))
-      \<or> height (Node l (n, h) r) = height(del_root (Node l (n, h) r)) + 1"
-      by (subst height_del_root,simp_all)
-    with True show ?thesis by simp
+    case True then show ?thesis using 1 avl_split_max[of l]
+      by(auto simp: balR_def max_absorb2 split!: if_splits prod.split tree.split)
   next
     case False
     show ?thesis 
--- a/src/HOL/Data_Structures/AVL_Set.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Data_Structures/AVL_Set.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1,6 +1,6 @@
 (*
 Author:     Tobias Nipkow, Daniel Stüwe
-Largely derived from AFP entry AVL.
+Based on the AFP entry AVL.
 *)
 
 section "AVL Tree Implementation of Sets"
@@ -33,26 +33,26 @@
 "node l a r = Node l (a, max (ht l) (ht r) + 1) r"
 
 definition balL :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
-"balL l a r =
-  (if ht l = ht r + 2 then
-     case l of 
-       Node bl (b, _) br \<Rightarrow>
-         if ht bl < ht br then
-           case br of
-             Node cl (c, _) cr \<Rightarrow> node (node bl b cl) c (node cr a r)
-         else node bl b (node br a r)
-   else node l a r)"
+"balL AB b C =
+  (if ht AB = ht C + 2 then
+     case AB of 
+       Node A (a, _) B \<Rightarrow>
+         if ht A \<ge> ht B then node A a (node B b C)
+         else
+           case B of
+             Node B\<^sub>1 (ab, _) B\<^sub>2 \<Rightarrow> node (node A a B\<^sub>1) ab (node B\<^sub>2 b C)
+   else node AB b C)"
 
 definition balR :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
-"balR l a r =
-   (if ht r = ht l + 2 then
-      case r of
-        Node bl (b, _) br \<Rightarrow>
-          if ht bl > ht br then
-            case bl of
-              Node cl (c, _) cr \<Rightarrow> node (node l a cl) c (node cr b br)
-          else node (node l a bl) b br
-  else node l a r)"
+"balR A a BC =
+   (if ht BC = ht A + 2 then
+      case BC of
+        Node B (b, _) C \<Rightarrow>
+          if ht B \<le> ht C then node (node A a B) b C
+          else
+            case B of
+              Node B\<^sub>1 (ab, _) B\<^sub>2 \<Rightarrow> node (node A a B\<^sub>1) ab (node B\<^sub>2 b C)
+  else node A a BC)"
 
 fun insert :: "'a::linorder \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 "insert x Leaf = Node Leaf (x, 1) Leaf" |
@@ -67,18 +67,12 @@
 
 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 Leaf (a,_) r) = r" |
-"del_root (Node l (a,_) Leaf) = l" |
-"del_root (Node l (a,_) r) = (let (l', a') = split_max l in balR l' a' r)"
-
-lemmas del_root_cases = del_root.cases[split_format(complete), case_names Leaf_t Node_Leaf Node_Node]
-
 fun delete :: "'a::linorder \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
 "delete _ Leaf = Leaf" |
 "delete x (Node l (a, n) r) =
   (case cmp x a of
-     EQ \<Rightarrow> del_root (Node l (a, n) r) |
+     EQ \<Rightarrow> if l = Leaf then r
+           else let (l', a') = split_max l in balR l' a' r |
      LT \<Rightarrow> balR (delete x l) a r |
      GT \<Rightarrow> balL l a (delete x r))"
 
@@ -112,16 +106,10 @@
 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 l ah r)) = inorder l @ inorder r"
-by(cases "Node l ah r" rule: del_root.cases)
-  (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_split_maxD split: prod.splits)
+  (auto simp: del_list_simps inorder_balL inorder_balR inorder_split_maxD split: prod.splits)
 
 
 subsection \<open>AVL invariants\<close>
@@ -136,14 +124,16 @@
 lemma ht_height[simp]: "avl t \<Longrightarrow> ht t = height t"
 by (cases t rule: tree2_cases) simp_all
 
+text \<open>First, a fast but relatively manual proof with many lemmas:\<close>
+
 lemma height_balL:
-  "\<lbrakk> height l = height r + 2; avl l; avl r \<rbrakk> \<Longrightarrow>
+  "\<lbrakk> avl l; avl r; height l = height r + 2 \<rbrakk> \<Longrightarrow>
    height (balL l a r) = height r + 2 \<or>
    height (balL l a r) = height r + 3"
 by (cases l) (auto simp:node_def balL_def split:tree.split)
        
 lemma height_balR:
-  "\<lbrakk> height r = height l + 2; avl l; avl r \<rbrakk> \<Longrightarrow>
+  "\<lbrakk> avl l; avl r; height r = height l + 2 \<rbrakk> \<Longrightarrow>
    height (balR l a r) = height l + 2 \<or>
    height (balR l a r) = height l + 3"
 by (cases r) (auto simp add:node_def balR_def split:tree.split)
@@ -151,65 +141,27 @@
 lemma height_node[simp]: "height(node l a r) = max (height l) (height r) + 1"
 by (simp add: node_def)
 
-lemma avl_node:
-  "\<lbrakk> avl l; avl r;
-     height l = height r \<or> height l = height r + 1 \<or> height r = height l + 1
-   \<rbrakk> \<Longrightarrow> avl(node l a r)"
-by (auto simp add:max_def node_def)
-
 lemma height_balL2:
   "\<lbrakk> avl l; avl r; height l \<noteq> height r + 2 \<rbrakk> \<Longrightarrow>
-   height (balL l a r) = (1 + max (height l) (height r))"
+   height (balL l a r) = 1 + max (height l) (height r)"
 by (cases l, cases r) (simp_all add: balL_def)
 
 lemma height_balR2:
   "\<lbrakk> avl l;  avl r;  height r \<noteq> height l + 2 \<rbrakk> \<Longrightarrow>
-   height (balR l a r) = (1 + max (height l) (height r))"
+   height (balR l a r) = 1 + max (height l) (height r)"
 by (cases l, cases r) (simp_all add: balR_def)
 
 lemma avl_balL: 
-  assumes "avl l" "avl r" and "height l = height r \<or> height l = height r + 1
-    \<or> height r = height l + 1 \<or> height l = height r + 2" 
-  shows "avl(balL l a r)"
-proof(cases l rule: tree2_cases)
-  case Leaf
-  with assms show ?thesis by (simp add: node_def balL_def)
-next
-  case Node
-  with assms show ?thesis
-  proof(cases "height l = height r + 2")
-    case True
-    from True Node assms show ?thesis
-      by (auto simp: balL_def intro!: avl_node split: tree.split) arith+
-  next
-    case False
-    with assms show ?thesis by (simp add: avl_node balL_def)
-  qed
-qed
+  "\<lbrakk> avl l; avl r; height r - 1 \<le> height l \<and> height l \<le> height r + 2 \<rbrakk> \<Longrightarrow> avl(balL l a r)"
+by(cases l, cases r)
+  (auto simp: balL_def node_def split!: if_split tree.split)
 
 lemma avl_balR: 
-  assumes "avl l" and "avl r" and "height l = height r \<or> height l = height r + 1
-    \<or> height r = height l + 1 \<or> height r = height l + 2" 
-  shows "avl(balR l a r)"
-proof(cases r rule: tree2_cases)
-  case Leaf
-  with assms show ?thesis by (simp add: node_def balR_def)
-next
-  case Node
-  with assms show ?thesis
-  proof(cases "height r = height l + 2")
-    case True
-      from True Node assms show ?thesis
-        by (auto simp: balR_def intro!: avl_node split: tree.split) arith+
-  next
-    case False
-    with assms show ?thesis by (simp add: balR_def avl_node)
-  qed
-qed
+  "\<lbrakk> avl l; avl r; height l - 1 \<le> height r \<and> height r \<le> height l + 2 \<rbrakk> \<Longrightarrow> avl(balR l a r)"
+by(cases r, cases l)
+  (auto simp: balR_def node_def split!: if_split tree.split)
 
-(* It appears that these two properties need to be proved simultaneously: *)
-
-text\<open>Insertion maintains the AVL property:\<close>
+text\<open>Insertion maintains the AVL property. Requires simultaneous proof.\<close>
 
 theorem avl_insert:
   assumes "avl t"
@@ -276,79 +228,24 @@
   qed
 qed simp_all
 
+text \<open>Now an automatic proof without lemmas:\<close>
+
+theorem avl_insert_auto: "avl t \<Longrightarrow>
+  avl(insert x t) \<and> height (insert x t) \<in> {height t, height t + 1}"
+apply (induction t rule: tree2_induct)
+ (* if you want to save a few secs: apply (auto split!: if_split) *)
+ apply (auto simp: balL_def balR_def node_def max_absorb2 split!: if_split tree.split)
+done
+
 
 subsubsection \<open>Deletion maintains AVL balance\<close>
 
 lemma avl_split_max:
-  assumes "avl x" and "x \<noteq> Leaf"
-  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: split_max_induct)
-  case Node
-  case 1
-  thus ?case using Node
-    by (auto simp: height_balL height_balL2 avl_balL split:prod.split)
-next
-  case (Node l a _ r)
-  case 2
-  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+
-qed auto
-
-lemma avl_del_root:
-  assumes "avl t" and "t \<noteq> Leaf"
-  shows "avl(del_root t)" 
-using assms
-proof (cases t rule:del_root_cases)
-  case (Node_Node ll ln lh lr n h rl rn rh rr)
-  let ?l = "Node ll (ln, lh) lr"
-  let ?r = "Node rl (rn, rh) rr"
-  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_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(split_max ?l)) ?r)"
-    by (rule avl_balR)
-  with Node_Node show ?thesis by (auto split:prod.splits)
-qed simp_all
-
-lemma height_del_root:
-  assumes "avl t" and "t \<noteq> Leaf" 
-  shows "height t = height(del_root t) \<or> height t = height(del_root t) + 1"
-using assms
-proof (cases t rule: del_root_cases)
-  case (Node_Node ll lx lh lr x h rl rx rh rr)
-  let ?l = "Node ll (lx, lh) lr"
-  let ?r = "Node rl (rx, rh) rr"
-  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_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")
-    case False
-    show ?thesis using l'_height t_height False
-      by (subst height_balR2[OF \<open>avl ?l'\<close> \<open>avl ?r\<close> False])+ arith
-  next
-    case True
-    show ?thesis
-    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
-      case 2 thus ?thesis using l'_height t_height True by arith
-    qed
-  qed
-  thus ?thesis using Node_Node by (auto split:prod.splits)
-qed simp_all
+  "\<lbrakk> avl t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
+  avl (fst (split_max t)) \<and>
+  height t \<in> {height(fst (split_max t)), height(fst (split_max t)) + 1}"
+by(induct t rule: split_max_induct)
+  (auto simp: balL_def node_def max_absorb2 split!: prod.split if_split tree.split)
 
 text\<open>Deletion maintains the AVL property:\<close>
 
@@ -359,26 +256,13 @@
 proof (induct t rule: tree2_induct)
   case (Node l a n r)
   case 1
-  show ?case
-  proof(cases "x = a")
-    case True with Node 1 show ?thesis by (auto simp:avl_del_root)
-  next
-    case False
-    show ?thesis 
-    proof(cases "x<a")
-      case True with Node 1 show ?thesis by (auto simp add:avl_balR)
-    next
-      case False with Node 1 \<open>x\<noteq>a\<close> show ?thesis by (auto simp add:avl_balL)
-    qed
-  qed
+  thus ?case
+    using Node avl_split_max[of l] by (auto simp: avl_balL avl_balR split: prod.split)
   case 2
   show ?case
   proof(cases "x = a")
-    case True
-    with 1 have "height (Node l (a,n) r) = height(del_root (Node l (a,n) r))
-      \<or> height (Node l (a,n) r) = height(del_root (Node l (a,n) r)) + 1"
-      by (subst height_del_root,simp_all)
-    with True show ?thesis by simp
+    case True then show ?thesis using 1 avl_split_max[of l]
+      by(auto simp: balR_def max_absorb2 split!: if_splits prod.split tree.split)
   next
     case False
     show ?thesis 
@@ -420,6 +304,37 @@
   qed
 qed simp_all
 
+text \<open>A more automatic proof.
+Complete automation as for insertion seems hard due to resource requirements.\<close>
+
+theorem avl_delete_auto:
+  assumes "avl t" 
+  shows "avl(delete x t)" and "height t \<in> {height (delete x t), height (delete x t) + 1}"
+using assms
+proof (induct t rule: tree2_induct)
+  case (Node l a n r)
+  case 1
+  show ?case
+  proof(cases "x = a")
+    case True thus ?thesis
+      using 1 avl_split_max[of l] by (auto simp: avl_balR split: prod.split)
+  next
+    case False thus ?thesis
+      using Node 1 by (auto simp: avl_balL avl_balR)
+  qed
+  case 2
+  show ?case
+  proof(cases "x = a")
+    case True thus ?thesis
+      using 1 avl_split_max[of l]
+      by(auto simp: balR_def max_absorb2 split!: if_splits prod.split tree.split)
+  next
+    case False thus ?thesis 
+      using height_balL[of l "delete x r" a] height_balR[of "delete x l" r a] Node 1
+        by(auto simp: balL_def balR_def split!: if_split)
+  qed
+qed simp_all
+
 
 subsection "Overall correctness"
 
--- a/src/HOL/Equiv_Relations.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Equiv_Relations.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -34,21 +34,21 @@
   unfolding refl_on_def by blast
 
 lemma equiv_comp_eq: "equiv A r \<Longrightarrow> r\<inverse> O r = r"
-  apply (unfold equiv_def)
-  apply clarify
-  apply (rule equalityI)
-   apply (iprover intro: sym_trans_comp_subset refl_on_comp_subset)+
-  done
+  unfolding equiv_def
+  by (iprover intro: sym_trans_comp_subset refl_on_comp_subset equalityI)
 
 text \<open>Second half.\<close>
 
-lemma comp_equivI: "r\<inverse> O r = r \<Longrightarrow> Domain r = A \<Longrightarrow> equiv A r"
-  apply (unfold equiv_def refl_on_def sym_def trans_def)
-  apply (erule equalityE)
-  apply (subgoal_tac "\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r")
-   apply fast
-  apply fast
-  done
+lemma comp_equivI:
+  assumes "r\<inverse> O r = r" "Domain r = A"
+  shows "equiv A r"
+proof -
+  have *: "\<And>x y. (x, y) \<in> r \<Longrightarrow> (y, x) \<in> r"
+    using assms by blast
+  show ?thesis
+    unfolding equiv_def refl_on_def sym_def trans_def
+    using assms by (auto intro: *)
+qed
 
 
 subsection \<open>Equivalence classes\<close>
@@ -58,10 +58,7 @@
   unfolding equiv_def trans_def sym_def by blast
 
 theorem equiv_class_eq: "equiv A r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> r``{a} = r``{b}"
-  apply (assumption | rule equalityI equiv_class_subset)+
-  apply (unfold equiv_def sym_def)
-  apply blast
-  done
+  by (intro equalityI equiv_class_subset; force simp add: equiv_def sym_def)
 
 lemma equiv_class_self: "equiv A r \<Longrightarrow> a \<in> A \<Longrightarrow> a \<in> r``{a}"
   unfolding equiv_def refl_on_def by blast
@@ -91,7 +88,7 @@
 definition quotient :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set set"  (infixl "'/'/" 90)
   where "A//r = (\<Union>x \<in> A. {r``{x}})"  \<comment> \<open>set of equiv classes\<close>
 
-lemma quotientI: "x \<in> A ==> r``{x} \<in> A//r"
+lemma quotientI: "x \<in> A \<Longrightarrow> r``{x} \<in> A//r"
   unfolding quotient_def by blast
 
 lemma quotientE: "X \<in> A//r \<Longrightarrow> (\<And>x. X = r``{x} \<Longrightarrow> x \<in> A \<Longrightarrow> P) \<Longrightarrow> P"
@@ -101,32 +98,31 @@
   unfolding equiv_def refl_on_def quotient_def by blast
 
 lemma quotient_disj: "equiv A r \<Longrightarrow> X \<in> A//r \<Longrightarrow> Y \<in> A//r \<Longrightarrow> X = Y \<or> X \<inter> Y = {}"
-  apply (unfold quotient_def)
-  apply clarify
-  apply (rule equiv_class_eq)
-   apply assumption
-  apply (unfold equiv_def trans_def sym_def)
-  apply blast
-  done
+  unfolding quotient_def equiv_def trans_def sym_def by blast
 
 lemma quotient_eqI:
-  "equiv A r \<Longrightarrow> X \<in> A//r \<Longrightarrow> Y \<in> A//r \<Longrightarrow> x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> X = Y"
-  apply (clarify elim!: quotientE)
-  apply (rule equiv_class_eq)
-   apply assumption
-  apply (unfold equiv_def sym_def trans_def)
-  apply blast
-  done
+  assumes "equiv A r" "X \<in> A//r" "Y \<in> A//r" and xy: "x \<in> X" "y \<in> Y" "(x, y) \<in> r"
+  shows "X = Y"
+proof -
+  obtain a b where "a \<in> A" and a: "X = r `` {a}" and "b \<in> A" and b: "Y = r `` {b}"
+    using assms by (auto elim!: quotientE)
+  then have "(a,b) \<in> r"
+      using xy \<open>equiv A r\<close> unfolding equiv_def sym_def trans_def by blast
+  then show ?thesis
+    unfolding a b by (rule equiv_class_eq [OF \<open>equiv A r\<close>])
+qed
 
 lemma quotient_eq_iff:
-  "equiv A r \<Longrightarrow> X \<in> A//r \<Longrightarrow> Y \<in> A//r \<Longrightarrow> x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> X = Y \<longleftrightarrow> (x, y) \<in> r"
-  apply (rule iffI)
-   prefer 2
-   apply (blast del: equalityI intro: quotient_eqI)
-  apply (clarify elim!: quotientE)
-  apply (unfold equiv_def sym_def trans_def)
-  apply blast
-  done
+  assumes "equiv A r" "X \<in> A//r" "Y \<in> A//r" and xy: "x \<in> X" "y \<in> Y" 
+  shows "X = Y \<longleftrightarrow> (x, y) \<in> r"
+proof
+  assume L: "X = Y" 
+  with assms show "(x, y) \<in> r" 
+    unfolding equiv_def sym_def trans_def by (blast elim!: quotientE)
+next
+  assume \<section>: "(x, y) \<in> r" show "X = Y"
+    by (rule quotient_eqI) (use \<section> assms in \<open>blast+\<close>)
+qed
 
 lemma eq_equiv_class_iff2: "equiv A r \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> {x}//r = {y}//r \<longleftrightarrow> (x, y) \<in> r"
   by (simp add: quotient_def eq_equiv_class_iff)
@@ -189,22 +185,22 @@
   \<comment> \<open>lemma required to prove \<open>UN_equiv_class\<close>\<close>
   by auto
 
-lemma UN_equiv_class: "equiv A r \<Longrightarrow> f respects r \<Longrightarrow> a \<in> A \<Longrightarrow> (\<Union>x \<in> r``{a}. f x) = f a"
+lemma UN_equiv_class:
+  assumes "equiv A r" "f respects r" "a \<in> A"
+  shows "(\<Union>x \<in> r``{a}. f x) = f a"
   \<comment> \<open>Conversion rule\<close>
-  apply (rule equiv_class_self [THEN UN_constant_eq])
-    apply assumption
-   apply assumption
-  apply (unfold equiv_def congruent_def sym_def)
-  apply (blast del: equalityI)
-  done
+proof -
+  have \<section>: "\<forall>x\<in>r `` {a}. f x = f a"
+    using assms unfolding equiv_def congruent_def sym_def by blast
+  show ?thesis
+    by (iprover intro: assms UN_constant_eq [OF equiv_class_self \<section>])
+qed
 
 lemma UN_equiv_class_type:
-  "equiv A r \<Longrightarrow> f respects r \<Longrightarrow> X \<in> A//r \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> (\<Union>x \<in> X. f x) \<in> B"
-  apply (unfold quotient_def)
-  apply clarify
-  apply (subst UN_equiv_class)
-     apply auto
-  done
+  assumes r: "equiv A r" "f respects r" and X: "X \<in> A//r" and AB: "\<And>x. x \<in> A \<Longrightarrow> f x \<in> B"
+  shows "(\<Union>x \<in> X. f x) \<in> B"
+  using assms unfolding quotient_def
+  by (auto simp: UN_equiv_class [OF r])
 
 text \<open>
   Sufficient conditions for injectiveness.  Could weaken premises!
@@ -213,19 +209,23 @@
 \<close>
 
 lemma UN_equiv_class_inject:
-  "equiv A r \<Longrightarrow> f respects r \<Longrightarrow>
-    (\<Union>x \<in> X. f x) = (\<Union>y \<in> Y. f y) \<Longrightarrow> X \<in> A//r ==> Y \<in> A//r
-    \<Longrightarrow> (\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<Longrightarrow> (x, y) \<in> r)
-    \<Longrightarrow> X = Y"
-  apply (unfold quotient_def)
-  apply clarify
-  apply (rule equiv_class_eq)
-   apply assumption
-  apply (subgoal_tac "f x = f xa")
-   apply blast
-  apply (erule box_equals)
-   apply (assumption | rule UN_equiv_class)+
-  done
+  assumes "equiv A r" "f respects r"
+    and eq: "(\<Union>x \<in> X. f x) = (\<Union>y \<in> Y. f y)" 
+    and X: "X \<in> A//r" and Y: "Y \<in> A//r" 
+    and fr: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<Longrightarrow> (x, y) \<in> r"
+  shows "X = Y"
+proof -
+  obtain a b where "a \<in> A" and a: "X = r `` {a}" and "b \<in> A" and b: "Y = r `` {b}"
+    using assms by (auto elim!: quotientE)
+  then have "\<Union> (f ` r `` {a}) = f a" "\<Union> (f ` r `` {b}) = f b"
+    by (iprover intro: UN_equiv_class [OF \<open>equiv A r\<close>] assms)+
+  then have "f a = f b"
+    using eq unfolding a b by (iprover intro: trans sym)
+  then have "(a,b) \<in> r"
+    using fr \<open>a \<in> A\<close> \<open>b \<in> A\<close> by blast
+  then show ?thesis
+    unfolding a b by (rule equiv_class_eq [OF \<open>equiv A r\<close>])
+qed
 
 
 subsection \<open>Defining binary operations upon equivalence classes\<close>
@@ -253,15 +253,20 @@
   unfolding congruent_def congruent2_def equiv_def refl_on_def by blast
 
 lemma congruent2_implies_congruent_UN:
-  "equiv A1 r1 \<Longrightarrow> equiv A2 r2 \<Longrightarrow> congruent2 r1 r2 f \<Longrightarrow> a \<in> A2 \<Longrightarrow>
-    congruent r1 (\<lambda>x1. \<Union>x2 \<in> r2``{a}. f x1 x2)"
-  apply (unfold congruent_def)
-  apply clarify
-  apply (rule equiv_type [THEN subsetD, THEN SigmaE2], assumption+)
-  apply (simp add: UN_equiv_class congruent2_implies_congruent)
-  apply (unfold congruent2_def equiv_def refl_on_def)
-  apply (blast del: equalityI)
-  done
+  assumes "equiv A1 r1" "equiv A2 r2" "congruent2 r1 r2 f" "a \<in> A2" 
+  shows "congruent r1 (\<lambda>x1. \<Union>x2 \<in> r2``{a}. f x1 x2)"
+  unfolding congruent_def
+proof clarify
+  fix c d
+  assume cd: "(c,d) \<in> r1"
+  then have "c \<in> A1" "d \<in> A1"
+    using \<open>equiv A1 r1\<close> by (auto elim!: equiv_type [THEN subsetD, THEN SigmaE2])
+  with assms show "\<Union> (f c ` r2 `` {a}) = \<Union> (f d ` r2 `` {a})"
+  proof (simp add: UN_equiv_class congruent2_implies_congruent)
+    show "f c a = f d a"
+      using assms cd unfolding congruent2_def equiv_def refl_on_def by blast
+  qed
+qed
 
 lemma UN_equiv_class2:
   "equiv A1 r1 \<Longrightarrow> equiv A2 r2 \<Longrightarrow> congruent2 r1 r2 f \<Longrightarrow> a1 \<in> A1 \<Longrightarrow> a2 \<in> A2 \<Longrightarrow>
@@ -273,11 +278,10 @@
     \<Longrightarrow> X1 \<in> A1//r1 \<Longrightarrow> X2 \<in> A2//r2
     \<Longrightarrow> (\<And>x1 x2. x1 \<in> A1 \<Longrightarrow> x2 \<in> A2 \<Longrightarrow> f x1 x2 \<in> B)
     \<Longrightarrow> (\<Union>x1 \<in> X1. \<Union>x2 \<in> X2. f x1 x2) \<in> B"
-  apply (unfold quotient_def)
-  apply clarify
-  apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN
-      congruent2_implies_congruent quotientI)
-  done
+  unfolding quotient_def
+  by (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN
+                   congruent2_implies_congruent quotientI)
+
 
 lemma UN_UN_split_split_eq:
   "(\<Union>(x1, x2) \<in> X. \<Union>(y1, y2) \<in> Y. A x1 x2 y1 y2) =
@@ -293,60 +297,63 @@
     \<Longrightarrow> congruent2 r1 r2 f"
   \<comment> \<open>Suggested by John Harrison -- the two subproofs may be\<close>
   \<comment> \<open>\<^emph>\<open>much\<close> simpler than the direct proof.\<close>
-  apply (unfold congruent2_def equiv_def refl_on_def)
-  apply clarify
-  apply (blast intro: trans)
-  done
+  unfolding congruent2_def equiv_def refl_on_def
+  by (blast intro: trans)
 
 lemma congruent2_commuteI:
   assumes equivA: "equiv A r"
     and commute: "\<And>y z. y \<in> A \<Longrightarrow> z \<in> A \<Longrightarrow> f y z = f z y"
     and congt: "\<And>y z w. w \<in> A \<Longrightarrow> (y,z) \<in> r \<Longrightarrow> f w y = f w z"
   shows "f respects2 r"
-  apply (rule congruent2I [OF equivA equivA])
-   apply (rule commute [THEN trans])
-     apply (rule_tac [3] commute [THEN trans, symmetric])
-       apply (rule_tac [5] sym)
-       apply (rule congt | assumption |
-         erule equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2])+
-  done
+proof (rule congruent2I [OF equivA equivA])
+  note eqv = equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2]
+  show "\<And>y z w. \<lbrakk>w \<in> A; (y, z) \<in> r\<rbrakk> \<Longrightarrow> f y w = f z w"
+    by (iprover intro: commute [THEN trans] sym congt elim: eqv)
+  show "\<And>y z w. \<lbrakk>w \<in> A; (y, z) \<in> r\<rbrakk> \<Longrightarrow> f w y = f w z"
+    by (iprover intro: congt elim: eqv)
+qed
 
 
 subsection \<open>Quotients and finiteness\<close>
 
 text \<open>Suggested by Florian Kammüller\<close>
 
-lemma finite_quotient: "finite A \<Longrightarrow> r \<subseteq> A \<times> A \<Longrightarrow> finite (A//r)"
-  \<comment> \<open>recall @{thm equiv_type}\<close>
-  apply (rule finite_subset)
-   apply (erule_tac [2] finite_Pow_iff [THEN iffD2])
-  apply (unfold quotient_def)
-  apply blast
-  done
+lemma finite_quotient:
+  assumes "finite A" "r \<subseteq> A \<times> A"
+  shows "finite (A//r)"
+    \<comment> \<open>recall @{thm equiv_type}\<close>
+proof -
+  have "A//r \<subseteq> Pow A"
+    using assms unfolding quotient_def by blast
+  moreover have "finite (Pow A)"
+    using assms by simp
+  ultimately show ?thesis
+    by (iprover intro: finite_subset)
+qed
 
 lemma finite_equiv_class: "finite A \<Longrightarrow> r \<subseteq> A \<times> A \<Longrightarrow> X \<in> A//r \<Longrightarrow> finite X"
-  apply (unfold quotient_def)
-  apply (rule finite_subset)
-   prefer 2 apply assumption
-  apply blast
-  done
+  unfolding quotient_def
+  by (erule rev_finite_subset) blast
 
-lemma equiv_imp_dvd_card: "finite A \<Longrightarrow> equiv A r \<Longrightarrow> \<forall>X \<in> A//r. k dvd card X \<Longrightarrow> k dvd card A"
-  apply (rule Union_quotient [THEN subst [where P="\<lambda>A. k dvd card A"]])
-   apply assumption
-  apply (rule dvd_partition)
-    prefer 3 apply (blast dest: quotient_disj)
-   apply (simp_all add: Union_quotient equiv_type)
-  done
+lemma equiv_imp_dvd_card:
+  assumes "finite A" "equiv A r" "\<And>X. X \<in> A//r \<Longrightarrow> k dvd card X"
+  shows "k dvd card A"
+proof (rule Union_quotient [THEN subst])
+  show "k dvd card (\<Union> (A // r))"
+    apply (rule dvd_partition)
+    using assms
+    by (auto simp: Union_quotient dest: quotient_disj)
+qed (use assms in blast)
 
-lemma card_quotient_disjoint: "finite A \<Longrightarrow> inj_on (\<lambda>x. {x} // r) A \<Longrightarrow> card (A//r) = card A"
-  apply (simp add:quotient_def)
-  apply (subst card_UN_disjoint)
-     apply assumption
-    apply simp
-   apply (fastforce simp add:inj_on_def)
-  apply simp
-  done
+lemma card_quotient_disjoint:
+  assumes "finite A" "inj_on (\<lambda>x. {x} // r) A"
+  shows "card (A//r) = card A"
+proof -
+  have "\<forall>i\<in>A. \<forall>j\<in>A. i \<noteq> j \<longrightarrow> r `` {j} \<noteq> r `` {i}"
+    using assms by (fastforce simp add: quotient_def inj_on_def)
+  with assms show ?thesis
+    by (simp add: quotient_def card_UN_disjoint)
+qed
 
 
 subsection \<open>Projection\<close>
--- a/src/HOL/Fields.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Fields.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -125,11 +125,14 @@
 qed
 
 lemma inverse_zero_imp_zero:
-  "inverse a = 0 \<Longrightarrow> a = 0"
-apply (rule classical)
-apply (drule nonzero_imp_inverse_nonzero)
-apply auto
-done
+  assumes "inverse a = 0" shows "a = 0"
+proof (rule ccontr)
+  assume "a \<noteq> 0"
+  then have "inverse a \<noteq> 0"
+    by (simp add: nonzero_imp_inverse_nonzero)
+  with assms show False
+    by auto
+qed
 
 lemma inverse_unique:
   assumes ab: "a * b = 1"
@@ -209,10 +212,10 @@
 lemma minus_divide_left: "- (a / b) = (-a) / b"
   by (simp add: divide_inverse)
 
-lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a / b) = a / (- b)"
+lemma nonzero_minus_divide_right: "b \<noteq> 0 \<Longrightarrow> - (a / b) = a / (- b)"
   by (simp add: divide_inverse nonzero_inverse_minus_eq)
 
-lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
+lemma nonzero_minus_divide_divide: "b \<noteq> 0 \<Longrightarrow> (-a) / (-b) = a / b"
   by (simp add: divide_inverse nonzero_inverse_minus_eq)
 
 lemma divide_minus_left [simp]: "(-a) / b = - (a / b)"
@@ -712,10 +715,16 @@
 qed
 
 lemma inverse_less_imp_less:
-  "inverse a < inverse b \<Longrightarrow> 0 < a \<Longrightarrow> b < a"
-apply (simp add: less_le [of "inverse a"] less_le [of "b"])
-apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq)
-done
+  assumes "inverse a < inverse b" "0 < a"
+  shows "b < a"
+proof -
+  have "a \<noteq> b"
+    using assms by (simp add: less_le)
+  moreover have "b \<le> a"
+    using assms by (force simp: less_le dest: inverse_le_imp_le)
+  ultimately show ?thesis
+    by (simp add: less_le)
+qed
 
 text\<open>Both premises are essential. Consider -1 and 1.\<close>
 lemma inverse_less_iff_less [simp]:
@@ -734,41 +743,44 @@
 text\<open>These results refer to both operands being negative.  The opposite-sign
 case is trivial, since inverse preserves signs.\<close>
 lemma inverse_le_imp_le_neg:
-  "inverse a \<le> inverse b \<Longrightarrow> b < 0 \<Longrightarrow> b \<le> a"
-apply (rule classical)
-apply (subgoal_tac "a < 0")
- prefer 2 apply force
-apply (insert inverse_le_imp_le [of "-b" "-a"])
-apply (simp add: nonzero_inverse_minus_eq)
-done
+  assumes "inverse a \<le> inverse b" "b < 0"
+  shows "b \<le> a"
+proof (rule classical)
+  assume "\<not> b \<le> a"
+  with \<open>b < 0\<close> have "a < 0"
+    by force
+  with assms show "b \<le> a"
+    using inverse_le_imp_le [of "-b" "-a"] by (simp add: nonzero_inverse_minus_eq)
+qed
 
 lemma less_imp_inverse_less_neg:
-   "a < b \<Longrightarrow> b < 0 \<Longrightarrow> inverse b < inverse a"
-apply (subgoal_tac "a < 0")
- prefer 2 apply (blast intro: less_trans)
-apply (insert less_imp_inverse_less [of "-b" "-a"])
-apply (simp add: nonzero_inverse_minus_eq)
-done
+  assumes "a < b" "b < 0"
+  shows "inverse b < inverse a"
+proof -
+  have "a < 0"
+    using assms by (blast intro: less_trans)
+  with less_imp_inverse_less [of "-b" "-a"] show ?thesis
+    by (simp add: nonzero_inverse_minus_eq assms)
+qed
 
 lemma inverse_less_imp_less_neg:
-   "inverse a < inverse b \<Longrightarrow> b < 0 \<Longrightarrow> b < a"
-apply (rule classical)
-apply (subgoal_tac "a < 0")
- prefer 2
- apply force
-apply (insert inverse_less_imp_less [of "-b" "-a"])
-apply (simp add: nonzero_inverse_minus_eq)
-done
+  assumes "inverse a < inverse b" "b < 0"
+  shows "b < a"
+proof (rule classical)
+  assume "\<not> b < a"
+  with \<open>b < 0\<close> have "a < 0"
+    by force
+  with inverse_less_imp_less [of "-b" "-a"] show ?thesis
+    by (simp add: nonzero_inverse_minus_eq assms)
+qed
 
 lemma inverse_less_iff_less_neg [simp]:
   "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> inverse a < inverse b \<longleftrightarrow> b < a"
-apply (insert inverse_less_iff_less [of "-b" "-a"])
-apply (simp del: inverse_less_iff_less
-            add: nonzero_inverse_minus_eq)
-done
+  using inverse_less_iff_less [of "-b" "-a"]
+  by (simp del: inverse_less_iff_less add: nonzero_inverse_minus_eq)
 
 lemma le_imp_inverse_le_neg:
-  "a \<le> b \<Longrightarrow> b < 0 ==> inverse b \<le> inverse a"
+  "a \<le> b \<Longrightarrow> b < 0 \<Longrightarrow> inverse b \<le> inverse a"
   by (force simp add: le_less less_imp_inverse_less_neg)
 
 lemma inverse_le_iff_le_neg [simp]:
@@ -907,113 +919,125 @@
   by (subst le_iff_diff_le_0) (simp add: diff_frac_eq )
 
 lemma divide_pos_pos[simp]:
-  "0 < x ==> 0 < y ==> 0 < x / y"
+  "0 < x \<Longrightarrow> 0 < y \<Longrightarrow> 0 < x / y"
 by(simp add:field_simps)
 
 lemma divide_nonneg_pos:
-  "0 <= x ==> 0 < y ==> 0 <= x / y"
+  "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> 0 \<le> x / y"
 by(simp add:field_simps)
 
 lemma divide_neg_pos:
-  "x < 0 ==> 0 < y ==> x / y < 0"
-by(simp add:field_simps)
+  "x < 0 \<Longrightarrow> 0 < y \<Longrightarrow> x / y < 0"
+  by(simp add:field_simps)
 
 lemma divide_nonpos_pos:
-  "x <= 0 ==> 0 < y ==> x / y <= 0"
-by(simp add:field_simps)
+  "x \<le> 0 \<Longrightarrow> 0 < y \<Longrightarrow> x / y \<le> 0"
+  by(simp add:field_simps)
 
 lemma divide_pos_neg:
-  "0 < x ==> y < 0 ==> x / y < 0"
-by(simp add:field_simps)
+  "0 < x \<Longrightarrow> y < 0 \<Longrightarrow> x / y < 0"
+  by(simp add:field_simps)
 
 lemma divide_nonneg_neg:
-  "0 <= x ==> y < 0 ==> x / y <= 0"
-by(simp add:field_simps)
+  "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> x / y \<le> 0"
+  by(simp add:field_simps)
 
 lemma divide_neg_neg:
-  "x < 0 ==> y < 0 ==> 0 < x / y"
-by(simp add:field_simps)
+  "x < 0 \<Longrightarrow> y < 0 \<Longrightarrow> 0 < x / y"
+  by(simp add:field_simps)
 
 lemma divide_nonpos_neg:
-  "x <= 0 ==> y < 0 ==> 0 <= x / y"
-by(simp add:field_simps)
+  "x \<le> 0 \<Longrightarrow> y < 0 \<Longrightarrow> 0 \<le> x / y"
+  by(simp add:field_simps)
 
 lemma divide_strict_right_mono:
-     "[|a < b; 0 < c|] ==> a / c < b / c"
-by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono
-              positive_imp_inverse_positive)
+  "\<lbrakk>a < b; 0 < c\<rbrakk> \<Longrightarrow> a / c < b / c"
+  by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono
+      positive_imp_inverse_positive)
 
 
 lemma divide_strict_right_mono_neg:
-     "[|b < a; c < 0|] ==> a / c < b / c"
-apply (drule divide_strict_right_mono [of _ _ "-c"], simp)
-apply (simp add: less_imp_not_eq nonzero_minus_divide_right [symmetric])
-done
+  assumes "b < a" "c < 0" shows "a / c < b / c"
+proof -
+  have "b / - c < a / - c"
+    by (rule divide_strict_right_mono) (use assms in auto)
+  then show ?thesis
+    by (simp add: less_imp_not_eq)
+qed
 
 text\<open>The last premise ensures that \<^term>\<open>a\<close> and \<^term>\<open>b\<close>
       have the same sign\<close>
 lemma divide_strict_left_mono:
-  "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / b"
+  "\<lbrakk>b < a; 0 < c; 0 < a*b\<rbrakk> \<Longrightarrow> c / a < c / b"
   by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono)
 
 lemma divide_left_mono:
-  "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / b"
+  "\<lbrakk>b \<le> a; 0 \<le> c; 0 < a*b\<rbrakk> \<Longrightarrow> c / a \<le> c / b"
   by (auto simp: field_simps zero_less_mult_iff mult_right_mono)
 
 lemma divide_strict_left_mono_neg:
-  "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / b"
+  "\<lbrakk>a < b; c < 0; 0 < a*b\<rbrakk> \<Longrightarrow> c / a < c / b"
   by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono_neg)
 
-lemma mult_imp_div_pos_le: "0 < y ==> x <= z * y ==>
-    x / y <= z"
+lemma mult_imp_div_pos_le: "0 < y \<Longrightarrow> x \<le> z * y \<Longrightarrow> x / y \<le> z"
 by (subst pos_divide_le_eq, assumption+)
 
-lemma mult_imp_le_div_pos: "0 < y ==> z * y <= x ==>
-    z <= x / y"
+lemma mult_imp_le_div_pos: "0 < y \<Longrightarrow> z * y \<le> x \<Longrightarrow> z \<le> x / y"
 by(simp add:field_simps)
 
-lemma mult_imp_div_pos_less: "0 < y ==> x < z * y ==>
-    x / y < z"
+lemma mult_imp_div_pos_less: "0 < y \<Longrightarrow> x < z * y \<Longrightarrow> x / y < z"
 by(simp add:field_simps)
 
-lemma mult_imp_less_div_pos: "0 < y ==> z * y < x ==>
-    z < x / y"
+lemma mult_imp_less_div_pos: "0 < y \<Longrightarrow> z * y < x \<Longrightarrow> z < x / y"
 by(simp add:field_simps)
 
-lemma frac_le: "0 <= x ==>
-    x <= y ==> 0 < w ==> w <= z  ==> x / z <= y / w"
-  apply (rule mult_imp_div_pos_le)
-  apply simp
-  apply (subst times_divide_eq_left)
-  apply (rule mult_imp_le_div_pos, assumption)
-  apply (rule mult_mono)
-  apply simp_all
-done
+lemma frac_le:
+  assumes "0 \<le> y" "x \<le> y" "0 < w" "w \<le> z"
+  shows "x / z \<le> y / w"
+proof (rule mult_imp_div_pos_le)
+  show "z > 0"
+    using assms by simp
+  have "x \<le> y * z / w"
+  proof (rule mult_imp_le_div_pos [OF \<open>0 < w\<close>])
+    show "x * w \<le> y * z"
+      using assms by (auto intro: mult_mono)
+  qed
+  also have "... = y / w * z"
+    by simp
+  finally show "x \<le> y / w * z" .
+qed
 
-lemma frac_less: "0 <= x ==>
-    x < y ==> 0 < w ==> w <= z  ==> x / z < y / w"
-  apply (rule mult_imp_div_pos_less)
-  apply simp
-  apply (subst times_divide_eq_left)
-  apply (rule mult_imp_less_div_pos, assumption)
-  apply (erule mult_less_le_imp_less)
-  apply simp_all
-done
+lemma frac_less:
+  assumes "0 \<le> x" "x < y" "0 < w" "w \<le> z"
+  shows "x / z < y / w"
+proof (rule mult_imp_div_pos_less)
+  show "z > 0"
+    using assms by simp
+  have "x < y * z / w"
+  proof (rule mult_imp_less_div_pos [OF \<open>0 < w\<close>])
+    show "x * w < y * z"
+      using assms by (auto intro: mult_less_le_imp_less)
+  qed
+  also have "... = y / w * z"
+    by simp
+  finally show "x < y / w * z" .
+qed
 
-lemma frac_less2: "0 < x ==>
-    x <= y ==> 0 < w ==> w < z  ==> x / z < y / w"
-  apply (rule mult_imp_div_pos_less)
-  apply simp_all
-  apply (rule mult_imp_less_div_pos, assumption)
-  apply (erule mult_le_less_imp_less)
-  apply simp_all
-done
+lemma frac_less2:
+  assumes "0 < x" "x \<le> y" "0 < w" "w < z"
+  shows "x / z < y / w"
+proof (rule mult_imp_div_pos_less)
+  show "z > 0"
+    using assms by simp
+  show "x < y / w * z"
+    using assms by (force intro: mult_imp_less_div_pos mult_le_less_imp_less)
+qed
 
-lemma less_half_sum: "a < b ==> a < (a+b) / (1+1)"
-by (simp add: field_simps zero_less_two)
+lemma less_half_sum: "a < b \<Longrightarrow> a < (a+b) / (1+1)"
+  by (simp add: field_simps zero_less_two)
 
-lemma gt_half_sum: "a < b ==> (a+b)/(1+1) < b"
-by (simp add: field_simps zero_less_two)
+lemma gt_half_sum: "a < b \<Longrightarrow> (a+b)/(1+1) < b"
+  by (simp add: field_simps zero_less_two)
 
 subclass unbounded_dense_linorder
 proof
@@ -1037,11 +1061,11 @@
   by (cases b 0 rule: linorder_cases) simp_all
 
 lemma nonzero_abs_inverse:
-  "a \<noteq> 0 ==> \<bar>inverse a\<bar> = inverse \<bar>a\<bar>"
+  "a \<noteq> 0 \<Longrightarrow> \<bar>inverse a\<bar> = inverse \<bar>a\<bar>"
   by (rule abs_inverse)
 
 lemma nonzero_abs_divide:
-  "b \<noteq> 0 ==> \<bar>a / b\<bar> = \<bar>a\<bar> / \<bar>b\<bar>"
+  "b \<noteq> 0 \<Longrightarrow> \<bar>a / b\<bar> = \<bar>a\<bar> / \<bar>b\<bar>"
   by (rule abs_divide)
 
 lemma field_le_epsilon:
@@ -1055,24 +1079,24 @@
   then show "t \<le> y" by (simp add: algebra_simps)
 qed
 
-lemma inverse_positive_iff_positive [simp]:
-  "(0 < inverse a) = (0 < a)"
-apply (cases "a = 0", simp)
-apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
-done
+lemma inverse_positive_iff_positive [simp]: "(0 < inverse a) = (0 < a)"
+proof (cases "a = 0")
+  case False
+  then show ?thesis
+    by (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
+qed auto
 
-lemma inverse_negative_iff_negative [simp]:
-  "(inverse a < 0) = (a < 0)"
-apply (cases "a = 0", simp)
-apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative)
-done
+lemma inverse_negative_iff_negative [simp]: "(inverse a < 0) = (a < 0)"
+proof (cases "a = 0")
+  case False
+  then show ?thesis
+    by (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative)
+qed auto
 
-lemma inverse_nonnegative_iff_nonnegative [simp]:
-  "0 \<le> inverse a \<longleftrightarrow> 0 \<le> a"
+lemma inverse_nonnegative_iff_nonnegative [simp]: "0 \<le> inverse a \<longleftrightarrow> 0 \<le> a"
   by (simp add: not_less [symmetric])
 
-lemma inverse_nonpositive_iff_nonpositive [simp]:
-  "inverse a \<le> 0 \<longleftrightarrow> a \<le> 0"
+lemma inverse_nonpositive_iff_nonpositive [simp]: "inverse a \<le> 0 \<longleftrightarrow> a \<le> 0"
   by (simp add: not_less [symmetric])
 
 lemma one_less_inverse_iff: "1 < inverse x \<longleftrightarrow> 0 < x \<and> x < 1"
@@ -1144,20 +1168,14 @@
   by (simp add: divide_less_0_iff)
 
 lemma divide_right_mono:
-     "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/c"
-by (force simp add: divide_strict_right_mono le_less)
+  "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> a/c \<le> b/c"
+  by (force simp add: divide_strict_right_mono le_less)
 
-lemma divide_right_mono_neg: "a <= b
-    ==> c <= 0 ==> b / c <= a / c"
-apply (drule divide_right_mono [of _ _ "- c"])
-apply auto
-done
+lemma divide_right_mono_neg: "a \<le> b \<Longrightarrow> c \<le> 0 \<Longrightarrow> b / c \<le> a / c"
+  by (auto dest: divide_right_mono [of _ _ "- c"])
 
-lemma divide_left_mono_neg: "a <= b
-    ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b"
-  apply (drule divide_left_mono [of _ _ "- c"])
-  apply (auto simp add: mult.commute)
-done
+lemma divide_left_mono_neg: "a \<le> b \<Longrightarrow> c \<le> 0 \<Longrightarrow> 0 < a * b \<Longrightarrow> c / a \<le> c / b"
+  by (auto simp add: mult.commute dest: divide_left_mono [of _ _ "- c"])
 
 lemma inverse_le_iff: "inverse a \<le> inverse b \<longleftrightarrow> (0 < a * b \<longrightarrow> b \<le> a) \<and> (a * b \<le> 0 \<longrightarrow> a \<le> b)"
   by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases])
@@ -1176,19 +1194,19 @@
 
 lemma le_divide_eq_1:
   "(1 \<le> b / a) = ((0 < a \<and> a \<le> b) \<or> (a < 0 \<and> b \<le> a))"
-by (auto simp add: le_divide_eq)
+  by (auto simp add: le_divide_eq)
 
 lemma divide_le_eq_1:
   "(b / a \<le> 1) = ((0 < a \<and> b \<le> a) \<or> (a < 0 \<and> a \<le> b) \<or> a=0)"
-by (auto simp add: divide_le_eq)
+  by (auto simp add: divide_le_eq)
 
 lemma less_divide_eq_1:
   "(1 < b / a) = ((0 < a \<and> a < b) \<or> (a < 0 \<and> b < a))"
-by (auto simp add: less_divide_eq)
+  by (auto simp add: less_divide_eq)
 
 lemma divide_less_eq_1:
   "(b / a < 1) = ((0 < a \<and> b < a) \<or> (a < 0 \<and> a < b) \<or> a=0)"
-by (auto simp add: divide_less_eq)
+  by (auto simp add: divide_less_eq)
 
 lemma divide_nonneg_nonneg [simp]:
   "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x / y"
@@ -1210,55 +1228,52 @@
 
 lemma le_divide_eq_1_pos [simp]:
   "0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)"
-by (auto simp add: le_divide_eq)
+  by (auto simp add: le_divide_eq)
 
 lemma le_divide_eq_1_neg [simp]:
   "a < 0 \<Longrightarrow> (1 \<le> b/a) = (b \<le> a)"
-by (auto simp add: le_divide_eq)
+  by (auto simp add: le_divide_eq)
 
 lemma divide_le_eq_1_pos [simp]:
   "0 < a \<Longrightarrow> (b/a \<le> 1) = (b \<le> a)"
-by (auto simp add: divide_le_eq)
+  by (auto simp add: divide_le_eq)
 
 lemma divide_le_eq_1_neg [simp]:
   "a < 0 \<Longrightarrow> (b/a \<le> 1) = (a \<le> b)"
-by (auto simp add: divide_le_eq)
+  by (auto simp add: divide_le_eq)
 
 lemma less_divide_eq_1_pos [simp]:
   "0 < a \<Longrightarrow> (1 < b/a) = (a < b)"
-by (auto simp add: less_divide_eq)
+  by (auto simp add: less_divide_eq)
 
 lemma less_divide_eq_1_neg [simp]:
   "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)"
-by (auto simp add: less_divide_eq)
+  by (auto simp add: less_divide_eq)
 
 lemma divide_less_eq_1_pos [simp]:
   "0 < a \<Longrightarrow> (b/a < 1) = (b < a)"
-by (auto simp add: divide_less_eq)
+  by (auto simp add: divide_less_eq)
 
 lemma divide_less_eq_1_neg [simp]:
   "a < 0 \<Longrightarrow> b/a < 1 \<longleftrightarrow> a < b"
-by (auto simp add: divide_less_eq)
+  by (auto simp add: divide_less_eq)
 
 lemma eq_divide_eq_1 [simp]:
   "(1 = b/a) = ((a \<noteq> 0 \<and> a = b))"
-by (auto simp add: eq_divide_eq)
+  by (auto simp add: eq_divide_eq)
 
 lemma divide_eq_eq_1 [simp]:
   "(b/a = 1) = ((a \<noteq> 0 \<and> a = b))"
-by (auto simp add: divide_eq_eq)
+  by (auto simp add: divide_eq_eq)
 
-lemma abs_div_pos: "0 < y ==>
-    \<bar>x\<bar> / y = \<bar>x / y\<bar>"
-  apply (subst abs_divide)
-  apply (simp add: order_less_imp_le)
-done
+lemma abs_div_pos: "0 < y \<Longrightarrow> \<bar>x\<bar> / y = \<bar>x / y\<bar>"
+  by (simp add: order_less_imp_le)
 
 lemma zero_le_divide_abs_iff [simp]: "(0 \<le> a / \<bar>b\<bar>) = (0 \<le> a \<or> b = 0)"
-by (auto simp: zero_le_divide_iff)
+  by (auto simp: zero_le_divide_iff)
 
 lemma divide_le_0_abs_iff [simp]: "(a / \<bar>b\<bar> \<le> 0) = (a \<le> 0 \<or> b = 0)"
-by (auto simp: divide_le_0_iff)
+  by (auto simp: divide_le_0_iff)
 
 lemma field_le_mult_one_interval:
   assumes *: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
@@ -1279,13 +1294,14 @@
 text\<open>For creating values between \<^term>\<open>u\<close> and \<^term>\<open>v\<close>.\<close>
 lemma scaling_mono:
   assumes "u \<le> v" "0 \<le> r" "r \<le> s"
-    shows "u + r * (v - u) / s \<le> v"
+  shows "u + r * (v - u) / s \<le> v"
 proof -
   have "r/s \<le> 1" using assms
     using divide_le_eq_1 by fastforce
-  then have "(r/s) * (v - u) \<le> 1 * (v - u)"
-    apply (rule mult_right_mono)
+  moreover have "0 \<le> v - u"
     using assms by simp
+  ultimately have "(r/s) * (v - u) \<le> 1 * (v - u)"
+    by (rule mult_right_mono)
   then show ?thesis
     by (simp add: field_simps)
 qed
--- a/src/HOL/Filter.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Filter.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -300,20 +300,15 @@
   "eventually P (inf F F') \<longleftrightarrow>
    (\<exists>Q R. eventually Q F \<and> eventually R F' \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))"
   unfolding inf_filter_def
-  apply (rule eventually_Abs_filter, rule is_filter.intro)
-  apply (fast intro: eventually_True)
-  apply clarify
-  apply (intro exI conjI)
-  apply (erule (1) eventually_conj)
-  apply (erule (1) eventually_conj)
-  apply simp
-  apply auto
+  apply (rule eventually_Abs_filter [OF is_filter.intro])
+  apply (blast intro: eventually_True)
+   apply (force elim!: eventually_conj)+
   done
 
 lemma eventually_Sup:
   "eventually P (Sup S) \<longleftrightarrow> (\<forall>F\<in>S. eventually P F)"
   unfolding Sup_filter_def
-  apply (rule eventually_Abs_filter, rule is_filter.intro)
+  apply (rule eventually_Abs_filter [OF is_filter.intro])
   apply (auto intro: eventually_conj elim!: eventually_rev_mp)
   done
 
@@ -557,8 +552,7 @@
 lemma eventually_filtermap:
   "eventually P (filtermap f F) = eventually (\<lambda>x. P (f x)) F"
   unfolding filtermap_def
-  apply (rule eventually_Abs_filter)
-  apply (rule is_filter.intro)
+  apply (rule eventually_Abs_filter [OF is_filter.intro])
   apply (auto elim!: eventually_rev_mp)
   done
 
@@ -746,10 +740,7 @@
 
 lemma le_principal: "F \<le> principal A \<longleftrightarrow> eventually (\<lambda>x. x \<in> A) F"
   unfolding le_filter_def eventually_principal
-  apply safe
-  apply (erule_tac x="\<lambda>x. x \<in> A" in allE)
-  apply (auto elim: eventually_mono)
-  done
+  by (force elim: eventually_mono)
 
 lemma principal_inject[iff]: "principal A = principal B \<longleftrightarrow> A = B"
   unfolding eq_iff by simp
@@ -1172,11 +1163,7 @@
 
 lemma eventually_prod_same: "eventually P (F \<times>\<^sub>F F) \<longleftrightarrow>
     (\<exists>Q. eventually Q F \<and> (\<forall>x y. Q x \<longrightarrow> Q y \<longrightarrow> P (x, y)))"
-  unfolding eventually_prod_filter
-  apply safe
-  apply (rule_tac x="inf Pf Pg" in exI)
-  apply (auto simp: inf_fun_def intro!: eventually_conj)
-  done
+  unfolding eventually_prod_filter by (blast intro!: eventually_conj)
 
 lemma eventually_prod_sequentially:
   "eventually P (sequentially \<times>\<^sub>F sequentially) \<longleftrightarrow> (\<exists>N. \<forall>m \<ge> N. \<forall>n \<ge> N. P (n, m))"
@@ -1227,13 +1214,15 @@
   using prod_filter_INF[of "{A}" J "\<lambda>x. x" B] by simp
 
 lemma prod_filtermap1: "prod_filter (filtermap f F) G = filtermap (apfst f) (prod_filter F G)"
-  apply(clarsimp simp add: filter_eq_iff eventually_filtermap eventually_prod_filter; safe)
+  unfolding filter_eq_iff eventually_filtermap eventually_prod_filter
+  apply safe
   subgoal by auto
   subgoal for P Q R by(rule exI[where x="\<lambda>y. \<exists>x. y = f x \<and> Q x"])(auto intro: eventually_mono)
   done
 
 lemma prod_filtermap2: "prod_filter F (filtermap g G) = filtermap (apsnd g) (prod_filter F G)"
-  apply(clarsimp simp add: filter_eq_iff eventually_filtermap eventually_prod_filter; safe)
+  unfolding filter_eq_iff eventually_filtermap eventually_prod_filter
+  apply safe
   subgoal by auto
   subgoal for P Q R  by(auto intro: exI[where x="\<lambda>y. \<exists>x. y = g x \<and> R x"] eventually_mono)
   done
@@ -1291,10 +1280,12 @@
   assumes "filterlim f F G" and ord: "F \<le> F'" "G' \<le> G"
   assumes eq: "eventually (\<lambda>x. f x = f' x) G'"
   shows "filterlim f' F' G'"
-  apply (rule filterlim_cong[OF refl refl eq, THEN iffD1])
-  apply (rule filterlim_mono[OF _ ord])
-  apply fact
-  done
+proof -
+  have "filterlim f F' G'"
+    by (simp add: filterlim_mono[OF _ ord] assms)
+  then show ?thesis
+    by (rule filterlim_cong[OF refl refl eq, THEN iffD1])
+qed
 
 lemma filtermap_mono_strong: "inj f \<Longrightarrow> filtermap f F \<le> filtermap f G \<longleftrightarrow> F \<le> G"
   apply (safe intro!: filtermap_mono)
@@ -1458,8 +1449,7 @@
   assume "\<forall>Z. eventually (\<lambda>x. f x \<le> Z) F"
   hence "eventually (\<lambda>x. f x \<le> Z') F" by auto
   thus "eventually (\<lambda>x. f x < Z) F"
-    apply (rule eventually_mono)
-    using 1 by auto
+    by (rule eventually_mono) (use 1 in auto)
   next
     fix Z :: 'b
     show "\<forall>Z. eventually (\<lambda>x. f x < Z) F \<Longrightarrow> eventually (\<lambda>x. f x \<le> Z) F"
--- a/src/HOL/Fun.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Fun.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -738,7 +738,7 @@
   by (simp add: fun_eq_iff)
 
 lemma fun_upd_twist: "a \<noteq> c \<Longrightarrow> (m(a := b))(c := d) = (m(c := d))(a := b)"
-  by (rule ext) auto
+  by auto
 
 lemma inj_on_fun_updI: "inj_on f A \<Longrightarrow> y \<notin> f ` A \<Longrightarrow> inj_on (f(x := y)) A"
   by (auto simp: inj_on_def)
--- a/src/HOL/Groups.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Groups.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -635,9 +635,7 @@
 
 text \<open>non-strict, in both arguments\<close>
 lemma add_mono: "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c \<le> b + d"
-  apply (erule add_right_mono [THEN order_trans])
-  apply (simp add: add.commute add_left_mono)
-  done
+  by (simp add: add.commute add_left_mono add_right_mono [THEN order_trans])
 
 end
 
@@ -656,20 +654,16 @@
   by (simp add: add.commute [of _ c] add_strict_left_mono)
 
 subclass strict_ordered_ab_semigroup_add
-  apply standard
-  apply (erule add_strict_right_mono [THEN less_trans])
-  apply (erule add_strict_left_mono)
-  done
+proof
+  show "\<And>a b c d. \<lbrakk>a < b; c < d\<rbrakk> \<Longrightarrow> a + c < b + d"
+    by (iprover intro: add_strict_left_mono add_strict_right_mono less_trans)
+qed
 
 lemma add_less_le_mono: "a < b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c < b + d"
-  apply (erule add_strict_right_mono [THEN less_le_trans])
-  apply (erule add_left_mono)
-  done
+  by (iprover intro: add_left_mono add_strict_right_mono less_le_trans)
 
 lemma add_le_less_mono: "a \<le> b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"
-  apply (erule add_right_mono [THEN le_less_trans])
-  apply (erule add_strict_left_mono)
-  done
+  by (iprover intro: add_strict_left_mono add_right_mono less_le_trans)
 
 end
 
@@ -684,12 +678,7 @@
   from less have le: "c + a \<le> c + b"
     by (simp add: order_le_less)
   have "a \<le> b"
-    apply (insert le)
-    apply (drule add_le_imp_le_left)
-    apply (insert le)
-    apply (drule add_le_imp_le_left)
-    apply assumption
-    done
+    using add_le_imp_le_left [OF le] .
   moreover have "a \<noteq> b"
   proof (rule ccontr)
     assume "\<not> ?thesis"
@@ -711,10 +700,7 @@
   by (blast intro: add_less_imp_less_right add_strict_right_mono)
 
 lemma add_le_cancel_left [simp]: "c + a \<le> c + b \<longleftrightarrow> a \<le> b"
-  apply auto
-   apply (drule add_le_imp_le_left)
-   apply (simp_all add: add_left_mono)
-  done
+  by (auto simp: dest: add_le_imp_le_left add_left_mono)
 
 lemma add_le_cancel_right [simp]: "a + c \<le> b + c \<longleftrightarrow> a \<le> b"
   by (simp add: add.commute [of a c] add.commute [of b c])
@@ -927,17 +913,7 @@
 qed
 
 lemma le_minus_iff: "a \<le> - b \<longleftrightarrow> b \<le> - a"
-proof -
-  have mm: "- (- a) < - b \<Longrightarrow> - (- b) < -a" for a b :: 'a
-    by (simp only: minus_less_iff)
-  have "- (- a) \<le> - b \<longleftrightarrow> b \<le> - a"
-    apply (auto simp only: le_less)
-      apply (drule mm)
-      apply (simp_all)
-    apply (drule mm[simplified], assumption)
-    done
-  then show ?thesis by simp
-qed
+  by (auto simp: order.order_iff_strict less_minus_iff)
 
 lemma minus_le_iff: "- a \<le> b \<longleftrightarrow> - b \<le> a"
   by (auto simp add: le_less minus_less_iff)
@@ -955,17 +931,17 @@
 
 lemma diff_less_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]:
   "a - b < c \<longleftrightarrow> a < c + b"
-  apply (subst less_iff_diff_less_0 [of a])
-  apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
-  apply (simp add: algebra_simps)
-  done
+proof (subst less_iff_diff_less_0 [of a])
+  show "(a - b < c) = (a - (c + b) < 0)"
+    by (simp add: algebra_simps less_iff_diff_less_0 [of _ c])
+qed
 
 lemma less_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]:
   "a < c - b \<longleftrightarrow> a + b < c"
-  apply (subst less_iff_diff_less_0 [of "a + b"])
-  apply (subst less_iff_diff_less_0 [of a])
-  apply (simp add: algebra_simps)
-  done
+proof (subst less_iff_diff_less_0 [of "a + b"])
+  show "(a < c - b) = (a + b - c < 0)"
+    by (simp add: algebra_simps less_iff_diff_less_0 [of a])
+qed
 
 lemma diff_gt_0_iff_gt [simp]: "a - b > 0 \<longleftrightarrow> a > b"
   by (simp add: less_diff_eq)
@@ -1070,11 +1046,10 @@
     assume *: "\<not> ?thesis"
     then have "b \<le> a" by (simp add: linorder_not_le)
     then have "c + b \<le> c + a" by (rule add_left_mono)
-    with le1 have "a = b"
-      apply -
-      apply (drule antisym)
-       apply simp_all
-      done
+    then have "c + a = c + b"
+      using le1 by (iprover intro: antisym)
+    then have "a = b"
+      by simp
     with * show False
       by (simp add: linorder_not_le [symmetric])
   qed
@@ -1146,10 +1121,7 @@
 qed
 
 lemma double_zero_sym [simp]: "0 = a + a \<longleftrightarrow> a = 0"
-  apply (rule iffI)
-   apply (drule sym)
-   apply simp_all
-  done
+  using double_zero [of a] by (simp only: eq_commute)
 
 lemma zero_less_double_add_iff_zero_less_single_add [simp]: "0 < a + a \<longleftrightarrow> 0 < a"
 proof
@@ -1358,13 +1330,17 @@
 
 lemma dense_eq0_I:
   fixes x::"'a::{dense_linorder,ordered_ab_group_add_abs}"
-  shows "(\<And>e. 0 < e \<Longrightarrow> \<bar>x\<bar> \<le> e) \<Longrightarrow> x = 0"
-  apply (cases "\<bar>x\<bar> = 0")
-   apply simp
-  apply (simp only: zero_less_abs_iff [symmetric])
-  apply (drule dense)
-  apply (auto simp add: not_less [symmetric])
-  done
+  assumes "\<And>e. 0 < e \<Longrightarrow> \<bar>x\<bar> \<le> e"
+  shows "x = 0"
+proof (cases "\<bar>x\<bar> = 0")
+  case False
+  then have "\<bar>x\<bar> > 0"
+    by simp
+  then obtain z where "0 < z" "z < \<bar>x\<bar>"
+    using dense by force
+  then show ?thesis
+    using assms by (simp flip: not_less)
+qed auto
 
 hide_fact (open) ab_diff_conv_add_uminus add_0 mult_1 ab_left_minus
 
--- a/src/HOL/HOL.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/HOL.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -249,11 +249,7 @@
         |   |
         c = d   *)
 lemma box_equals: "\<lbrakk>a = b; a = c; b = d\<rbrakk> \<Longrightarrow> c = d"
-  apply (rule trans)
-   apply (rule trans)
-    apply (rule sym)
-    apply assumption+
-  done
+  by (iprover intro: sym trans)
 
 text \<open>For calculational reasoning:\<close>
 
@@ -268,25 +264,17 @@
 
 text \<open>Similar to \<open>AP_THM\<close> in Gordon's HOL.\<close>
 lemma fun_cong: "(f :: 'a \<Rightarrow> 'b) = g \<Longrightarrow> f x = g x"
-  apply (erule subst)
-  apply (rule refl)
-  done
+  by (iprover intro: refl elim: subst)
 
 text \<open>Similar to \<open>AP_TERM\<close> in Gordon's HOL and FOL's \<open>subst_context\<close>.\<close>
 lemma arg_cong: "x = y \<Longrightarrow> f x = f y"
-  apply (erule subst)
-  apply (rule refl)
-  done
+  by (iprover intro: refl elim: subst)
 
 lemma arg_cong2: "\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> f a c = f b d"
-  apply (erule ssubst)+
-  apply (rule refl)
-  done
+  by (iprover intro: refl elim: subst)
 
 lemma cong: "\<lbrakk>f = g; (x::'a) = y\<rbrakk> \<Longrightarrow> f x = g y"
-  apply (erule subst)+
-  apply (rule refl)
-  done
+  by (iprover intro: refl elim: subst)
 
 ML \<open>fun cong_tac ctxt = Cong_Tac.cong_tac ctxt @{thm cong}\<close>
 
@@ -324,20 +312,15 @@
 subsubsection \<open>Universal quantifier (1)\<close>
 
 lemma spec: "\<forall>x::'a. P x \<Longrightarrow> P x"
-  apply (unfold All_def)
-  apply (rule eqTrueE)
-  apply (erule fun_cong)
-  done
+  unfolding All_def by (iprover intro: eqTrueE fun_cong)
 
 lemma allE:
-  assumes major: "\<forall>x. P x"
-    and minor: "P x \<Longrightarrow> R"
+  assumes major: "\<forall>x. P x" and minor: "P x \<Longrightarrow> R"
   shows R
   by (iprover intro: minor major [THEN spec])
 
 lemma all_dupE:
-  assumes major: "\<forall>x. P x"
-    and minor: "\<lbrakk>P x; \<forall>x. P x\<rbrakk> \<Longrightarrow> R"
+  assumes major: "\<forall>x. P x" and minor: "\<lbrakk>P x; \<forall>x. P x\<rbrakk> \<Longrightarrow> R"
   shows R
   by (iprover intro: minor major major [THEN spec])
 
@@ -350,9 +333,7 @@
 \<close>
 
 lemma FalseE: "False \<Longrightarrow> P"
-  apply (unfold False_def)
-  apply (erule spec)
-  done
+  unfolding False_def by (erule spec)
 
 lemma False_neq_True: "False = True \<Longrightarrow> P"
   by (erule eqTrueE [THEN FalseE])
@@ -363,29 +344,17 @@
 lemma notI:
   assumes "P \<Longrightarrow> False"
   shows "\<not> P"
-  apply (unfold not_def)
-  apply (iprover intro: impI assms)
-  done
+  unfolding not_def by (iprover intro: impI assms)
 
 lemma False_not_True: "False \<noteq> True"
-  apply (rule notI)
-  apply (erule False_neq_True)
-  done
+  by (iprover intro: notI elim: False_neq_True)
 
 lemma True_not_False: "True \<noteq> False"
-  apply (rule notI)
-  apply (drule sym)
-  apply (erule False_neq_True)
-  done
+  by (iprover intro: notI dest: sym elim: False_neq_True)
 
 lemma notE: "\<lbrakk>\<not> P; P\<rbrakk> \<Longrightarrow> R"
-  apply (unfold not_def)
-  apply (erule mp [THEN FalseE])
-  apply assumption
-  done
-
-lemma notI2: "(P \<Longrightarrow> \<not> Pa) \<Longrightarrow> (P \<Longrightarrow> Pa) \<Longrightarrow> \<not> P"
-  by (erule notE [THEN notI]) (erule meta_mp)
+  unfolding not_def
+  by (iprover intro: mp [THEN FalseE])
 
 
 subsubsection \<open>Implication\<close>
@@ -397,7 +366,7 @@
 
 text \<open>Reduces \<open>Q\<close> to \<open>P \<longrightarrow> Q\<close>, allowing substitution in \<open>P\<close>.\<close>
 lemma rev_mp: "\<lbrakk>P; P \<longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-  by (iprover intro: mp)
+  by (rule mp)
 
 lemma contrapos_nn:
   assumes major: "\<not> Q"
@@ -510,10 +479,10 @@
   assumes major: "P \<and> Q"
     and minor: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R"
   shows R
-  apply (rule minor)
-   apply (rule major [THEN conjunct1])
-  apply (rule major [THEN conjunct2])
-  done
+proof (rule minor)
+  show P by (rule major [THEN conjunct1])
+  show Q by (rule major [THEN conjunct2])
+qed
 
 lemma context_conjI:
   assumes P "P \<Longrightarrow> Q"
@@ -533,14 +502,18 @@
 subsubsection \<open>Classical logic\<close>
 
 lemma classical:
-  assumes prem: "\<not> P \<Longrightarrow> P"
+  assumes "\<not> P \<Longrightarrow> P"
   shows P
-  apply (rule True_or_False [THEN disjE, THEN eqTrueE])
-   apply assumption
-  apply (rule notI [THEN prem, THEN eqTrueI])
-  apply (erule subst)
-  apply assumption
-  done
+proof (rule True_or_False [THEN disjE])
+  show P if "P = True"
+    using that by (iprover intro: eqTrueE)
+  show P if "P = False"
+  proof (intro notI assms)
+    assume P
+    with that show False
+      by (iprover elim: subst)
+  qed
+qed
 
 lemmas ccontr = FalseE [THEN classical]
 
@@ -550,16 +523,12 @@
   assumes premp: P
     and premnot: "\<not> R \<Longrightarrow> \<not> P"
   shows R
-  apply (rule ccontr)
-  apply (erule notE [OF premnot premp])
-  done
+  by (iprover intro: ccontr notE [OF premnot premp])
+
 
 text \<open>Double negation law.\<close>
 lemma notnotD: "\<not>\<not> P \<Longrightarrow> P"
-  apply (rule classical)
-  apply (erule notE)
-  apply assumption
-  done
+  by (iprover intro: ccontr notE )
 
 lemma contrapos_pp:
   assumes p1: Q
@@ -583,19 +552,15 @@
   by (iprover intro: ex_prem [THEN exE] ex1I eq)
 
 lemma ex1E:
-  assumes major: "\<exists>!x. P x"
-    and minor: "\<And>x. \<lbrakk>P x; \<forall>y. P y \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R"
+  assumes major: "\<exists>!x. P x" and minor: "\<And>x. \<lbrakk>P x; \<forall>y. P y \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R"
   shows R
-  apply (rule major [unfolded Ex1_def, THEN exE])
-  apply (erule conjE)
-  apply (iprover intro: minor)
-  done
+proof (rule major [unfolded Ex1_def, THEN exE])
+  show "\<And>x. P x \<and> (\<forall>y. P y \<longrightarrow> y = x) \<Longrightarrow> R"
+    by (iprover intro: minor elim: conjE)
+qed
 
 lemma ex1_implies_ex: "\<exists>!x. P x \<Longrightarrow> \<exists>x. P x"
-  apply (erule ex1E)
-  apply (rule exI)
-  apply assumption
-  done
+  by (iprover intro: exI elim: ex1E)
 
 
 subsubsection \<open>Classical intro rules for disjunction and existential quantifiers\<close>
@@ -613,22 +578,18 @@
   Note that \<open>\<not> P\<close> is the second case, not the first.
 \<close>
 lemma case_split [case_names True False]:
-  assumes prem1: "P \<Longrightarrow> Q"
-    and prem2: "\<not> P \<Longrightarrow> Q"
+  assumes "P \<Longrightarrow> Q" "\<not> P \<Longrightarrow> Q"
   shows Q
-  apply (rule excluded_middle [THEN disjE])
-   apply (erule prem2)
-  apply (erule prem1)
-  done
+  using excluded_middle [of P]
+    by (iprover intro: assms elim: disjE)
 
 text \<open>Classical implies (\<open>\<longrightarrow>\<close>) elimination.\<close>
 lemma impCE:
   assumes major: "P \<longrightarrow> Q"
     and minor: "\<not> P \<Longrightarrow> R" "Q \<Longrightarrow> R"
   shows R
-  apply (rule excluded_middle [of P, THEN disjE])
-   apply (iprover intro: minor major [THEN mp])+
-  done
+  using excluded_middle [of P]
+  by (iprover intro: minor major [THEN mp] elim: disjE)+
 
 text \<open>
   This version of \<open>\<longrightarrow>\<close> elimination works on \<open>Q\<close> before \<open>P\<close>.  It works best for
@@ -639,9 +600,8 @@
   assumes major: "P \<longrightarrow> Q"
     and minor: "Q \<Longrightarrow> R" "\<not> P \<Longrightarrow> R"
   shows R
-  apply (rule excluded_middle [of P, THEN disjE])
-   apply (iprover intro: minor major [THEN mp])+
-  done
+  using assms by (elim impCE)
+
 
 text \<open>Classical \<open>\<longleftrightarrow>\<close> elimination.\<close>
 lemma iffCE:
@@ -874,9 +834,7 @@
 ML \<open>val HOL_cs = claset_of \<^context>\<close>
 
 lemma contrapos_np: "\<not> Q \<Longrightarrow> (\<not> P \<Longrightarrow> Q) \<Longrightarrow> P"
-  apply (erule swap)
-  apply (erule (1) meta_mp)
-  done
+  by (erule swap)
 
 declare ex_ex1I [rule del, intro! 2]
   and ex1I [intro]
@@ -889,15 +847,13 @@
 text \<open>Better than \<open>ex1E\<close> for classical reasoner: needs no quantifier duplication!\<close>
 lemma alt_ex1E [elim!]:
   assumes major: "\<exists>!x. P x"
-    and prem: "\<And>x. \<lbrakk>P x; \<forall>y y'. P y \<and> P y' \<longrightarrow> y = y'\<rbrakk> \<Longrightarrow> R"
+    and minor: "\<And>x. \<lbrakk>P x; \<forall>y y'. P y \<and> P y' \<longrightarrow> y = y'\<rbrakk> \<Longrightarrow> R"
   shows R
-  apply (rule ex1E [OF major])
-  apply (rule prem)
-   apply assumption
-  apply (rule allI)+
-  apply (tactic \<open>eresolve_tac \<^context> [Classical.dup_elim \<^context> @{thm allE}] 1\<close>)
-  apply iprover
-  done
+proof (rule ex1E [OF major minor])
+  show "\<forall>y y'. P y \<and> P y' \<longrightarrow> y = y'" if "P x" and \<section>: "\<forall>y. P y \<longrightarrow> y = x" for x
+    using \<open>P x\<close> \<section> \<section> by fast
+qed assumption
+
 
 ML \<open>
   structure Blast = Blast
@@ -1101,12 +1057,12 @@
   unfolding If_def by blast
 
 lemma if_split: "P (if Q then x else y) = ((Q \<longrightarrow> P x) \<and> (\<not> Q \<longrightarrow> P y))"
-  apply (rule case_split [of Q])
-   apply (simplesubst if_P)
-    prefer 3
-    apply (simplesubst if_not_P)
-     apply blast+
-  done
+proof (rule case_split [of Q])
+  show ?thesis if Q
+    using that by (simplesubst if_P) blast+
+  show ?thesis if "\<not> Q"
+    using that by (simplesubst if_not_P) blast+
+qed
 
 lemma if_split_asm: "P (if Q then x else y) = (\<not> ((Q \<and> \<not> P x) \<or> (\<not> Q \<and> \<not> P y)))"
   by (simplesubst if_split) blast
@@ -1150,20 +1106,15 @@
 lemma simp_impliesI:
   assumes PQ: "(PROP P \<Longrightarrow> PROP Q)"
   shows "PROP P =simp=> PROP Q"
-  apply (unfold simp_implies_def)
-  apply (rule PQ)
-  apply assumption
-  done
+  unfolding simp_implies_def
+  by (iprover intro: PQ)
 
 lemma simp_impliesE:
   assumes PQ: "PROP P =simp=> PROP Q"
     and P: "PROP P"
     and QR: "PROP Q \<Longrightarrow> PROP R"
   shows "PROP R"
-  apply (rule QR)
-  apply (rule PQ [unfolded simp_implies_def])
-  apply (rule P)
-  done
+  by (iprover intro: QR P PQ [unfolded simp_implies_def])
 
 lemma simp_implies_cong:
   assumes PP' :"PROP P \<equiv> PROP P'"
@@ -1658,22 +1609,32 @@
 lemma ex1_eq [iff]: "\<exists>!x. x = t" "\<exists>!x. t = x"
   by blast+
 
-lemma choice_eq: "(\<forall>x. \<exists>!y. P x y) = (\<exists>!f. \<forall>x. P x (f x))"
-  apply (rule iffI)
-   apply (rule_tac a = "\<lambda>x. THE y. P x y" in ex1I)
-    apply (fast dest!: theI')
-   apply (fast intro: the1_equality [symmetric])
-  apply (erule ex1E)
-  apply (rule allI)
-  apply (rule ex1I)
-   apply (erule spec)
-  apply (erule_tac x = "\<lambda>z. if z = x then y else f z" in allE)
-  apply (erule impE)
-   apply (rule allI)
-   apply (case_tac "xa = x")
-    apply (drule_tac [3] x = x in fun_cong)
-    apply simp_all
-  done
+lemma choice_eq: "(\<forall>x. \<exists>!y. P x y) = (\<exists>!f. \<forall>x. P x (f x))" (is "?lhs = ?rhs")
+proof (intro iffI allI)
+  assume L: ?lhs
+  then have \<section>: "\<forall>x. P x (THE y. P x y)"
+    by (best intro: theI')
+  show ?rhs
+    by (rule ex1I) (use L \<section> in \<open>fast+\<close>)
+next
+  fix x
+  assume R: ?rhs
+  then obtain f where f: "\<forall>x. P x (f x)" and f1: "\<And>y. (\<forall>x. P x (y x)) \<Longrightarrow> y = f"
+    by (blast elim: ex1E)
+  show "\<exists>!y. P x y"
+  proof (rule ex1I)
+    show "P x (f x)"
+      using f by blast
+    show "y = f x" if "P x y" for y
+    proof -
+      have "P z (if z = x then y else f z)" for z
+        using f that by (auto split: if_split)
+      with f1 [of "\<lambda>z. if z = x then y else f z"] f
+      show ?thesis
+        by (auto simp add: split: if_split_asm dest: fun_cong)
+    qed
+  qed
+qed
 
 lemmas eq_sym_conv = eq_commute
 
--- a/src/HOL/Hilbert_Choice.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Hilbert_Choice.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -49,27 +49,25 @@
   using ext[of P Q, OF assms] by simp
 
 text \<open>
-  Easier to apply than \<open>someI\<close> if the witness comes from an
+  Easier to use than \<open>someI\<close> if the witness comes from an
   existential formula.
 \<close>
 lemma someI_ex [elim?]: "\<exists>x. P x \<Longrightarrow> P (SOME x. P x)"
-  apply (erule exE)
-  apply (erule someI)
-  done
+  by (elim exE someI)
 
 lemma some_eq_imp:
   assumes "Eps P = a" "P b" shows "P a"
   using assms someI_ex by force
 
 text \<open>
-  Easier to apply than \<open>someI\<close> because the conclusion has only one
+  Easier to use than \<open>someI\<close> because the conclusion has only one
   occurrence of \<^term>\<open>P\<close>.
 \<close>
 lemma someI2: "P a \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> Q (SOME x. P x)"
   by (blast intro: someI)
 
 text \<open>
-  Easier to apply than \<open>someI2\<close> if the witness comes from an
+  Easier to use than \<open>someI2\<close> if the witness comes from an
   existential formula.
 \<close>
 lemma someI2_ex: "\<exists>a. P a \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> Q (SOME x. P x)"
@@ -94,10 +92,7 @@
   by (rule some_equality) (rule refl)
 
 lemma some_sym_eq_trivial [simp]: "(SOME y. x = y) = x"
-  apply (rule some_equality)
-   apply (rule refl)
-  apply (erule sym)
-  done
+  by (iprover intro: some_equality)
 
 
 subsection \<open>Axiom of Choice, Proved Using the Description Operator\<close>
@@ -240,11 +235,16 @@
 lemma surj_iff_all: "surj f \<longleftrightarrow> (\<forall>x. f (inv f x) = x)"
   by (simp add: o_def surj_iff fun_eq_iff)
 
-lemma surj_imp_inv_eq: "surj f \<Longrightarrow> \<forall>x. g (f x) = x \<Longrightarrow> inv f = g"
-  apply (rule ext)
-  apply (drule_tac x = "inv f x" in spec)
-  apply (simp add: surj_f_inv_f)
-  done
+lemma surj_imp_inv_eq:
+  assumes "surj f" and gf: "\<And>x. g (f x) = x"
+  shows "inv f = g"
+proof (rule ext)
+  fix x
+  have "g (f (inv f x)) = inv f x"
+    by (rule gf)
+  then show "inv f x = g x"
+    by (simp add: surj_f_inv_f \<open>surj f\<close>)
+qed
 
 lemma bij_imp_bij_inv: "bij f \<Longrightarrow> bij (inv f)"
   by (simp add: bij_def inj_imp_surj_inv surj_imp_inj_inv)
@@ -266,11 +266,7 @@
 lemma inv_into_comp:
   "inj_on f (g ` A) \<Longrightarrow> inj_on g A \<Longrightarrow> x \<in> f ` g ` A \<Longrightarrow>
     inv_into A (f \<circ> g) x = (inv_into A g \<circ> inv_into (g ` A) f) x"
-  apply (rule inv_into_f_eq)
-    apply (fast intro: comp_inj_on)
-   apply (simp add: inv_into_into)
-  apply (simp add: f_inv_into_f inv_into_into)
-  done
+  by (auto simp: f_inv_into_f inv_into_into intro: inv_into_f_eq comp_inj_on)
 
 lemma o_inv_distrib: "bij f \<Longrightarrow> bij g \<Longrightarrow> inv (f \<circ> g) = inv g \<circ> inv f"
   by (rule inv_equality) (auto simp add: bij_def surj_f_inv_f)
@@ -281,16 +277,25 @@
 lemma image_inv_f_f: "inj f \<Longrightarrow> inv f ` (f ` A) = A"
   by simp
 
-lemma bij_image_Collect_eq: "bij f \<Longrightarrow> f ` Collect P = {y. P (inv f y)}"
-  apply auto
-   apply (force simp add: bij_is_inj)
-  apply (blast intro: bij_is_surj [THEN surj_f_inv_f, symmetric])
-  done
+lemma bij_image_Collect_eq:
+  assumes "bij f"
+  shows "f ` Collect P = {y. P (inv f y)}"
+proof
+  show "f ` Collect P \<subseteq> {y. P (inv f y)}"
+    using assms by (force simp add: bij_is_inj)
+  show "{y. P (inv f y)} \<subseteq> f ` Collect P"
+    using assms by (blast intro: bij_is_surj [THEN surj_f_inv_f, symmetric])
+qed
 
-lemma bij_vimage_eq_inv_image: "bij f \<Longrightarrow> f -` A = inv f ` A"
-  apply (auto simp add: bij_is_surj [THEN surj_f_inv_f])
-  apply (blast intro: bij_is_inj [THEN inv_into_f_f, symmetric])
-  done
+lemma bij_vimage_eq_inv_image:
+  assumes "bij f"
+  shows "f -` A = inv f ` A"
+proof
+  show "f -` A \<subseteq> inv f ` A"
+    using assms by (blast intro: bij_is_inj [THEN inv_into_f_f, symmetric])
+  show "inv f ` A \<subseteq> f -` A"
+    using assms by (auto simp add: bij_is_surj [THEN surj_f_inv_f])
+qed
 
 lemma inv_fn_o_fn_is_id:
   fixes f::"'a \<Rightarrow> 'a"
@@ -338,11 +343,16 @@
   shows "inv (f^^n) = ((inv f)^^n)"
 proof -
   have "inv (f^^n) x = ((inv f)^^n) x" for x
-  apply (rule inv_into_f_eq, auto simp add: inj_fn[OF bij_is_inj[OF assms]])
-  using fn_o_inv_fn_is_id[OF assms, of n, THEN fun_cong] by (simp)
+  proof (rule inv_into_f_eq)
+    show "inj (f ^^ n)"
+      by (simp add: inj_fn[OF bij_is_inj [OF assms]])
+    show "(f ^^ n) ((inv f ^^ n) x) = x"
+      using fn_o_inv_fn_is_id[OF assms, THEN fun_cong] by force
+  qed auto
   then show ?thesis by auto
 qed
 
+
 lemma mono_inv:
   fixes f::"'a::linorder \<Rightarrow> 'b::linorder"
   assumes "mono f" "bij f"
@@ -746,13 +756,16 @@
   qed
   then have "N \<le> card (f N)" by simp
   also have "\<dots> \<le> card S" using S by (intro card_mono)
-  finally have "f (card S) = f N" using eq by auto
-  then show ?thesis
-    using eq inj [of N]
-    apply auto
-    apply (case_tac "n < N")
-     apply (auto simp: not_less)
-    done
+  finally have \<section>: "f (card S) = f N" using eq by auto
+  moreover have "\<Union> (range f) \<subseteq> f N"
+  proof clarify
+    fix x n
+    assume "x \<in> f n"
+    with eq inj [of N] show "x \<in> f N"
+      by (cases "n < N") (auto simp: not_less)
+  qed
+  ultimately show ?thesis
+    by auto
 qed
 
 
@@ -822,28 +835,13 @@
   case True
   with infinite have "\<not> finite (A - {a})" by auto
   with infinite_iff_countable_subset[of "A - {a}"]
-  obtain f :: "nat \<Rightarrow> 'a" where 1: "inj f" and 2: "f ` UNIV \<subseteq> A - {a}" by blast
+  obtain f :: "nat \<Rightarrow> 'a" where "inj f" and f: "f ` UNIV \<subseteq> A - {a}" by blast
   define g where "g n = (if n = 0 then a else f (Suc n))" for n
   define A' where "A' = g ` UNIV"
-  have *: "\<forall>y. f y \<noteq> a" using 2 by blast
+  have *: "\<forall>y. f y \<noteq> a" using f by blast
   have 3: "inj_on g UNIV \<and> g ` UNIV \<subseteq> A \<and> a \<in> g ` UNIV"
-    apply (auto simp add: True g_def [abs_def])
-     apply (unfold inj_on_def)
-     apply (intro ballI impI)
-     apply (case_tac "x = 0")
-      apply (auto simp add: 2)
-  proof -
-    fix y
-    assume "a = (if y = 0 then a else f (Suc y))"
-    then show "y = 0" by (cases "y = 0") (use * in auto)
-  next
-    fix x y
-    assume "f (Suc x) = (if y = 0 then a else f (Suc y))"
-    with 1 * show "x = y" by (cases "y = 0") (auto simp: inj_on_def)
-  next
-    fix n
-    from 2 show "f (Suc n) \<in> A" by blast
-  qed
+    using \<open>inj f\<close> f * unfolding inj_on_def g_def
+    by (auto simp add: True image_subset_iff)
   then have 4: "bij_betw g UNIV A' \<and> a \<in> A' \<and> A' \<subseteq> A"
     using inj_on_imp_bij_betw[of g] by (auto simp: A'_def)
   then have 5: "bij_betw (inv g) A' UNIV"
@@ -852,38 +850,14 @@
   have 6: "bij_betw g (UNIV - {n}) (A' - {a})"
     by (rule bij_betw_subset) (use 3 4 n in \<open>auto simp: image_set_diff A'_def\<close>)
   define v where "v m = (if m < n then m else Suc m)" for m
-  have 7: "bij_betw v UNIV (UNIV - {n})"
-  proof (unfold bij_betw_def inj_on_def, intro conjI, clarify)
-    fix m1 m2
-    assume "v m1 = v m2"
-    then show "m1 = m2"
-      apply (cases "m1 < n")
-       apply (cases "m2 < n")
-        apply (auto simp: inj_on_def v_def [abs_def])
-      apply (cases "m2 < n")
-       apply auto
-      done
-  next
-    show "v ` UNIV = UNIV - {n}"
-    proof (auto simp: v_def [abs_def])
-      fix m
-      assume "m \<noteq> n"
-      assume *: "m \<notin> Suc ` {m'. \<not> m' < n}"
-      have False if "n \<le> m"
-      proof -
-        from \<open>m \<noteq> n\<close> that have **: "Suc n \<le> m" by auto
-        from Suc_le_D [OF this] obtain m' where m': "m = Suc m'" ..
-        with ** have "n \<le> m'" by auto
-        with m' * show ?thesis by auto
-      qed
-      then show "m < n" by force
-    qed
-  qed
+  have "m < n \<or> m = n" if "\<And>k. k < n \<or> m \<noteq> Suc k" for m
+    using that [of "m-1"] by auto
+  then have 7: "bij_betw v UNIV (UNIV - {n})"
+    unfolding bij_betw_def inj_on_def v_def by auto
   define h' where "h' = g \<circ> v \<circ> (inv g)"
   with 5 6 7 have 8: "bij_betw h' A' (A' - {a})"
     by (auto simp add: bij_betw_trans)
   define h where "h b = (if b \<in> A' then h' b else b)" for b
-  then have "\<forall>b \<in> A'. h b = h' b" by simp
   with 8 have "bij_betw h  A' (A' - {a})"
     using bij_betw_cong[of A' h] by auto
   moreover
@@ -943,14 +917,14 @@
 lemma Sup_Inf: "\<Squnion> (Inf ` A) = \<Sqinter> (Sup ` {f ` A |f. \<forall>B\<in>A. f B \<in> B})"
 proof (rule antisym)
   show "\<Squnion> (Inf ` A) \<le> \<Sqinter> (Sup ` {f ` A |f. \<forall>B\<in>A. f B \<in> B})"
-    apply (rule Sup_least, rule INF_greatest)
-    using Inf_lower2 Sup_upper by auto
+    using Inf_lower2 Sup_upper
+    by (fastforce simp add: intro: Sup_least INF_greatest)
 next
   show "\<Sqinter> (Sup ` {f ` A |f. \<forall>B\<in>A. f B \<in> B}) \<le> \<Squnion> (Inf ` A)"
   proof (simp add:  Inf_Sup, rule SUP_least, simp, safe)
     fix f
     assume "\<forall>Y. (\<exists>f. Y = f ` A \<and> (\<forall>Y\<in>A. f Y \<in> Y)) \<longrightarrow> f Y \<in> Y"
-    from this have B: "\<And> F . (\<forall> Y \<in> A . F Y \<in> Y) \<Longrightarrow> \<exists> Z \<in> A . f (F ` A) = F Z"
+    then have B: "\<And> F . (\<forall> Y \<in> A . F Y \<in> Y) \<Longrightarrow> \<exists> Z \<in> A . f (F ` A) = F Z"
       by auto
     show "\<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> \<Squnion>(Inf ` A)"
     proof (cases "\<exists> Z \<in> A . \<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> Inf Z")
@@ -963,21 +937,20 @@
         by simp
     next
       case False
-      from this have X: "\<And> Z . Z \<in> A \<Longrightarrow> \<exists> x . x \<in> Z \<and> \<not> \<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> x"
+      then have X: "\<And> Z . Z \<in> A \<Longrightarrow> \<exists> x . x \<in> Z \<and> \<not> \<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> x"
         using Inf_greatest by blast
       define F where "F = (\<lambda> Z . SOME x . x \<in> Z \<and> \<not> \<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> x)"
-      have C: "\<And> Y . Y \<in> A \<Longrightarrow> F Y \<in> Y"
+      have C: "\<And>Y. Y \<in> A \<Longrightarrow> F Y \<in> Y"
         using X by (simp add: F_def, rule someI2_ex, auto)
-      have E: "\<And> Y . Y \<in> A \<Longrightarrow> \<not> \<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> F Y"
+      have E: "\<And>Y. Y \<in> A \<Longrightarrow> \<not> \<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> F Y"
         using X by (simp add: F_def, rule someI2_ex, auto)
       from C and B obtain  Z where D: "Z \<in> A " and Y: "f (F ` A) = F Z"
         by blast
       from E and D have W: "\<not> \<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> F Z"
         by simp
       have "\<Sqinter>(f ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) \<le> f (F ` A)"
-        apply (rule INF_lower)
-        using C by blast
-      from this and W and Y show ?thesis
+        using C by (blast intro: INF_lower)
+      with W Y show ?thesis
         by simp
     qed
   qed
@@ -985,15 +958,13 @@
   
 lemma dual_complete_distrib_lattice:
   "class.complete_distrib_lattice Sup Inf sup (\<ge>) (>) inf \<top> \<bottom>"
-  apply (rule class.complete_distrib_lattice.intro)
-   apply (fact dual_complete_lattice)
-  by (simp add: class.complete_distrib_lattice_axioms_def Sup_Inf)
+  by (simp add: class.complete_distrib_lattice.intro [OF dual_complete_lattice] 
+                class.complete_distrib_lattice_axioms_def Sup_Inf)
 
 lemma sup_Inf: "a \<squnion> \<Sqinter>B = \<Sqinter>((\<squnion>) a ` B)"
 proof (rule antisym)
   show "a \<squnion> \<Sqinter>B \<le> \<Sqinter>((\<squnion>) a ` B)"
-    apply (rule INF_greatest)
-    using Inf_lower sup.mono by fastforce
+    using Inf_lower sup.mono by (fastforce intro: INF_greatest)
 next
   have "\<Sqinter>((\<squnion>) a ` B) \<le> \<Sqinter>(Sup ` {{f {a}, f B} |f. f {a} = a \<and> f B \<in> B})"
     by (rule INF_greatest, auto simp add: INF_lower)
@@ -1034,8 +1005,7 @@
         have "(INF x\<in>{uu. \<exists>y. uu = {uu. \<exists>x. uu = P x y}}. f x) \<le> f {uu. \<exists>x. uu = P x y}"
           by (rule INF_lower, blast)
         also have "... \<le> P (SOME x. f {uu . \<exists>x. uu = P x y} = P x y) y"
-          apply (rule someI2_ex)
-          using A by auto
+          by (rule someI2_ex) (use A in auto)
         finally show "\<Sqinter>(f ` {uu. \<exists>y. uu = {uu. \<exists>x. uu = P x y}}) \<le>
           P (SOME x. f {uu. \<exists>x. uu = P x y} = P x y) y"
           by simp
@@ -1050,70 +1020,46 @@
 qed
 
 lemma INF_SUP_set: "(\<Sqinter>B\<in>A. \<Squnion>(g ` B)) = (\<Squnion>B\<in>{f ` A |f. \<forall>C\<in>A. f C \<in> C}. \<Sqinter>(g ` B))"
+                    (is "_ = (\<Squnion>B\<in>?F. _)")
 proof (rule antisym)
-  have "\<Sqinter> ((g \<circ> f) ` A) \<le> \<Squnion> (g ` B)" if "\<And>B. B \<in> A \<Longrightarrow> f B \<in> B" and "B \<in> A"
-    for f and B
+  have "\<Sqinter> ((g \<circ> f) ` A) \<le> \<Squnion> (g ` B)" if "\<And>B. B \<in> A \<Longrightarrow> f B \<in> B" "B \<in> A" for f B
     using that by (auto intro: SUP_upper2 INF_lower2)
-  then show "(\<Squnion>x\<in>{f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}. \<Sqinter>a\<in>x. g a) \<le> (\<Sqinter>x\<in>A. \<Squnion>a\<in>x. g a)"
+  then show "(\<Squnion>x\<in>?F. \<Sqinter>a\<in>x. g a) \<le> (\<Sqinter>x\<in>A. \<Squnion>a\<in>x. g a)"
     by (auto intro!: SUP_least INF_greatest simp add: image_comp)
 next
-  show "(\<Sqinter>x\<in>A. \<Squnion>a\<in>x. g a) \<le> (\<Squnion>x\<in>{f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}. \<Sqinter>a\<in>x. g a)"
+  show "(\<Sqinter>x\<in>A. \<Squnion>a\<in>x. g a) \<le> (\<Squnion>x\<in>?F. \<Sqinter>a\<in>x. g a)"
   proof (cases "{} \<in> A")
     case True
     then show ?thesis 
       by (rule INF_lower2) simp_all
   next
     case False
-    have *: "\<And>f B. B \<in> A \<Longrightarrow> f B \<in> B \<Longrightarrow>
-      (\<Sqinter>B. if B \<in> A then if f B \<in> B then g (f B) else \<bottom> else \<top>) \<le> g (f B)"
-      by (rule INF_lower2, auto)
-    have **: "\<And>f B. B \<in> A \<Longrightarrow> f B \<notin> B \<Longrightarrow>
-      (\<Sqinter>B. if B \<in> A then if f B \<in> B then g (f B) else \<bottom> else \<top>) \<le> g (SOME x. x \<in> B)"
-      by (rule INF_lower2, auto)
-    have ****: "\<And>f B. B \<in> A \<Longrightarrow>
-      (\<Sqinter>B. if B \<in> A then if f B \<in> B then g (f B) else \<bottom> else \<top>)
-        \<le> (if f B \<in> B then g (f B) else g (SOME x. x \<in> B))"
-      by (rule INF_lower2) auto
-    have ***: "\<And>x. (\<Sqinter>B. if B \<in> A then if x B \<in> B then g (x B) else \<bottom> else \<top>)
-        \<le> (\<Squnion>x\<in>{f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}. \<Sqinter>x\<in>x. g x)"
-    proof -
-      fix x
-      define F where "F = (\<lambda> (y::'b set) . if x y \<in> y then x y else (SOME x . x \<in>y))"
-      have B: "(\<forall>Y\<in>A. F Y \<in> Y)"
-        using False some_in_eq F_def by auto
-      have A: "F ` A \<in> {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}"
-        using B by blast
-      show "(\<Sqinter>xa. if xa \<in> A then if x xa \<in> xa then g (x xa) else \<bottom> else \<top>) \<le> (\<Squnion>x\<in>{f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}. \<Sqinter>x\<in>x. g x)"
-        using A apply (rule SUP_upper2)
-        apply (rule INF_greatest)
-        using * **
-        apply (auto simp add: F_def)
-        done
-    qed
-
     {fix x
-      have "(\<Sqinter>x\<in>A. \<Squnion>x\<in>x. g x) \<le> (\<Squnion>xa. if x \<in> A then if xa \<in> x then g xa else \<bottom> else \<top>)"
+      have "(\<Sqinter>x\<in>A. \<Squnion>x\<in>x. g x) \<le> (\<Squnion>u. if x \<in> A then if u \<in> x then g u else \<bottom> else \<top>)"
       proof (cases "x \<in> A")
         case True
         then show ?thesis
-          apply (rule INF_lower2)
-          apply (rule SUP_least)
-          apply (rule SUP_upper2)
-           apply auto
-          done
-      next
-        case False
-        then show ?thesis by simp
+          by (intro INF_lower2 SUP_least SUP_upper2) auto
+      qed auto
+    }
+    then have "(\<Sqinter>Y\<in>A. \<Squnion>a\<in>Y. g a) \<le> (\<Sqinter>Y. \<Squnion>y. if Y \<in> A then if y \<in> Y then g y else \<bottom> else \<top>)"
+      by (rule INF_greatest)
+    also have "... = (\<Squnion>x. \<Sqinter>Y. if Y \<in> A then if x Y \<in> Y then g (x Y) else \<bottom> else \<top>)"
+      by (simp only: INF_SUP)
+    also have "... \<le> (\<Squnion>x\<in>?F. \<Sqinter>a\<in>x. g a)"
+    proof (rule SUP_least)
+      show "(\<Sqinter>B. if B \<in> A then if x B \<in> B then g (x B) else \<bottom> else \<top>)
+               \<le> (\<Squnion>x\<in>?F. \<Sqinter>x\<in>x. g x)" for x
+      proof -
+        define G where "G \<equiv> \<lambda>Y. if x Y \<in> Y then x Y else (SOME x. x \<in>Y)"
+        have "\<forall>Y\<in>A. G Y \<in> Y"
+          using False some_in_eq G_def by auto
+        then have A: "G ` A \<in> ?F"
+          by blast
+        show "(\<Sqinter>Y. if Y \<in> A then if x Y \<in> Y then g (x Y) else \<bottom> else \<top>) \<le> (\<Squnion>x\<in>?F. \<Sqinter>x\<in>x. g x)"
+          by (fastforce simp: G_def intro: SUP_upper2 [OF A] INF_greatest INF_lower2)
       qed
-    }
-    from this have "(\<Sqinter>x\<in>A. \<Squnion>a\<in>x. g a) \<le> (\<Sqinter>x. \<Squnion>xa. if x \<in> A then if xa \<in> x then g xa else \<bottom> else \<top>)"
-      by (rule INF_greatest)
-    also have "... = (\<Squnion>x. \<Sqinter>xa. if xa \<in> A then if x xa \<in> xa then g (x xa) else \<bottom> else \<top>)"
-      by (simp only: INF_SUP)
-    also have "... \<le> (\<Squnion>x\<in>{f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}. \<Sqinter>a\<in>x. g a)"
-      apply (rule SUP_least)
-      using *** apply simp
-      done
+    qed
     finally show ?thesis by simp
   qed
 qed
@@ -1181,22 +1127,15 @@
 instance proof (standard, clarsimp)
   fix A :: "(('a set) set) set"
   fix x::'a
-  define F where "F = (\<lambda> Y . (SOME X . (Y \<in> A \<and> X \<in> Y \<and> x \<in> X)))"
-  assume A: "\<forall>xa\<in>A. \<exists>X\<in>xa. x \<in> X"
-    
-  from this have B: " (\<forall>xa \<in> F ` A. x \<in> xa)"
-    apply (safe, simp add: F_def)
-    by (rule someI2_ex, auto)
-
-  have C: "(\<forall>Y\<in>A. F Y \<in> Y)"
-    apply (simp  add: F_def, safe)
-    apply (rule someI2_ex)
-    using A by auto
-
-  have "(\<exists>f. F ` A  = f ` A \<and> (\<forall>Y\<in>A. f Y \<in> Y))"
-    using C by blast
-    
-  from B and this show "\<exists>X. (\<exists>f. X = f ` A \<and> (\<forall>Y\<in>A. f Y \<in> Y)) \<and> (\<forall>xa\<in>X. x \<in> xa)"
+  assume A: "\<forall>\<S>\<in>A. \<exists>X\<in>\<S>. x \<in> X"
+  define F where "F \<equiv> \<lambda>Y. SOME X. Y \<in> A \<and> X \<in> Y \<and> x \<in> X"
+  have "(\<forall>S \<in> F ` A. x \<in> S)"
+    using A unfolding F_def by (fastforce intro: someI2_ex)
+  moreover have "\<forall>Y\<in>A. F Y \<in> Y"
+    using A unfolding F_def by (fastforce intro: someI2_ex)
+  then have "\<exists>f. F ` A  = f ` A \<and> (\<forall>Y\<in>A. f Y \<in> Y)"
+    by blast
+  ultimately show "\<exists>X. (\<exists>f. X = f ` A \<and> (\<forall>Y\<in>A. f Y \<in> Y)) \<and> (\<forall>S\<in>X. x \<in> S)"
     by auto
 qed
 end
@@ -1212,85 +1151,56 @@
 
 context complete_linorder
 begin
-  
+
 subclass complete_distrib_lattice
 proof (standard, rule ccontr)
-  fix A
-  assume "\<not> \<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
-  then have C: "\<Sqinter>(Sup ` A) > \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
+  fix A :: "'a set set"
+  let ?F = "{f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}"
+  assume "\<not> \<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` ?F)"
+  then have C: "\<Sqinter>(Sup ` A) > \<Squnion>(Inf ` ?F)"
     by (simp add: not_le)
   show False
-    proof (cases "\<exists> z . \<Sqinter>(Sup ` A) > z \<and> z > \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})")
-      case True
-      from this obtain z where A: "z < \<Sqinter>(Sup ` A)" and X: "z > \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
-        by blast
-          
-      from A have "\<And> Y . Y \<in> A \<Longrightarrow> z < Sup Y"
-        by (simp add: less_INF_D)
-    
-      from this have B: "\<And> Y . Y \<in> A \<Longrightarrow> \<exists> k \<in>Y . z < k"
-        using local.less_Sup_iff by blast
-          
-      define F where "F = (\<lambda> Y . SOME k . k \<in> Y \<and> z < k)"
-        
-      have D: "\<And> Y . Y \<in> A \<Longrightarrow> z < F Y"
-        using B apply (simp add: F_def)
-        by (rule someI2_ex, auto)
+  proof (cases "\<exists> z . \<Sqinter>(Sup ` A) > z \<and> z > \<Squnion>(Inf ` ?F)")
+    case True
+    then obtain z where A: "z < \<Sqinter>(Sup ` A)" and X: "z > \<Squnion>(Inf ` ?F)"
+      by blast
+    then have B: "\<And>Y. Y \<in> A \<Longrightarrow> \<exists>k \<in>Y . z < k"
+      using local.less_Sup_iff by(force dest: less_INF_D)
+
+    define G where "G \<equiv> \<lambda>Y. SOME k . k \<in> Y \<and> z < k"
+    have E: "\<And>Y. Y \<in> A \<Longrightarrow> G Y \<in> Y"
+      using B unfolding G_def by (fastforce intro: someI2_ex)
+    have "z \<le> Inf (G ` A)"
+    proof (rule INF_greatest)
+      show  "\<And>Y. Y \<in> A \<Longrightarrow> z \<le> G Y"
+        using B unfolding G_def by (fastforce intro: someI2_ex)
+    qed
+    also have "... \<le> \<Squnion>(Inf ` ?F)"
+      by (rule SUP_upper) (use E in blast)
+    finally have "z \<le> \<Squnion>(Inf ` ?F)"
+      by simp
 
-    
-      have E: "\<And> Y . Y \<in> A \<Longrightarrow> F Y \<in> Y"
-        using B apply (simp add: F_def)
-        by (rule someI2_ex, auto)
-    
-      have "z \<le> Inf (F ` A)"
-        by (simp add: D local.INF_greatest local.order.strict_implies_order)
-    
-      also have "... \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
-        apply (rule SUP_upper, safe)
-        using E by blast
-      finally have "z \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
-        by simp
-          
-      from X and this show ?thesis
-        using local.not_less by blast
-    next
-      case False
-      from this have A: "\<And> z . \<Sqinter>(Sup ` A) \<le> z \<or> z \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
-        using local.le_less_linear by blast
-
-      from C have "\<And> Y . Y \<in> A \<Longrightarrow> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) < Sup Y"
-        by (simp add: less_INF_D)
-
-      from this have B: "\<And> Y . Y \<in> A \<Longrightarrow> \<exists> k \<in>Y . \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) < k"
-        using local.less_Sup_iff by blast
-          
-      define F where "F = (\<lambda> Y . SOME k . k \<in> Y \<and> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) < k)"
-
-      have D: "\<And> Y . Y \<in> A \<Longrightarrow> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}) < F Y"
-        using B apply (simp add: F_def)
-        by (rule someI2_ex, auto)
-    
-      have E: "\<And> Y . Y \<in> A \<Longrightarrow> F Y \<in> Y"
-        using B apply (simp add: F_def)
-        by (rule someI2_ex, auto)
-          
-      have "\<And> Y . Y \<in> A \<Longrightarrow> \<Sqinter>(Sup ` A) \<le> F Y"
-        using D False local.leI by blast
-         
-      from this have "\<Sqinter>(Sup ` A) \<le> Inf (F ` A)"
-        by (simp add: local.INF_greatest)
-          
-      also have "Inf (F ` A) \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
-        apply (rule SUP_upper, safe)
-        using E by blast
-
-      finally have "\<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
-        by simp
-        
-      from C and this show ?thesis
-        using not_less by blast
-    qed
+    with X show ?thesis
+      using local.not_less by blast
+  next
+    case False
+    have B: "\<And>Y. Y \<in> A \<Longrightarrow> \<exists> k \<in>Y . \<Squnion>(Inf ` ?F) < k"
+      using C local.less_Sup_iff by(force dest: less_INF_D)
+    define G where "G \<equiv> \<lambda> Y . SOME k . k \<in> Y \<and> \<Squnion>(Inf ` ?F) < k"
+    have E: "\<And>Y. Y \<in> A \<Longrightarrow> G Y \<in> Y"
+      using B unfolding G_def by (fastforce intro: someI2_ex)
+    have "\<And>Y. Y \<in> A \<Longrightarrow> \<Sqinter>(Sup ` A) \<le> G Y"
+      using B False local.leI unfolding G_def by (fastforce intro: someI2_ex)
+    then have "\<Sqinter>(Sup ` A) \<le> Inf (G ` A)"
+      by (simp add: local.INF_greatest)
+    also have "Inf (G ` A) \<le> \<Squnion>(Inf ` ?F)"
+      by (rule SUP_upper) (use E in blast)
+    finally have "\<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` ?F)"
+      by simp
+    with C show ?thesis
+      using not_less by blast
   qed
+qed
 end
 
 
--- a/src/HOL/Induct/Comb.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Induct/Comb.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -10,12 +10,10 @@
 begin
 
 text \<open>
-  Curiously, combinators do not include free variables.
-
+  Combinator terms do not have free variables.
   Example taken from @{cite camilleri92}.
 \<close>
 
-
 subsection \<open>Definitions\<close>
 
 text \<open>Datatype definition of combinators \<open>S\<close> and \<open>K\<close>.\<close>
@@ -29,36 +27,32 @@
   (multi-step) reductions, \<open>\<rightarrow>\<close>.
 \<close>
 
-inductive_set contract :: "(comb*comb) set"
-  and contract_rel1 :: "[comb,comb] \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sup>1" 50)
-where
-  "x \<rightarrow>\<^sup>1 y == (x,y) \<in> contract"
-| K:     "K\<bullet>x\<bullet>y \<rightarrow>\<^sup>1 x"
-| S:     "S\<bullet>x\<bullet>y\<bullet>z \<rightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
-| Ap1:   "x\<rightarrow>\<^sup>1y \<Longrightarrow> x\<bullet>z \<rightarrow>\<^sup>1 y\<bullet>z"
-| Ap2:   "x\<rightarrow>\<^sup>1y \<Longrightarrow> z\<bullet>x \<rightarrow>\<^sup>1 z\<bullet>y"
+inductive contract1 :: "[comb,comb] \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sup>1" 50)
+  where
+    K:     "K\<bullet>x\<bullet>y \<rightarrow>\<^sup>1 x"
+  | S:     "S\<bullet>x\<bullet>y\<bullet>z \<rightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
+  | Ap1:   "x \<rightarrow>\<^sup>1 y \<Longrightarrow> x\<bullet>z \<rightarrow>\<^sup>1 y\<bullet>z"
+  | Ap2:   "x \<rightarrow>\<^sup>1 y \<Longrightarrow> z\<bullet>x \<rightarrow>\<^sup>1 z\<bullet>y"
 
 abbreviation
-  contract_rel :: "[comb,comb] \<Rightarrow> bool"   (infixl "\<rightarrow>" 50) where
-  "x \<rightarrow> y == (x,y) \<in> contract\<^sup>*"
+  contract :: "[comb,comb] \<Rightarrow> bool"   (infixl "\<rightarrow>" 50) where
+  "contract \<equiv> contract1\<^sup>*\<^sup>*"
 
 text \<open>
   Inductive definition of parallel contractions, \<open>\<Rrightarrow>\<^sup>1\<close> and
   (multi-step) parallel reductions, \<open>\<Rrightarrow>\<close>.
 \<close>
 
-inductive_set parcontract :: "(comb*comb) set"
-  and parcontract_rel1 :: "[comb,comb] \<Rightarrow> bool"  (infixl "\<Rrightarrow>\<^sup>1" 50)
-where
-  "x \<Rrightarrow>\<^sup>1 y == (x,y) \<in> parcontract"
-| refl:  "x \<Rrightarrow>\<^sup>1 x"
-| K:     "K\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 x"
-| S:     "S\<bullet>x\<bullet>y\<bullet>z \<Rrightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
-| Ap:    "[| x\<Rrightarrow>\<^sup>1y;  z\<Rrightarrow>\<^sup>1w |] ==> x\<bullet>z \<Rrightarrow>\<^sup>1 y\<bullet>w"
+inductive parcontract1 :: "[comb,comb] \<Rightarrow> bool"  (infixl "\<Rrightarrow>\<^sup>1" 50)
+  where
+    refl:  "x \<Rrightarrow>\<^sup>1 x"
+  | K:     "K\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 x"
+  | S:     "S\<bullet>x\<bullet>y\<bullet>z \<Rrightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
+  | Ap:    "\<lbrakk>x \<Rrightarrow>\<^sup>1 y; z \<Rrightarrow>\<^sup>1 w\<rbrakk> \<Longrightarrow> x\<bullet>z \<Rrightarrow>\<^sup>1 y\<bullet>w"
 
 abbreviation
-  parcontract_rel :: "[comb,comb] \<Rightarrow> bool"   (infixl "\<Rrightarrow>" 50) where
-  "x \<Rrightarrow> y == (x,y) \<in> parcontract\<^sup>*"
+  parcontract :: "[comb,comb] \<Rightarrow> bool"   (infixl "\<Rrightarrow>" 50) where
+  "parcontract \<equiv> parcontract1\<^sup>*\<^sup>*"
 
 text \<open>
   Misc definitions.
@@ -66,37 +60,55 @@
 
 definition
   I :: comb where
-  "I = S\<bullet>K\<bullet>K"
+  "I \<equiv> S\<bullet>K\<bullet>K"
 
 definition
-  diamond   :: "('a * 'a)set \<Rightarrow> bool" where
+  diamond   :: "([comb,comb] \<Rightarrow> bool) \<Rightarrow> bool" where
     \<comment> \<open>confluence; Lambda/Commutation treats this more abstractly\<close>
-  "diamond(r) = (\<forall>x y. (x,y) \<in> r --> 
-                  (\<forall>y'. (x,y') \<in> r --> 
-                    (\<exists>z. (y,z) \<in> r & (y',z) \<in> r)))"
+  "diamond r \<equiv> \<forall>x y. r x y \<longrightarrow>
+                  (\<forall>y'. r x y' \<longrightarrow> 
+                    (\<exists>z. r y z \<and> r y' z))"
 
 
 subsection \<open>Reflexive/Transitive closure preserves Church-Rosser property\<close>
 
-text\<open>So does the Transitive closure, with a similar proof\<close>
+text\<open>Remark: So does the Transitive closure, with a similar proof\<close>
 
 text\<open>Strip lemma.  
    The induction hypothesis covers all but the last diamond of the strip.\<close>
-lemma diamond_strip_lemmaE [rule_format]: 
-    "[| diamond(r);  (x,y) \<in> r\<^sup>* |] ==>   
-          \<forall>y'. (x,y') \<in> r --> (\<exists>z. (y',z) \<in> r\<^sup>* & (y,z) \<in> r)"
-apply (unfold diamond_def)
-apply (erule rtrancl_induct)
-apply (meson rtrancl_refl)
-apply (meson rtrancl_trans r_into_rtrancl)
-done
+lemma strip_lemma [rule_format]: 
+  assumes "diamond r" and r: "r\<^sup>*\<^sup>* x y" "r x y'"
+  shows "\<exists>z. r\<^sup>*\<^sup>* y' z \<and> r y z"
+  using r
+proof (induction rule: rtranclp_induct)
+  case base
+  then show ?case
+    by blast
+next
+  case (step y z)
+  then show ?case
+    using \<open>diamond r\<close> unfolding diamond_def
+    by (metis rtranclp.rtrancl_into_rtrancl)
+qed
 
-lemma diamond_rtrancl: "diamond(r) \<Longrightarrow> diamond(r\<^sup>*)"
-apply (simp (no_asm_simp) add: diamond_def)
-apply (rule impI [THEN allI, THEN allI])
-apply (erule rtrancl_induct, blast)
-apply (meson rtrancl_trans r_into_rtrancl diamond_strip_lemmaE)
-done
+proposition diamond_rtrancl:
+  assumes "diamond r" 
+  shows "diamond(r\<^sup>*\<^sup>*)"
+  unfolding diamond_def
+proof (intro strip)
+  fix x y y'
+  assume "r\<^sup>*\<^sup>* x y" "r\<^sup>*\<^sup>* x y'"
+  then show "\<exists>z. r\<^sup>*\<^sup>* y z \<and> r\<^sup>*\<^sup>* y' z"
+  proof (induction rule: rtranclp_induct)
+    case base
+    then show ?case
+      by blast
+  next
+    case (step y z)
+    then show ?case
+      by (meson assms strip_lemma rtranclp.rtrancl_into_rtrancl)
+  qed
+qed
 
 
 subsection \<open>Non-contraction results\<close>
@@ -104,33 +116,29 @@
 text \<open>Derive a case for each combinator constructor.\<close>
 
 inductive_cases
-      K_contractE [elim!]: "K \<rightarrow>\<^sup>1 r"
+  K_contractE [elim!]: "K \<rightarrow>\<^sup>1 r"
   and S_contractE [elim!]: "S \<rightarrow>\<^sup>1 r"
   and Ap_contractE [elim!]: "p\<bullet>q \<rightarrow>\<^sup>1 r"
 
-declare contract.K [intro!] contract.S [intro!]
-declare contract.Ap1 [intro] contract.Ap2 [intro]
+declare contract1.K [intro!] contract1.S [intro!]
+declare contract1.Ap1 [intro] contract1.Ap2 [intro]
 
-lemma I_contract_E [elim!]: "I \<rightarrow>\<^sup>1 z \<Longrightarrow> P"
-by (unfold I_def, blast)
+lemma I_contract_E [iff]: "\<not> I \<rightarrow>\<^sup>1 z"
+  unfolding I_def by blast
 
-lemma K1_contractD [elim!]: "K\<bullet>x \<rightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' & x \<rightarrow>\<^sup>1 x')"
-by blast
+lemma K1_contractD [elim!]: "K\<bullet>x \<rightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' \<and> x \<rightarrow>\<^sup>1 x')"
+  by blast
 
 lemma Ap_reduce1 [intro]: "x \<rightarrow> y \<Longrightarrow> x\<bullet>z \<rightarrow> y\<bullet>z"
-apply (erule rtrancl_induct)
-apply (blast intro: rtrancl_trans)+
-done
+  by (induction rule: rtranclp_induct; blast intro: rtranclp_trans)
 
 lemma Ap_reduce2 [intro]: "x \<rightarrow> y \<Longrightarrow> z\<bullet>x \<rightarrow> z\<bullet>y"
-apply (erule rtrancl_induct)
-apply (blast intro: rtrancl_trans)+
-done
+  by (induction rule: rtranclp_induct; blast intro: rtranclp_trans)
 
 text \<open>Counterexample to the diamond property for \<^term>\<open>x \<rightarrow>\<^sup>1 y\<close>\<close>
 
-lemma not_diamond_contract: "~ diamond(contract)"
-by (unfold diamond_def, metis S_contractE contract.K) 
+lemma not_diamond_contract: "\<not> diamond(contract1)"
+  unfolding diamond_def by (metis S_contractE contract1.K) 
 
 
 subsection \<open>Results about Parallel Contraction\<close>
@@ -142,55 +150,52 @@
   and S_parcontractE [elim!]: "S \<Rrightarrow>\<^sup>1 r"
   and Ap_parcontractE [elim!]: "p\<bullet>q \<Rrightarrow>\<^sup>1 r"
 
-declare parcontract.intros [intro]
-
-(*** Basic properties of parallel contraction ***)
+declare parcontract1.intros [intro]
 
 subsection \<open>Basic properties of parallel contraction\<close>
+text\<open>The rules below are not essential but make proofs much faster\<close>
 
-lemma K1_parcontractD [dest!]: "K\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' & x \<Rrightarrow>\<^sup>1 x')"
-by blast
+lemma K1_parcontractD [dest!]: "K\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' \<and> x \<Rrightarrow>\<^sup>1 x')"
+  by blast
 
-lemma S1_parcontractD [dest!]: "S\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = S\<bullet>x' & x \<Rrightarrow>\<^sup>1 x')"
-by blast
+lemma S1_parcontractD [dest!]: "S\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = S\<bullet>x' \<and> x \<Rrightarrow>\<^sup>1 x')"
+  by blast
 
-lemma S2_parcontractD [dest!]:
-     "S\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x' y'. z = S\<bullet>x'\<bullet>y' & x \<Rrightarrow>\<^sup>1 x' & y \<Rrightarrow>\<^sup>1 y')"
-by blast
-
-text\<open>The rules above are not essential but make proofs much faster\<close>
+lemma S2_parcontractD [dest!]: "S\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x' y'. z = S\<bullet>x'\<bullet>y' \<and> x \<Rrightarrow>\<^sup>1 x' \<and> y \<Rrightarrow>\<^sup>1 y')"
+  by blast
 
 text\<open>Church-Rosser property for parallel contraction\<close>
-lemma diamond_parcontract: "diamond parcontract"
-apply (unfold diamond_def)
-apply (rule impI [THEN allI, THEN allI])
-apply (erule parcontract.induct, fast+)
-done
+proposition diamond_parcontract: "diamond parcontract1"
+proof -
+  have "(\<exists>z. w \<Rrightarrow>\<^sup>1 z \<and> y' \<Rrightarrow>\<^sup>1 z)" if "y \<Rrightarrow>\<^sup>1 w" "y \<Rrightarrow>\<^sup>1 y'" for w y y'
+    using that by (induction arbitrary: y' rule: parcontract1.induct) fast+
+  then show ?thesis
+    by (auto simp: diamond_def)
+qed
 
-text \<open>
-  \<^medskip>
-  Equivalence of \<^prop>\<open>p \<rightarrow> q\<close> and \<^prop>\<open>p \<Rrightarrow> q\<close>.
-\<close>
+subsection \<open>Equivalence of \<^prop>\<open>p \<rightarrow> q\<close> and \<^prop>\<open>p \<Rrightarrow> q\<close>.\<close>
 
-lemma contract_subset_parcontract: "contract \<subseteq> parcontract"
-by (auto, erule contract.induct, blast+)
+lemma contract_imp_parcontract: "x \<rightarrow>\<^sup>1 y \<Longrightarrow> x \<Rrightarrow>\<^sup>1 y"
+  by (induction rule: contract1.induct; blast)
 
 text\<open>Reductions: simply throw together reflexivity, transitivity and
   the one-step reductions\<close>
 
-declare r_into_rtrancl [intro]  rtrancl_trans [intro]
-
-(*Example only: not used*)
-lemma reduce_I: "I\<bullet>x \<rightarrow> x"
-by (unfold I_def, blast)
+proposition reduce_I: "I\<bullet>x \<rightarrow> x"
+  unfolding I_def
+  by (meson contract1.K contract1.S r_into_rtranclp rtranclp.rtrancl_into_rtrancl)
 
-lemma parcontract_subset_reduce: "parcontract \<subseteq> contract\<^sup>*"
-by (auto, erule parcontract.induct, blast+)
+lemma parcontract_imp_reduce: "x \<Rrightarrow>\<^sup>1 y \<Longrightarrow> x \<rightarrow> y"
+proof (induction rule: parcontract1.induct)
+  case (Ap x y z w)
+  then show ?case
+    by (meson Ap_reduce1 Ap_reduce2 rtranclp_trans)
+qed auto
 
-lemma reduce_eq_parreduce: "contract\<^sup>* = parcontract\<^sup>*"
-by (metis contract_subset_parcontract parcontract_subset_reduce rtrancl_subset)
+lemma reduce_eq_parreduce: "x \<rightarrow> y  \<longleftrightarrow>  x \<Rrightarrow> y"
+  by (metis contract_imp_parcontract parcontract_imp_reduce predicate2I rtranclp_subset)
 
-theorem diamond_reduce: "diamond(contract\<^sup>*)"
-by (simp add: reduce_eq_parreduce diamond_rtrancl diamond_parcontract)
+theorem diamond_reduce: "diamond(contract)"
+  using diamond_parcontract diamond_rtrancl reduce_eq_parreduce by presburger
 
 end
--- a/src/HOL/Int.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Int.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -136,21 +136,23 @@
     by (cases "k = 0") (simp_all add: distrib_right add_strict_mono)
 qed
 
-lemma zero_le_imp_eq_int: "0 \<le> k \<Longrightarrow> \<exists>n. k = int n"
-  for k :: int
-  apply transfer
-  apply clarsimp
-  apply (rule_tac x="a - b" in exI)
-  apply simp
-  done
+lemma zero_le_imp_eq_int:
+  assumes "k \<ge> (0::int)" shows "\<exists>n. k = int n"
+proof -
+  have "b \<le> a \<Longrightarrow> \<exists>n::nat. a = n + b" for a b
+    by (rule_tac x="a - b" in exI) simp
+  with assms show ?thesis
+    by transfer auto
+qed
 
-lemma zero_less_imp_eq_int: "0 < k \<Longrightarrow> \<exists>n>0. k = int n"
-  for k :: int
-  apply transfer
-  apply clarsimp
-  apply (rule_tac x="a - b" in exI)
-  apply simp
-  done
+lemma zero_less_imp_eq_int:
+  assumes "k > (0::int)" shows "\<exists>n>0. k = int n"
+proof -
+  have "b < a \<Longrightarrow> \<exists>n::nat. n>0 \<and> a = n + b" for a b
+    by (rule_tac x="a - b" in exI) simp
+  with assms show ?thesis
+    by transfer auto
+qed
 
 lemma zmult_zless_mono2: "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"
   for i j k :: int
@@ -185,12 +187,12 @@
 
 lemma zless_iff_Suc_zadd: "w < z \<longleftrightarrow> (\<exists>n. z = w + int (Suc n))"
   for w z :: int
-  apply transfer
-  apply auto
-  apply (rename_tac a b c d)
-  apply (rule_tac x="c+b - Suc(a+d)" in exI)
-  apply arith
-  done
+proof -
+  have "\<And>a b c d. a + d < c + b \<Longrightarrow> \<exists>n. c + b = Suc (a + n + d)"
+    by (rule_tac x="c+b - Suc(a+d)" in exI) arith
+  then show ?thesis
+    by transfer auto
+qed
 
 lemma zabs_less_one_iff [simp]: "\<bar>z\<bar> < 1 \<longleftrightarrow> z = 0" (is "?lhs \<longleftrightarrow> ?rhs")
   for z :: int
@@ -471,16 +473,16 @@
 qed
 
 instance int :: no_top
-  apply standard
-  apply (rule_tac x="x + 1" in exI)
-  apply simp
-  done
+proof
+  show "\<And>x::int. \<exists>y. x < y"
+    by (rule_tac x="x + 1" in exI) simp
+qed
 
 instance int :: no_bot
-  apply standard
-  apply (rule_tac x="x - 1" in exI)
-  apply simp
-  done
+proof
+  show "\<And>x::int. \<exists>y. y < x"
+    by (rule_tac x="x - 1" in exI) simp
+qed
 
 
 subsection \<open>Magnitude of an Integer, as a Natural Number: \<open>nat\<close>\<close>
@@ -732,12 +734,14 @@
   for a :: "'a::linordered_idom"
   by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
 
-lemma negD: "x < 0 \<Longrightarrow> \<exists>n. x = - (int (Suc n))"
-  apply transfer
-  apply clarsimp
-  apply (rule_tac x="b - Suc a" in exI)
-  apply arith
-  done
+lemma negD:
+  assumes "x < 0" shows "\<exists>n. x = - (int (Suc n))"
+proof -
+  have "\<And>a b. a < b \<Longrightarrow> \<exists>n. Suc (a + n) = b"
+    by (rule_tac x="b - Suc a" in exI) arith
+  with assms show ?thesis
+    by transfer auto
+qed
 
 
 subsection \<open>Cases and induction\<close>
@@ -754,13 +758,17 @@
 
 text \<open>This is the default, with a negative case.\<close>
 lemma int_cases [case_names nonneg neg, cases type: int]:
-  "(\<And>n. z = int n \<Longrightarrow> P) \<Longrightarrow> (\<And>n. z = - (int (Suc n)) \<Longrightarrow> P) \<Longrightarrow> P"
-  apply (cases "z < 0")
-   apply (blast dest!: negD)
-  apply (simp add: linorder_not_less del: of_nat_Suc)
-  apply auto
-  apply (blast dest: nat_0_le [THEN sym])
-  done
+  assumes pos: "\<And>n. z = int n \<Longrightarrow> P" and neg: "\<And>n. z = - (int (Suc n)) \<Longrightarrow> P"
+  shows P
+proof (cases "z < 0")
+  case True
+  with neg show ?thesis
+    by (blast dest!: negD)
+next
+  case False
+  with pos show ?thesis
+    by (force simp add: linorder_not_less dest: nat_0_le [THEN sym])
+qed
 
 lemma int_cases3 [case_names zero pos neg]:
   fixes k :: int
@@ -891,31 +899,19 @@
   by (subst of_nat_numeral [symmetric], rule Ints_of_nat)
 
 lemma Ints_add [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a + b \<in> \<int>"
-  apply (auto simp add: Ints_def)
-  apply (rule range_eqI)
-  apply (rule of_int_add [symmetric])
-  done
+  by (force simp add: Ints_def simp flip: of_int_add intro: range_eqI)
 
 lemma Ints_minus [simp]: "a \<in> \<int> \<Longrightarrow> -a \<in> \<int>"
-  apply (auto simp add: Ints_def)
-  apply (rule range_eqI)
-  apply (rule of_int_minus [symmetric])
-  done
+  by (force simp add: Ints_def simp flip: of_int_minus intro: range_eqI)
 
 lemma minus_in_Ints_iff: "-x \<in> \<int> \<longleftrightarrow> x \<in> \<int>"
   using Ints_minus[of x] Ints_minus[of "-x"] by auto
 
 lemma Ints_diff [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a - b \<in> \<int>"
-  apply (auto simp add: Ints_def)
-  apply (rule range_eqI)
-  apply (rule of_int_diff [symmetric])
-  done
+  by (force simp add: Ints_def simp flip: of_int_diff intro: range_eqI)
 
 lemma Ints_mult [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a * b \<in> \<int>"
-  apply (auto simp add: Ints_def)
-  apply (rule range_eqI)
-  apply (rule of_int_mult [symmetric])
-  done
+  by (force simp add: Ints_def simp flip: of_int_mult intro: range_eqI)
 
 lemma Ints_power [simp]: "a \<in> \<int> \<Longrightarrow> a ^ n \<in> \<int>"
   by (induct n) simp_all
@@ -1234,12 +1230,15 @@
          of_nat_nat [OF assms] of_nat_nat [OF ge_0], simp)
 qed
 
-lemma nat_mult_distrib_neg: "z \<le> 0 \<Longrightarrow> nat (z * z') = nat (- z) * nat (- z')"
-  for z z' :: int
-  apply (rule trans)
-   apply (rule_tac [2] nat_mult_distrib)
-   apply auto
-  done
+lemma nat_mult_distrib_neg:
+  assumes "z \<le> (0::int)" shows "nat (z * z') = nat (- z) * nat (- z')" (is "?L = ?R")
+proof -
+  have "?L = nat (- z * - z')"
+    using assms by auto
+  also have "... = ?R"
+    by (rule nat_mult_distrib) (use assms in auto)
+  finally show ?thesis .
+qed
 
 lemma nat_abs_mult_distrib: "nat \<bar>w * z\<bar> = nat \<bar>w\<bar> * nat \<bar>z\<bar>"
   by (cases "z = 0 \<or> w = 0")
@@ -1332,16 +1331,14 @@
 (* `set:int': dummy construction *)
 theorem int_gr_induct [case_names base step, induct set: int]:
   fixes i k :: int
-  assumes gr: "k < i"
-    and base: "P (k + 1)"
-    and step: "\<And>i. k < i \<Longrightarrow> P i \<Longrightarrow> P (i + 1)"
+  assumes "k < i" "P (k + 1)" "\<And>i. k < i \<Longrightarrow> P i \<Longrightarrow> P (i + 1)"
   shows "P i"
-  apply (rule int_ge_induct[of "k + 1"])
-  using gr apply arith
-   apply (rule base)
-  apply (rule step)
-   apply simp_all
-  done
+proof -
+  have "k+1 \<le> i"
+    using assms by auto
+  then show ?thesis
+    by (induction i rule: int_ge_induct) (auto simp: assms)
+qed
 
 theorem int_le_induct [consumes 1, case_names base step]:
   fixes i k :: int
@@ -1367,16 +1364,14 @@
 
 theorem int_less_induct [consumes 1, case_names base step]:
   fixes i k :: int
-  assumes less: "i < k"
-    and base: "P (k - 1)"
-    and step: "\<And>i. i < k \<Longrightarrow> P i \<Longrightarrow> P (i - 1)"
+  assumes "i < k" "P (k - 1)" "\<And>i. i < k \<Longrightarrow> P i \<Longrightarrow> P (i - 1)"
   shows "P i"
-  apply (rule int_le_induct[of _ "k - 1"])
-  using less apply arith
-   apply (rule base)
-  apply (rule step)
-   apply simp_all
-  done
+proof -
+  have "i \<le> k-1"
+    using assms by auto
+  then show ?thesis
+    by (induction i rule: int_le_induct) (auto simp: assms)
+qed
 
 theorem int_induct [case_names base step1 step2]:
   fixes k :: int
@@ -1401,26 +1396,30 @@
 
 subsection \<open>Intermediate value theorems\<close>
 
+lemma nat_ivt_aux: 
+  "\<lbrakk>\<forall>i<n. \<bar>f (Suc i) - f i\<bar> \<le> 1; f 0 \<le> k; k \<le> f n\<rbrakk> \<Longrightarrow> \<exists>i \<le> n. f i = k"
+  for m n :: nat and k :: int
+proof (induct n)
+  case (Suc n)
+  show ?case
+  proof (cases "k = f (Suc n)")
+    case False
+    with Suc have "k \<le> f n"
+      by auto
+    with Suc show ?thesis
+      by (auto simp add: abs_if split: if_split_asm intro: le_SucI)
+  qed (use Suc in auto)
+qed auto
+
 lemma nat_intermed_int_val:
-  "\<exists>i. m \<le> i \<and> i \<le> n \<and> f i = k"
-  if "\<forall>i. m \<le> i \<and> i < n \<longrightarrow> \<bar>f (Suc i) - f i\<bar> \<le> 1"
-    "m \<le> n" "f m \<le> k" "k \<le> f n"
-  for m n :: nat and k :: int
+  fixes m n :: nat and k :: int
+  assumes "\<forall>i. m \<le> i \<and> i < n \<longrightarrow> \<bar>f (Suc i) - f i\<bar> \<le> 1" "m \<le> n" "f m \<le> k" "k \<le> f n"
+  shows "\<exists>i. m \<le> i \<and> i \<le> n \<and> f i = k"
 proof -
-  have "(\<forall>i<n. \<bar>f (Suc i) - f i\<bar> \<le> 1) \<Longrightarrow> f 0 \<le> k \<Longrightarrow> k \<le> f n
-    \<Longrightarrow> (\<exists>i \<le> n. f i = k)"
-  for n :: nat and f
-    apply (induct n)
-     apply auto
-    apply (erule_tac x = n in allE)
-    apply (case_tac "k = f (Suc n)")
-     apply (auto simp add: abs_if split: if_split_asm intro: le_SucI)
-    done
-  from this [of "n - m" "f \<circ> plus m"] that show ?thesis
-    apply auto
-    apply (rule_tac x = "m + i" in exI)
-    apply auto
-    done
+  obtain i where "i \<le> n - m" "k = f (m + i)"
+    using nat_ivt_aux [of "n - m" "f \<circ> plus m" k] assms by auto
+  with assms show ?thesis
+    by (rule_tac x = "m + i" in exI) auto
 qed
 
 lemma nat0_intermed_int_val:
@@ -1465,14 +1464,12 @@
     by (auto dest: pos_zmult_eq_1_iff_lemma)
 qed
 
-lemma zmult_eq_1_iff: "m * n = 1 \<longleftrightarrow> (m = 1 \<and> n = 1) \<or> (m = - 1 \<and> n = - 1)"
+lemma zmult_eq_1_iff: "m * n = 1 \<longleftrightarrow> (m = 1 \<and> n = 1) \<or> (m = - 1 \<and> n = - 1)" (is "?L = ?R")
   for m n :: int
-  apply (rule iffI)
-   apply (frule pos_zmult_eq_1_iff_lemma)
-   apply (simp add: mult.commute [of m])
-   apply (frule pos_zmult_eq_1_iff_lemma)
-   apply auto
-  done
+proof
+  assume L: ?L show ?R
+    using pos_zmult_eq_1_iff_lemma [OF L] L by force
+qed auto
 
 lemma infinite_UNIV_int [simp]: "\<not> finite (UNIV::int set)"
 proof
@@ -1677,13 +1674,12 @@
   using nat_less_eq_zless[of a "numeral x ^ n"]
   by (cases "a < 0") (auto simp: nat_power_eq less_le_trans[where y=0])
 
-lemma zdvd_imp_le: "z dvd n \<Longrightarrow> 0 < n \<Longrightarrow> z \<le> n"
-  for n z :: int
-  apply (cases n)
-  apply auto
-  apply (cases z)
-   apply (auto simp add: dvd_imp_le)
-  done
+lemma zdvd_imp_le: "z \<le> n" if "z dvd n" "0 < n" for n z :: int
+proof (cases n)
+  case (nonneg n)
+  show ?thesis
+    by (cases z) (use nonneg dvd_imp_le that in auto)
+qed (use that in auto)
 
 lemma zdvd_period:
   fixes a d :: int
--- a/src/HOL/List.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/List.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -853,7 +853,7 @@
   "(Suc n \<le> length xs) = (\<exists>x ys. xs = x # ys \<and> n \<le> length ys)"
 by (metis Suc_le_D[of n] Suc_le_mono[of n] Suc_length_conv[of _ xs])
 
-lemma impossible_Cons: "length xs \<le> length ys ==> xs = x # ys = False"
+lemma impossible_Cons: "length xs \<le> length ys \<Longrightarrow> xs = x # ys = False"
 by (induct xs) auto
 
 lemma list_induct2 [consumes 1, case_names Nil Cons]:
@@ -972,7 +972,7 @@
 
 lemma append_eq_append_conv [simp]:
   "length xs = length ys \<or> length us = length vs
-  ==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
+  \<Longrightarrow> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
   by (induct xs arbitrary: ys; case_tac ys; force)
 
 lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) =
@@ -998,19 +998,19 @@
 lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
 using append_same_eq [of "[]"] by auto
 
-lemma hd_Cons_tl: "xs \<noteq> [] ==> hd xs # tl xs = xs"
+lemma hd_Cons_tl: "xs \<noteq> [] \<Longrightarrow> hd xs # tl xs = xs"
   by (fact list.collapse)
 
 lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
 by (induct xs) auto
 
-lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs"
+lemma hd_append2 [simp]: "xs \<noteq> [] \<Longrightarrow> hd (xs @ ys) = hd xs"
 by (simp add: hd_append split: list.split)
 
 lemma tl_append: "tl (xs @ ys) = (case xs of [] \<Rightarrow> tl ys | z#zs \<Rightarrow> zs @ ys)"
 by (simp split: list.split)
 
-lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys"
+lemma tl_append2 [simp]: "xs \<noteq> [] \<Longrightarrow> tl (xs @ ys) = tl xs @ ys"
 by (simp add: tl_append split: list.split)
 
 
@@ -1031,16 +1031,14 @@
 
 text \<open>Trivial rules for solving \<open>@\<close>-equations automatically.\<close>
 
-lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys"
-by simp
-
-lemma Cons_eq_appendI:
-"[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs"
-by (drule sym) simp
-
-lemma append_eq_appendI:
-"[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us"
-by (drule sym) simp
+lemma eq_Nil_appendI: "xs = ys \<Longrightarrow> xs = [] @ ys"
+  by simp
+
+lemma Cons_eq_appendI: "\<lbrakk>x # xs1 = ys; xs = xs1 @ zs\<rbrakk> \<Longrightarrow> x # xs = ys @ zs"
+  by auto
+
+lemma append_eq_appendI: "\<lbrakk>xs @ xs1 = zs; ys = xs1 @ us\<rbrakk> \<Longrightarrow> xs @ ys = zs @ us"
+  by auto
 
 
 text \<open>
@@ -1100,7 +1098,7 @@
 lemma map_tl: "map f (tl xs) = tl (map f xs)"
 by (cases xs) simp_all
 
-lemma map_ext: "(\<And>x. x \<in> set xs \<longrightarrow> f x = g x) ==> map f xs = map g xs"
+lemma map_ext: "(\<And>x. x \<in> set xs \<longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g xs"
 by (induct xs) simp_all
 
 lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)"
@@ -1175,16 +1173,16 @@
 by(blast dest:map_inj_on)
 
 lemma map_injective:
-  "map f xs = map f ys ==> inj f ==> xs = ys"
+  "map f xs = map f ys \<Longrightarrow> inj f \<Longrightarrow> xs = ys"
 by (induct ys arbitrary: xs) (auto dest!:injD)
 
 lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
 by(blast dest:map_injective)
 
-lemma inj_mapI: "inj f ==> inj (map f)"
+lemma inj_mapI: "inj f \<Longrightarrow> inj (map f)"
 by (iprover dest: map_injective injD intro: inj_onI)
 
-lemma inj_mapD: "inj (map f) ==> inj f"
+lemma inj_mapD: "inj (map f) \<Longrightarrow> inj f"
   by (metis (no_types, hide_lams) injI list.inject list.simps(9) the_inv_f_f)
 
 lemma inj_map[iff]: "inj (map f) = inj f"
@@ -1260,13 +1258,16 @@
 by(simp add:inj_on_def)
 
 lemma rev_induct [case_names Nil snoc]:
-  "[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs"
-apply(simplesubst rev_rev_ident[symmetric])
-apply(rule_tac list = "rev xs" in list.induct, simp_all)
-done
+  assumes "P []" and "\<And>x xs. P xs \<Longrightarrow> P (xs @ [x])"
+  shows "P xs"
+proof -
+  have "P (rev (rev xs))"
+    by (rule_tac list = "rev xs" in list.induct, simp_all add: assms)
+  then show ?thesis by simp
+qed
 
 lemma rev_exhaust [case_names Nil snoc]:
-  "(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P"
+  "(xs = [] \<Longrightarrow> P) \<Longrightarrow>(\<And>ys y. xs = ys @ [y] \<Longrightarrow> P) \<Longrightarrow> P"
 by (induct xs rule: rev_induct) auto
 
 lemmas rev_cases = rev_exhaust
@@ -1475,10 +1476,10 @@
   "length(filter P xs) + length(filter (\<lambda>x. \<not>P x) xs) = length xs"
 by(induct xs) simp_all
 
-lemma filter_True [simp]: "\<forall>x \<in> set xs. P x ==> filter P xs = xs"
+lemma filter_True [simp]: "\<forall>x \<in> set xs. P x \<Longrightarrow> filter P xs = xs"
 by (induct xs) auto
 
-lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x ==> filter P xs = []"
+lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x \<Longrightarrow> filter P xs = []"
 by (induct xs) auto
 
 lemma filter_empty_conv: "(filter P xs = []) = (\<forall>x\<in>set xs. \<not> P x)"
@@ -1664,13 +1665,13 @@
 lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))"
   by (induct xs) auto
 
-lemma concat_eq_concat_iff: "\<forall>(x, y) \<in> set (zip xs ys). length x = length y ==> length xs = length ys ==> (concat xs = concat ys) = (xs = ys)"
+lemma concat_eq_concat_iff: "\<forall>(x, y) \<in> set (zip xs ys). length x = length y \<Longrightarrow> length xs = length ys \<Longrightarrow> (concat xs = concat ys) = (xs = ys)"
 proof (induct xs arbitrary: ys)
   case (Cons x xs ys)
   thus ?case by (cases ys) auto
 qed (auto)
 
-lemma concat_injective: "concat xs = concat ys ==> length xs = length ys ==> \<forall>(x, y) \<in> set (zip xs ys). length x = length y ==> xs = ys"
+lemma concat_injective: "concat xs = concat ys \<Longrightarrow> length xs = length ys \<Longrightarrow> \<forall>(x, y) \<in> set (zip xs ys). length x = length y \<Longrightarrow> xs = ys"
   by (simp add: concat_eq_concat_iff)
 
 lemma concat_eq_appendD:
@@ -1725,7 +1726,7 @@
 lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n"
   by (induct xs) auto
 
-lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)"
+lemma nth_map [simp]: "n < length xs \<Longrightarrow> (map f xs)!n = f(xs!n)"
 proof (induct xs arbitrary: n)
   case (Cons x xs)
   then show ?case
@@ -1856,13 +1857,13 @@
   by (induct xs arbitrary: i) (auto split: nat.split)
 
 lemma nth_list_update:
-  "i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)"
+  "i < length xs\<Longrightarrow> (xs[i:=x])!j = (if i = j then x else xs!j)"
   by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
 
-lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x"
+lemma nth_list_update_eq [simp]: "i < length xs \<Longrightarrow> (xs[i:=x])!i = x"
   by (simp add: nth_list_update)
 
-lemma nth_list_update_neq [simp]: "i \<noteq> j ==> xs[i:=x]!j = xs!j"
+lemma nth_list_update_neq [simp]: "i \<noteq> j \<Longrightarrow> xs[i:=x]!j = xs!j"
   by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
 
 lemma list_update_id[simp]: "xs[i := xs!i] = xs"
@@ -1879,7 +1880,7 @@
   by (simp only: length_0_conv[symmetric] length_list_update)
 
 lemma list_update_same_conv:
-  "i < length xs ==> (xs[i := x] = xs) = (xs!i = x)"
+  "i < length xs \<Longrightarrow> (xs[i := x] = xs) = (xs!i = x)"
   by (induct xs arbitrary: i) (auto split: nat.split)
 
 lemma list_update_append1:
@@ -2105,10 +2106,10 @@
 lemma length_drop [simp]: "length (drop n xs) = (length xs - n)"
   by (induct n arbitrary: xs) (auto, case_tac xs, auto)
 
-lemma take_all [simp]: "length xs \<le> n ==> take n xs = xs"
+lemma take_all [simp]: "length xs \<le> n \<Longrightarrow> take n xs = xs"
   by (induct n arbitrary: xs) (auto, case_tac xs, auto)
 
-lemma drop_all [simp]: "length xs \<le> n ==> drop n xs = []"
+lemma drop_all [simp]: "length xs \<le> n \<Longrightarrow> drop n xs = []"
   by (induct n arbitrary: xs) (auto, case_tac xs, auto)
 
 lemma take_append [simp]:
@@ -2206,7 +2207,7 @@
 lemma take_rev: "take n (rev xs) = rev (drop (length xs - n) xs)"
   by (cases "length xs < n") (auto simp: rev_drop)
 
-lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i"
+lemma nth_take [simp]: "i < n \<Longrightarrow> (take n xs)!i = xs!i"
 proof (induct xs arbitrary: i n)
   case Nil
   then show ?case by simp
@@ -2216,7 +2217,7 @@
 qed
 
 lemma nth_drop [simp]:
-  "n \<le> length xs ==> (drop n xs)!i = xs!(n + i)"
+  "n \<le> length xs \<Longrightarrow> (drop n xs)!i = xs!(n + i)"
 proof (induct n arbitrary: xs)
   case 0
   then show ?case by simp
@@ -2226,13 +2227,13 @@
 qed
 
 lemma butlast_take:
-  "n \<le> length xs ==> butlast (take n xs) = take (n - 1) xs"
+  "n \<le> length xs \<Longrightarrow> butlast (take n xs) = take (n - 1) xs"
   by (simp add: butlast_conv_take min.absorb1 min.absorb2)
 
 lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)"
   by (simp add: butlast_conv_take drop_take ac_simps)
 
-lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs"
+lemma take_butlast: "n < length xs \<Longrightarrow> take n (butlast xs) = take n xs"
   by (simp add: butlast_conv_take min.absorb1)
 
 lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)"
@@ -2390,7 +2391,7 @@
   by (induct xs) auto
 
 lemma dropWhile_append2 [simp]:
-  "(\<And>x. x \<in> set xs \<Longrightarrow> P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys"
+  "(\<And>x. x \<in> set xs \<Longrightarrow> P(x)) \<Longrightarrow> dropWhile P (xs @ ys) = dropWhile P ys"
   by (induct xs) auto
 
 lemma dropWhile_append3:
@@ -2419,10 +2420,10 @@
   "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys \<and> \<not> P y)"
   by(induct xs, auto)
 
-lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)"
+lemma distinct_takeWhile[simp]: "distinct xs \<Longrightarrow> distinct (takeWhile P xs)"
   by (induct xs) (auto dest: set_takeWhileD)
 
-lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)"
+lemma distinct_dropWhile[simp]: "distinct xs \<Longrightarrow> distinct (dropWhile P xs)"
   by (induct xs) auto
 
 lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \<circ> f) xs)"
@@ -2582,12 +2583,12 @@
   by (induct xs ys rule:list_induct2') auto
 
 lemma zip_append [simp]:
-  "[| length xs = length us |] ==>
+  "\<lbrakk>length xs = length us\<rbrakk> \<Longrightarrow>
   zip (xs@ys) (us@vs) = zip xs us @ zip ys vs"
   by (simp add: zip_append1)
 
 lemma zip_rev:
-  "length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)"
+  "length xs = length ys \<Longrightarrow> zip (rev xs) (rev ys) = rev (zip xs ys)"
   by (induct rule:list_induct2, simp_all)
 
 lemma zip_map_map:
@@ -2622,7 +2623,7 @@
   by(induct xs) auto
 
 lemma nth_zip [simp]:
-  "[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)"
+  "\<lbrakk>i < length xs; i < length ys\<rbrakk> \<Longrightarrow> (zip xs ys)!i = (xs!i, ys!i)"
 proof (induct ys arbitrary: i xs)
   case (Cons y ys)
   then show ?case
@@ -2804,7 +2805,7 @@
 subsubsection \<open>\<^const>\<open>list_all2\<close>\<close>
 
 lemma list_all2_lengthD [intro?]:
-  "list_all2 P xs ys ==> length xs = length ys"
+  "list_all2 P xs ys \<Longrightarrow> length xs = length ys"
   by (simp add: list_all2_iff)
 
 lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])"
@@ -2892,8 +2893,8 @@
   by (force simp add: list_all2_iff set_zip)
 
 lemma list_all2_trans:
-  assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c"
-  shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs"
+  assumes tr: "!!a b c. P1 a b \<Longrightarrow> P2 b c \<Longrightarrow> P3 a c"
+  shows "!!bs cs. list_all2 P1 as bs \<Longrightarrow> list_all2 P2 bs cs \<Longrightarrow> list_all2 P3 as cs"
     (is "!!bs cs. PROP ?Q as bs cs")
 proof (induct as)
   fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs"
@@ -3224,7 +3225,7 @@
 
 lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n
 
-lemma upt_conv_Nil [simp]: "j \<le> i ==> [i..<j] = []"
+lemma upt_conv_Nil [simp]: "j \<le> i \<Longrightarrow> [i..<j] = []"
   by (subst upt_rec) simp
 
 lemma upt_eq_Nil_conv[simp]: "([i..<j] = []) = (j = 0 \<or> j \<le> i)"
@@ -3238,11 +3239,11 @@
     by (simp add: upt_rec)
 qed simp
 
-lemma upt_Suc_append: "i \<le> j ==> [i..<(Suc j)] = [i..<j]@[j]"
+lemma upt_Suc_append: "i \<le> j \<Longrightarrow> [i..<(Suc j)] = [i..<j]@[j]"
   \<comment> \<open>Only needed if \<open>upt_Suc\<close> is deleted from the simpset.\<close>
   by simp
 
-lemma upt_conv_Cons: "i < j ==> [i..<j] = i # [Suc i..<j]"
+lemma upt_conv_Cons: "i < j \<Longrightarrow> [i..<j] = i # [Suc i..<j]"
   by (simp add: upt_rec)
 
 lemma upt_conv_Cons_Cons: \<comment> \<open>no precondition\<close>
@@ -3253,14 +3254,14 @@
   case True then show ?thesis by (simp add: upt_conv_Cons)
 qed
 
-lemma upt_add_eq_append: "i<=j ==> [i..<j+k] = [i..<j]@[j..<j+k]"
+lemma upt_add_eq_append: "i<=j \<Longrightarrow> [i..<j+k] = [i..<j]@[j..<j+k]"
   \<comment> \<open>LOOPS as a simprule, since \<open>j \<le> j\<close>.\<close>
   by (induct k) auto
 
 lemma length_upt [simp]: "length [i..<j] = j - i"
   by (induct j) (auto simp add: Suc_diff_le)
 
-lemma nth_upt [simp]: "i + k < j ==> [i..<j] ! k = i + k"
+lemma nth_upt [simp]: "i + k < j \<Longrightarrow> [i..<j] ! k = i + k"
   by (induct j) (auto simp add: less_Suc_eq nth_append split: nat_diff_split)
 
 lemma hd_upt[simp]: "i < j \<Longrightarrow> hd[i..<j] = i"
@@ -3272,7 +3273,7 @@
 lemma last_upt[simp]: "i < j \<Longrightarrow> last[i..<j] = j - 1"
   by(cases j)(auto simp: upt_Suc_append)
 
-lemma take_upt [simp]: "i+m \<le> n ==> take m [i..<n] = [i..<i+m]"
+lemma take_upt [simp]: "i+m \<le> n \<Longrightarrow> take m [i..<n] = [i..<i+m]"
 proof (induct m arbitrary: i)
   case (Suc m)
   then show ?case
@@ -3288,7 +3289,7 @@
 lemma map_add_upt: "map (\<lambda>i. i + n) [0..<m] = [n..<m + n]"
   by (induct m) simp_all
 
-lemma nth_map_upt: "i < n-m ==> (map f [m..<n]) ! i = f(m+i)"
+lemma nth_map_upt: "i < n-m \<Longrightarrow> (map f [m..<n]) ! i = f(m+i)"
 proof (induct n m  arbitrary: i rule: diff_induct)
   case (3 x y)
   then show ?case
@@ -3324,7 +3325,7 @@
   \<Longrightarrow> xs = ys"
   by (simp add: list_all2_conv_all_nth nth_equalityI)
 
-lemma take_equalityI: "(\<forall>i. take i xs = take i ys) ==> xs = ys"
+lemma take_equalityI: "(\<forall>i. take i xs = take i ys) \<Longrightarrow> xs = ys"
 \<comment> \<open>The famous take-lemma.\<close>
   by (metis length_take min.commute order_refl take_all)
 
@@ -3403,10 +3404,11 @@
 qed
 
 lemma nth_upto[simp]: "i + int k \<le> j \<Longrightarrow> [i..j] ! k = i + int k"
-  apply(induction i j arbitrary: k rule: upto.induct)
-apply(subst upto_rec1)
-apply(auto simp add: nth_Cons')
-done
+proof(induction i j arbitrary: k rule: upto.induct)
+  case (1 i j)
+  then show ?case
+    by (auto simp add: upto_rec1 [of i j] nth_Cons')
+qed
 
 lemma upto_split1:
   "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> [i..k] = [i..j-1] @ [j..k]"
@@ -3454,7 +3456,7 @@
 lemma distinct_remdups [iff]: "distinct (remdups xs)"
 by (induct xs) auto
 
-lemma distinct_remdups_id: "distinct xs ==> remdups xs = xs"
+lemma distinct_remdups_id: "distinct xs \<Longrightarrow> remdups xs = xs"
 by (induct xs, auto)
 
 lemma remdups_id_iff_distinct [simp]: "remdups xs = xs \<longleftrightarrow> distinct xs"
@@ -3491,17 +3493,18 @@
   "distinct (map f xs) \<Longrightarrow> distinct (map f (filter P xs))"
 by (induct xs) auto
 
-lemma distinct_filter [simp]: "distinct xs ==> distinct (filter P xs)"
+lemma distinct_filter [simp]: "distinct xs \<Longrightarrow> distinct (filter P xs)"
 by (induct xs) auto
 
 lemma distinct_upt[simp]: "distinct[i..<j]"
 by (induct j) auto
 
 lemma distinct_upto[simp]: "distinct[i..j]"
-apply(induct i j rule:upto.induct)
-apply(subst upto.simps)
-apply(simp)
-done
+proof (induction i j rule: upto.induct)
+  case (1 i j)
+  then show ?case
+    by (simp add: upto.simps [of i])
+qed
 
 lemma distinct_take[simp]: "distinct xs \<Longrightarrow> distinct (take i xs)"
 proof (induct xs arbitrary: i)
@@ -3531,10 +3534,10 @@
       by auto
   next
     case False
+    have "set (take i xs) \<inter> set (drop (Suc i) xs) = {}"
+      by (metis True d disjoint_insert(1) distinct_append id_take_nth_drop list.set(2))
     then show ?thesis
-      using d anot \<open>i < length xs\<close>
-      apply (simp add: upd_conv_take_nth_drop)
-      by (metis disjoint_insert(1) distinct_append id_take_nth_drop set_simps(2))
+      using d False anot \<open>i < length xs\<close> by (simp add: upd_conv_take_nth_drop)
   qed
 next
   case False with d show ?thesis by auto
@@ -3547,8 +3550,9 @@
    \<rbrakk> \<Longrightarrow> distinct (concat xs)"
 by (induct xs) auto
 
-text \<open>It is best to avoid this indexed version of distinct, but
-sometimes it is useful.\<close>
+text \<open>An iff-version of @{thm distinct_concat} is available further down as \<open>distinct_concat_iff\<close>.\<close>
+
+text \<open>It is best to avoid the following indexed version of distinct, but sometimes it is useful.\<close>
 
 lemma distinct_conv_nth: "distinct xs = (\<forall>i < size xs. \<forall>j < size xs. i \<noteq> j \<longrightarrow> xs!i \<noteq> xs!j)"
 proof (induct xs)
@@ -3580,21 +3584,20 @@
   set(xs[n := x]) = insert x (set xs - {xs!n})"
 by(auto simp: set_eq_iff in_set_conv_nth nth_list_update nth_eq_iff_index_eq)
 
-lemma distinct_swap[simp]: "\<lbrakk> i < size xs; j < size xs \<rbrakk> \<Longrightarrow>
+lemma distinct_swap[simp]: "\<lbrakk> i < size xs; j < size xs\<rbrakk> \<Longrightarrow>
   distinct(xs[i := xs!j, j := xs!i]) = distinct xs"
   apply (simp add: distinct_conv_nth nth_list_update)
-apply safe
-apply metis+
-done
+  apply (safe; metis)
+  done
 
 lemma set_swap[simp]:
   "\<lbrakk> i < size xs; j < size xs \<rbrakk> \<Longrightarrow> set(xs[i := xs!j, j := xs!i]) = set xs"
 by(simp add: set_conv_nth nth_list_update) metis
 
-lemma distinct_card: "distinct xs ==> card (set xs) = size xs"
+lemma distinct_card: "distinct xs \<Longrightarrow> card (set xs) = size xs"
 by (induct xs) auto
 
-lemma card_distinct: "card (set xs) = size xs ==> distinct xs"
+lemma card_distinct: "card (set xs) = size xs \<Longrightarrow> distinct xs"
 proof (induct xs)
   case (Cons x xs)
   show ?case
@@ -3822,18 +3825,12 @@
           using f_mono by (simp add: mono_iff_le_Suc)
       next
         have "?f' ` {0 ..< length (x1 # xs)} = f ` {Suc 0 ..< length (x1 # x2 # xs)}"
-          apply safe
-           apply fastforce
-          subgoal for \<dots> x by (cases x) auto
-          done
+          using less_Suc_eq_0_disj by auto
         also have "\<dots> = f ` {0 ..< length (x1 # x2 # xs)}"
         proof -
           have "f 0 = f (Suc 0)" using \<open>x1 = x2\<close> f_nth[of 0] by simp
           then show ?thesis
-            apply safe
-             apply fastforce
-            subgoal for \<dots> x by (cases x) auto
-            done
+            using less_Suc_eq_0_disj by auto
         qed
         also have "\<dots> = {0 ..< length ys}" by fact
         finally show  "?f' ` {0 ..< length (x1 # xs)} = {0 ..< length ys}" .
@@ -3842,7 +3839,7 @@
     next
       assume "x1 \<noteq> x2"
 
-      have "2 \<le> length ys"
+      have two: "Suc (Suc 0) \<le> length ys"
       proof -
         have "2 = card {f 0, f 1}" using \<open>x1 \<noteq> x2\<close> f_nth[of 0] by auto
         also have "\<dots> \<le> card (f ` {0..< length (x1 # x2 # xs)})"
@@ -3862,26 +3859,18 @@
         then have "Suc 0 \<noteq> f i" for i using \<open>f 0 = 0\<close>
           by (cases i) fastforce+
         then have "Suc 0 \<notin> f ` {0 ..< length (x1 # x2 # xs)}" by auto
-        then show False using f_img \<open>2 \<le> length ys\<close> by auto
+        then show False using f_img two by auto
       qed
       obtain ys' where "ys = x1 # x2 # ys'"
-        using \<open>2 \<le> length ys\<close> f_nth[of 0] f_nth[of 1]
-        apply (cases ys)
-         apply (rename_tac [2] ys', case_tac [2] ys')
-          apply (auto simp: \<open>f 0 = 0\<close> \<open>f (Suc 0) = Suc 0\<close>)
-        done
+        using two f_nth[of 0] f_nth[of 1]
+        by (auto simp: Suc_le_length_iff  \<open>f 0 = 0\<close> \<open>f (Suc 0) = Suc 0\<close>)
+
+      have Suc0_le_f_Suc: "Suc 0 \<le> f (Suc i)" for i
+        by (metis Suc_le_mono \<open>f (Suc 0) = Suc 0\<close> f_mono le0 mono_def)
 
       define f' where "f' x = f (Suc x) - 1" for x
-
-      { fix i
-        have "Suc 0 \<le> f (Suc 0)" using f_nth[of 0] \<open>x1 \<noteq> x2\<close> \<open>f 0 = 0\<close>  by auto
-        also have "\<dots> \<le> f (Suc i)" using f_mono by (rule monoD) arith
-        finally have "Suc 0 \<le> f (Suc i)" .
-      } note Suc0_le_f_Suc = this
-
-      { fix i have "f (Suc i) = Suc (f' i)"
+      have f_Suc: "f (Suc i) = Suc (f' i)" for i
           using Suc0_le_f_Suc[of i] by (auto simp: f'_def)
-      } note f_Suc = this
 
       have "remdups_adj (x2 # xs) = (x2 # ys')"
       proof (intro "3.hyps" exI conjI impI allI)
@@ -3904,15 +3893,15 @@
 qed
 
 lemma hd_remdups_adj[simp]: "hd (remdups_adj xs) = hd xs"
-by (induction xs rule: remdups_adj.induct) simp_all
+  by (induction xs rule: remdups_adj.induct) simp_all
 
 lemma remdups_adj_Cons: "remdups_adj (x # xs) =
   (case remdups_adj xs of [] \<Rightarrow> [x] | y # xs \<Rightarrow> if x = y then y # xs else x # y # xs)"
-by (induct xs arbitrary: x) (auto split: list.splits)
+  by (induct xs arbitrary: x) (auto split: list.splits)
 
 lemma remdups_adj_append_two:
   "remdups_adj (xs @ [x,y]) = remdups_adj (xs @ [x]) @ (if x = y then [] else [y])"
-by (induct xs rule: remdups_adj.induct, simp_all)
+  by (induct xs rule: remdups_adj.induct, simp_all)
 
 lemma remdups_adj_adjacent:
   "Suc i < length (remdups_adj xs) \<Longrightarrow> remdups_adj xs ! i \<noteq> remdups_adj xs ! Suc i"
@@ -4059,8 +4048,8 @@
 proof (induction xs arbitrary: X)
   case (Cons x xs)
   then show ?case
-    apply (auto simp: sum.If_cases sum.remove)
-    by (metis (no_types) Cons.IH Cons.prems(2) diff_eq sum.remove)
+    using sum.remove [of X x "count_list xs"]
+    by (auto simp: sum.If_cases simp flip: diff_eq)
 qed simp
 
 
@@ -4191,6 +4180,14 @@
   "x \<in> set xs \<Longrightarrow> length (removeAll x xs) < length xs"
   by (auto dest: length_filter_less simp add: removeAll_filter_not_eq)
 
+lemma distinct_concat_iff: "distinct (concat xs) \<longleftrightarrow>
+  distinct (removeAll [] xs) \<and> 
+  (\<forall>ys. ys \<in> set xs \<longrightarrow> distinct ys) \<and>
+  (\<forall>ys zs. ys \<in> set xs \<and> zs \<in> set xs \<and> ys \<noteq> zs \<longrightarrow> set ys \<inter> set zs = {})"
+apply (induct xs)
+ apply(simp_all, safe, auto)
+by (metis Int_iff UN_I empty_iff equals0I set_empty)
+
 
 subsubsection \<open>\<^const>\<open>replicate\<close>\<close>
 
@@ -4237,16 +4234,16 @@
   "filter P (replicate n x) = (if P x then replicate n x else [])"
 by(induct n) auto
 
-lemma hd_replicate [simp]: "n \<noteq> 0 ==> hd (replicate n x) = x"
+lemma hd_replicate [simp]: "n \<noteq> 0 \<Longrightarrow> hd (replicate n x) = x"
 by (induct n) auto
 
 lemma tl_replicate [simp]: "tl (replicate n x) = replicate (n - 1) x"
 by (induct n) auto
 
-lemma last_replicate [simp]: "n \<noteq> 0 ==> last (replicate n x) = x"
+lemma last_replicate [simp]: "n \<noteq> 0 \<Longrightarrow> last (replicate n x) = x"
 by (atomize (full), induct n) auto
 
-lemma nth_replicate[simp]: "i < n ==> (replicate n x)!i = x"
+lemma nth_replicate[simp]: "i < n \<Longrightarrow> (replicate n x)!i = x"
 by (induct n arbitrary: i)(auto simp: nth_Cons split: nat.split)
 
 text\<open>Courtesy of Matthias Daum (2 lemmas):\<close>
@@ -4273,7 +4270,7 @@
 lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}"
 by (induct n) auto
 
-lemma set_replicate [simp]: "n \<noteq> 0 ==> set (replicate n x) = {x}"
+lemma set_replicate [simp]: "n \<noteq> 0 \<Longrightarrow> set (replicate n x) = {x}"
 by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc)
 
 lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
@@ -4495,8 +4492,7 @@
     proof (cases "n mod length xs = 0")
       case True
       then show ?thesis
-        apply (simp add: mod_Suc)
-        by (simp add: False Suc.hyps drop_Suc rotate1_hd_tl take_Suc)
+        by (auto simp add: mod_Suc False Suc.hyps drop_Suc rotate1_hd_tl take_Suc Suc_length_conv)
     next
       case False
       with \<open>xs \<noteq> []\<close> Suc
@@ -4508,37 +4504,37 @@
 qed simp
 
 lemma rotate_conv_mod: "rotate n xs = rotate (n mod length xs) xs"
-by(simp add:rotate_drop_take)
+  by(simp add:rotate_drop_take)
 
 lemma rotate_id[simp]: "n mod length xs = 0 \<Longrightarrow> rotate n xs = xs"
-by(simp add:rotate_drop_take)
+  by(simp add:rotate_drop_take)
 
 lemma length_rotate1[simp]: "length(rotate1 xs) = length xs"
-by (cases xs) simp_all
+  by (cases xs) simp_all
 
 lemma length_rotate[simp]: "length(rotate n xs) = length xs"
-by (induct n arbitrary: xs) (simp_all add:rotate_def)
+  by (induct n arbitrary: xs) (simp_all add:rotate_def)
 
 lemma distinct1_rotate[simp]: "distinct(rotate1 xs) = distinct xs"
-by (cases xs) auto
+  by (cases xs) auto
 
 lemma distinct_rotate[simp]: "distinct(rotate n xs) = distinct xs"
-by (induct n) (simp_all add:rotate_def)
+  by (induct n) (simp_all add:rotate_def)
 
 lemma rotate_map: "rotate n (map f xs) = map f (rotate n xs)"
-by(simp add:rotate_drop_take take_map drop_map)
+  by(simp add:rotate_drop_take take_map drop_map)
 
 lemma set_rotate1[simp]: "set(rotate1 xs) = set xs"
-by (cases xs) auto
+  by (cases xs) auto
 
 lemma set_rotate[simp]: "set(rotate n xs) = set xs"
-by (induct n) (simp_all add:rotate_def)
+  by (induct n) (simp_all add:rotate_def)
 
 lemma rotate1_is_Nil_conv[simp]: "(rotate1 xs = []) = (xs = [])"
-by (cases xs) auto
+  by (cases xs) auto
 
 lemma rotate_is_Nil_conv[simp]: "(rotate n xs = []) = (xs = [])"
-by (induct n) (simp_all add:rotate_def)
+  by (induct n) (simp_all add:rotate_def)
 
 lemma rotate_rev:
   "rotate n (rev xs) = rev(rotate (length xs - (n mod length xs)) xs)"
@@ -4564,21 +4560,20 @@
 subsubsection \<open>\<^const>\<open>nths\<close> --- a generalization of \<^const>\<open>nth\<close> to sets\<close>
 
 lemma nths_empty [simp]: "nths xs {} = []"
-by (auto simp add: nths_def)
+  by (auto simp add: nths_def)
 
 lemma nths_nil [simp]: "nths [] A = []"
-by (auto simp add: nths_def)
+  by (auto simp add: nths_def)
 
 lemma nths_all: "\<forall>i < length xs. i \<in> I \<Longrightarrow> nths xs I = xs"
 apply (simp add: nths_def)
 apply (subst filter_True)
- apply (clarsimp simp: in_set_zip subset_iff)
-apply simp
+ apply (auto simp: in_set_zip subset_iff)
 done
 
 lemma length_nths:
   "length (nths xs I) = card{i. i < length xs \<and> i \<in> I}"
-by(simp add: nths_def length_filter_conv_card cong:conj_cong)
+  by(simp add: nths_def length_filter_conv_card cong:conj_cong)
 
 lemma nths_shift_lemma_Suc:
   "map fst (filter (\<lambda>p. P(Suc(snd p))) (zip xs is)) =
@@ -4612,46 +4607,41 @@
 qed (auto simp: nths_def)
 
 lemma nths_map: "nths (map f xs) I = map f (nths xs I)"
-by(induction xs arbitrary: I) (simp_all add: nths_Cons)
+  by(induction xs arbitrary: I) (simp_all add: nths_Cons)
 
 lemma set_nths: "set(nths xs I) = {xs!i|i. i<size xs \<and> i \<in> I}"
-by (induct xs arbitrary: I) (auto simp: nths_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc)
+  by (induct xs arbitrary: I) (auto simp: nths_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc)
 
 lemma set_nths_subset: "set(nths xs I) \<subseteq> set xs"
-by(auto simp add:set_nths)
+  by(auto simp add:set_nths)
 
 lemma notin_set_nthsI[simp]: "x \<notin> set xs \<Longrightarrow> x \<notin> set(nths xs I)"
-by(auto simp add:set_nths)
+  by(auto simp add:set_nths)
 
 lemma in_set_nthsD: "x \<in> set(nths xs I) \<Longrightarrow> x \<in> set xs"
-by(auto simp add:set_nths)
+  by(auto simp add:set_nths)
 
 lemma nths_singleton [simp]: "nths [x] A = (if 0 \<in> A then [x] else [])"
-by (simp add: nths_Cons)
+  by (simp add: nths_Cons)
 
 lemma distinct_nthsI[simp]: "distinct xs \<Longrightarrow> distinct (nths xs I)"
-by (induct xs arbitrary: I) (auto simp: nths_Cons)
+  by (induct xs arbitrary: I) (auto simp: nths_Cons)
 
 lemma nths_upt_eq_take [simp]: "nths l {..<n} = take n l"
-by (induct l rule: rev_induct)
-   (simp_all split: nat_diff_split add: nths_append)
+  by (induct l rule: rev_induct) (simp_all split: nat_diff_split add: nths_append)
 
 lemma nths_nths: "nths (nths xs A) B = nths xs {i \<in> A. \<exists>j \<in> B. card {i' \<in> A. i' < i} = j}"
-apply(induction xs arbitrary: A B)
-apply(auto simp add: nths_Cons card_less_Suc card_less_Suc2)
-done
+  by (induction xs arbitrary: A B) (auto simp add: nths_Cons card_less_Suc card_less_Suc2)
 
 lemma drop_eq_nths: "drop n xs = nths xs {i. i \<ge> n}"
-apply(induction xs arbitrary: n)
-apply (auto simp add: nths_Cons nths_all drop_Cons' intro: arg_cong2[where f=nths, OF refl])
-done
+  by (induction xs arbitrary: n) (auto simp add: nths_Cons nths_all drop_Cons' intro: arg_cong2[where f=nths, OF refl])
 
 lemma nths_drop: "nths (drop n xs) I = nths xs ((+) n ` I)"
 by(force simp: drop_eq_nths nths_nths simp flip: atLeastLessThan_iff
          intro: arg_cong2[where f=nths, OF refl])
 
 lemma filter_eq_nths: "filter P xs = nths xs {i. i<length xs \<and> P(xs!i)}"
-by(induction xs) (auto simp: nths_Cons)
+  by(induction xs) (auto simp: nths_Cons)
 
 lemma filter_in_nths:
   "distinct xs \<Longrightarrow> filter (%x. x \<in> set (nths xs s)) xs = nths xs s"
@@ -4732,18 +4722,20 @@
 subsubsection \<open>\<^const>\<open>splice\<close>\<close>
 
 lemma splice_Nil2 [simp]: "splice xs [] = xs"
-by (cases xs) simp_all
+  by (cases xs) simp_all
 
 lemma length_splice[simp]: "length(splice xs ys) = length xs + length ys"
-by (induct xs ys rule: splice.induct) auto
+  by (induct xs ys rule: splice.induct) auto
 
 lemma split_Nil_iff[simp]: "splice xs ys = [] \<longleftrightarrow> xs = [] \<and> ys = []"
-by (induct xs ys rule: splice.induct) auto
+  by (induct xs ys rule: splice.induct) auto
 
 lemma splice_replicate[simp]: "splice (replicate m x) (replicate n x) = replicate (m+n) x"
-apply(induction "replicate m x" "replicate n x" arbitrary: m n rule: splice.induct)
-apply (auto simp add: Cons_replicate_eq dest: gr0_implies_Suc)
-done
+proof (induction "replicate m x" "replicate n x" arbitrary: m n rule: splice.induct)
+  case (2 x xs)
+  then show ?case 
+    by (auto simp add: Cons_replicate_eq dest: gr0_implies_Suc)
+qed auto
 
 subsubsection \<open>\<^const>\<open>shuffles\<close>\<close>
 
@@ -4768,7 +4760,7 @@
   by (induct xs ys rule: shuffles.induct) auto
 
 lemma splice_in_shuffles [simp, intro]: "splice xs ys \<in> shuffles xs ys"
-by (induction xs ys rule: splice.induct) (simp_all add: Cons_in_shuffles_iff shuffles_commutes)
+  by (induction xs ys rule: splice.induct) (simp_all add: Cons_in_shuffles_iff shuffles_commutes)
 
 lemma Nil_in_shufflesI: "xs = [] \<Longrightarrow> ys = [] \<Longrightarrow> [] \<in> shuffles xs ys"
   by simp
@@ -5001,13 +4993,13 @@
 subsubsection \<open>\<^const>\<open>min\<close> and \<^const>\<open>arg_min\<close>\<close>
 
 lemma min_list_Min: "xs \<noteq> [] \<Longrightarrow> min_list xs = Min (set xs)"
-by (induction xs rule: induct_list012)(auto)
+  by (induction xs rule: induct_list012)(auto)
 
 lemma f_arg_min_list_f: "xs \<noteq> [] \<Longrightarrow> f (arg_min_list f xs) = Min (f ` (set xs))"
-by(induction f xs rule: arg_min_list.induct) (auto simp: min_def intro!: antisym)
+  by(induction f xs rule: arg_min_list.induct) (auto simp: min_def intro!: antisym)
 
 lemma arg_min_list_in: "xs \<noteq> [] \<Longrightarrow> arg_min_list f xs \<in> set xs"
-by(induction xs rule: induct_list012) (auto simp: Let_def)
+  by(induction xs rule: induct_list012) (auto simp: Let_def)
 
 
 subsubsection \<open>(In)finiteness\<close>
@@ -5811,19 +5803,19 @@
 
 lemma sorted_list_of_set_empty:
   "sorted_list_of_set {} = []"
-by (fact sorted_list_of_set.empty)
+  by (fact sorted_list_of_set.empty)
 
 lemma sorted_list_of_set_insert [simp]:
   "finite A \<Longrightarrow> sorted_list_of_set (insert x A) = insort x (sorted_list_of_set (A - {x}))"
-by (fact sorted_list_of_set.insert_remove)
+  by (fact sorted_list_of_set.insert_remove)
 
 lemma sorted_list_of_set_eq_Nil_iff [simp]:
   "finite A \<Longrightarrow> sorted_list_of_set A = [] \<longleftrightarrow> A = {}"
-by (auto simp: sorted_list_of_set.remove)
+  by (auto simp: sorted_list_of_set.remove)
 
 lemma set_sorted_list_of_set [simp]:
   "finite A \<Longrightarrow> set (sorted_list_of_set A) = A"
-by(induct A rule: finite_induct) (simp_all add: set_insort_key)
+  by(induct A rule: finite_induct) (simp_all add: set_insort_key)
 
 lemma sorted_sorted_list_of_set [simp]: "sorted (sorted_list_of_set A)"
 proof (cases "finite A")
@@ -5894,13 +5886,13 @@
   "listsp A []"
   "listsp A (x # xs)"
 
-lemma listsp_mono [mono]: "A \<le> B ==> listsp A \<le> listsp B"
+lemma listsp_mono [mono]: "A \<le> B \<Longrightarrow> listsp A \<le> listsp B"
 by (rule predicate1I, erule listsp.induct, blast+)
 
 lemmas lists_mono = listsp_mono [to_set]
 
 lemma listsp_infI:
-  assumes l: "listsp A l" shows "listsp B l ==> listsp (inf A B) l" using l
+  assumes l: "listsp A l" shows "listsp B l \<Longrightarrow> listsp (inf A B) l" using l
 by induct blast+
 
 lemmas lists_IntI = listsp_infI [to_set]
@@ -5930,12 +5922,12 @@
 
 lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set]
 
-lemma in_listspD [dest!]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
+lemma in_listspD [dest!]: "listsp A xs \<Longrightarrow> \<forall>x\<in>set xs. A x"
 by (rule in_listsp_conv_set [THEN iffD1])
 
 lemmas in_listsD [dest!] = in_listspD [to_set]
 
-lemma in_listspI [intro!]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
+lemma in_listspI [intro!]: "\<forall>x\<in>set xs. A x \<Longrightarrow> listsp A xs"
 by (rule in_listsp_conv_set [THEN iffD2])
 
 lemmas in_listsI [intro!] = in_listspI [to_set]
@@ -6025,24 +6017,31 @@
 
 lemma lexn_length:
   "(xs, ys) \<in> lexn r n \<Longrightarrow> length xs = n \<and> length ys = n"
-by (induct n arbitrary: xs ys) auto
-
-lemma wf_lex [intro!]: "wf r ==> wf (lex r)"
+  by (induct n arbitrary: xs ys) auto
+
+lemma wf_lex [intro!]: 
+  assumes "wf r" shows "wf (lex r)"
   unfolding lex_def
-  apply (rule wf_UN)
-   apply (simp add: wf_lexn)
-  apply (metis DomainE Int_emptyI RangeE lexn_length)
-  done
+proof (rule wf_UN)
+  show "wf (lexn r i)" for i
+    by (simp add: assms wf_lexn)
+  show "\<And>i j. lexn r i \<noteq> lexn r j \<Longrightarrow> Domain (lexn r i) \<inter> Range (lexn r j) = {}"
+    by (metis DomainE Int_emptyI RangeE lexn_length)
+qed
+
 
 lemma lexn_conv:
   "lexn r n =
     {(xs,ys). length xs = n \<and> length ys = n \<and>
     (\<exists>xys x y xs' ys'. xs= xys @ x#xs' \<and> ys= xys @ y # ys' \<and> (x, y) \<in> r)}"
-apply (induct n, simp)
-apply (simp add: image_Collect lex_prod_def, safe, blast)
- apply (rule_tac x = "ab # xys" in exI, simp)
-apply (case_tac xys, simp_all, blast)
-done
+proof (induction n)
+  case (Suc n)
+  then show ?case
+    apply (simp add: image_Collect lex_prod_def, safe, blast)
+     apply (rule_tac x = "ab # xys" in exI, simp)
+    apply (case_tac xys; force)
+    done
+qed auto
 
 text\<open>By Mathias Fleury:\<close>
 proposition lexn_transI:
@@ -6121,7 +6120,7 @@
     (\<exists>xys x y xs' ys'. xs = xys @ x # xs' \<and> ys = xys @ y # ys' \<and> (x, y) \<in> r)}"
 by (force simp add: lex_def lexn_conv)
 
-lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)"
+lemma wf_lenlex [intro!]: "wf r \<Longrightarrow> wf (lenlex r)"
 by (unfold lenlex_def) blast
 
 lemma lenlex_conv:
@@ -6152,18 +6151,21 @@
   by (meson lex_transI trans_inv_image trans_less_than trans_lex_prod)
 
 lemma Nil_notin_lex [iff]: "([], ys) \<notin> lex r"
-by (simp add: lex_conv)
+  by (simp add: lex_conv)
 
 lemma Nil2_notin_lex [iff]: "(xs, []) \<notin> lex r"
-by (simp add:lex_conv)
+  by (simp add:lex_conv)
 
 lemma Cons_in_lex [simp]:
-  "((x # xs, y # ys) \<in> lex r) =
-      ((x, y) \<in> r \<and> length xs = length ys \<or> x = y \<and> (xs, ys) \<in> lex r)"
-  apply (simp add: lex_conv)
-  apply (rule iffI)
-   prefer 2 apply (blast intro: Cons_eq_appendI, clarify)
-  by (metis hd_append append_Nil list.sel(1) list.sel(3) tl_append2)
+  "(x # xs, y # ys) \<in> lex r \<longleftrightarrow> (x, y) \<in> r \<and> length xs = length ys \<or> x = y \<and> (xs, ys) \<in> lex r"
+ (is "?lhs = ?rhs")
+proof
+  assume ?lhs then show ?rhs
+    by (simp add: lex_conv) (metis hd_append list.sel(1) list.sel(3) tl_append2)
+next
+  assume ?rhs then show ?lhs
+    by (simp add: lex_conv) (blast intro: Cons_eq_appendI)
+qed
 
 lemma Nil_lenlex_iff1 [simp]: "([], ns) \<in> lenlex r \<longleftrightarrow> ns \<noteq> []" 
   and Nil_lenlex_iff2 [simp]: "(ns,[]) \<notin> lenlex r"
@@ -6181,24 +6183,23 @@
 
 lemma lex_append_rightI:
   "(xs, ys) \<in> lex r \<Longrightarrow> length vs = length us \<Longrightarrow> (xs @ us, ys @ vs) \<in> lex r"
-by (fastforce simp: lex_def lexn_conv)
+  by (fastforce simp: lex_def lexn_conv)
 
 lemma lex_append_leftI:
   "(ys, zs) \<in> lex r \<Longrightarrow> (xs @ ys, xs @ zs) \<in> lex r"
-by (induct xs) auto
+  by (induct xs) auto
 
 lemma lex_append_leftD:
   "\<forall>x. (x,x) \<notin> r \<Longrightarrow> (xs @ ys, xs @ zs) \<in> lex r \<Longrightarrow> (ys, zs) \<in> lex r"
-by (induct xs) auto
+  by (induct xs) auto
 
 lemma lex_append_left_iff:
   "\<forall>x. (x,x) \<notin> r \<Longrightarrow> (xs @ ys, xs @ zs) \<in> lex r \<longleftrightarrow> (ys, zs) \<in> lex r"
-by(metis lex_append_leftD lex_append_leftI)
+  by(metis lex_append_leftD lex_append_leftI)
 
 lemma lex_take_index:
   assumes "(xs, ys) \<in> lex r"
-  obtains i where "i < length xs" and "i < length ys" and "take i xs =
-take i ys"
+  obtains i where "i < length xs" and "i < length ys" and "take i xs = take i ys"
     and "(xs ! i, ys ! i) \<in> r"
 proof -
   obtain n us x xs' y ys' where "(xs, ys) \<in> lexn r n" and "length xs = n" and "length ys = n"
@@ -6225,40 +6226,47 @@
 by (unfold lexord_def, induct_tac x, auto)
 
 lemma lexord_cons_cons[simp]:
-  "((a # x, b # y) \<in> lexord r) = ((a,b)\<in> r \<or> (a = b \<and> (x,y)\<in> lexord r))"
-  unfolding lexord_def
-  apply (safe, simp_all)
-     apply (metis hd_append list.sel(1))
+  "(a # x, b # y) \<in> lexord r \<longleftrightarrow> (a,b)\<in> r \<or> (a = b \<and> (x,y)\<in> lexord r)"  (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    apply (simp add: lexord_def)
     apply (metis hd_append list.sel(1) list.sel(3) tl_append2)
-   apply blast
-  by (meson Cons_eq_appendI)
+    done
+qed (auto simp add: lexord_def; (blast | meson Cons_eq_appendI))
 
 lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons
 
 lemma lexord_append_rightI: "\<exists> b z. y = b # z \<Longrightarrow> (x, x @ y) \<in> lexord r"
-by (induct_tac x, auto)
+  by (induct_tac x, auto)
 
 lemma lexord_append_left_rightI:
-     "(a,b) \<in> r \<Longrightarrow> (u @ a # x, u @ b # y) \<in> lexord r"
-by (induct_tac u, auto)
+  "(a,b) \<in> r \<Longrightarrow> (u @ a # x, u @ b # y) \<in> lexord r"
+  by (induct_tac u, auto)
 
 lemma lexord_append_leftI: " (u,v) \<in> lexord r \<Longrightarrow> (x @ u, x @ v) \<in> lexord r"
-by (induct x, auto)
+  by (induct x, auto)
 
 lemma lexord_append_leftD:
-     "\<lbrakk> (x @ u, x @ v) \<in> lexord r; (\<forall>a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
-by (erule rev_mp, induct_tac x, auto)
+  "\<lbrakk>(x @ u, x @ v) \<in> lexord r; (\<forall>a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
+  by (erule rev_mp, induct_tac x, auto)
 
 lemma lexord_take_index_conv:
    "((x,y) \<in> lexord r) =
     ((length x < length y \<and> take (length x) y = x) \<or>
      (\<exists>i. i < min(length x)(length y) \<and> take i x = take i y \<and> (x!i,y!i) \<in> r))"
-  apply (unfold lexord_def Let_def, clarsimp)
-  apply (rule_tac f = "(% a b. a \<or> b)" in arg_cong2)
-  apply (metis Cons_nth_drop_Suc append_eq_conv_conj drop_all list.simps(3) not_le)
-  apply auto
-  apply (rule_tac x ="length u" in exI, simp)
-  by (metis id_take_nth_drop)
+proof -
+  have "(\<exists>a v. y = x @ a # v) = (length x < length y \<and> take (length x) y = x)"
+    by (metis Cons_nth_drop_Suc append_eq_conv_conj drop_all list.simps(3) not_le)
+  moreover
+  have "(\<exists>u a b. (a, b) \<in> r \<and> (\<exists>v. x = u @ a # v) \<and> (\<exists>w. y = u @ b # w)) =
+        (\<exists>i<length x. i < length y \<and> take i x = take i y \<and> (x ! i, y ! i) \<in> r)"
+    apply safe
+    using less_iff_Suc_add apply auto[1]
+    by (metis id_take_nth_drop)
+  ultimately show ?thesis
+    by (auto simp: lexord_def Let_def)
+qed
 
 \<comment> \<open>lexord is extension of partial ordering List.lex\<close>
 lemma lexord_lex: "(x,y) \<in> lex r = ((x,y) \<in> lexord r \<and> length x = length y)"
@@ -6539,10 +6547,10 @@
 unfolding measures_def
 by auto
 
-lemma measures_less: "f x < f y ==> (x, y) \<in> measures (f#fs)"
+lemma measures_less: "f x < f y \<Longrightarrow> (x, y) \<in> measures (f#fs)"
 by simp
 
-lemma measures_lesseq: "f x \<le> f y ==> (x, y) \<in> measures fs ==> (x, y) \<in> measures (f#fs)"
+lemma measures_lesseq: "f x \<le> f y \<Longrightarrow> (x, y) \<in> measures fs \<Longrightarrow> (x, y) \<in> measures (f#fs)"
 by auto
 
 
@@ -6685,7 +6693,7 @@
   for r :: "('a \<times> 'b) set"
 where
     Nil:  "([],[]) \<in> listrel r"
-  | Cons: "[| (x,y) \<in> r; (xs,ys) \<in> listrel r |] ==> (x#xs, y#ys) \<in> listrel r"
+  | Cons: "\<lbrakk>(x,y) \<in> r; (xs,ys) \<in> listrel r\<rbrakk> \<Longrightarrow> (x#xs, y#ys) \<in> listrel r"
 
 inductive_cases listrel_Nil1 [elim!]: "([],xs) \<in> listrel r"
 inductive_cases listrel_Nil2 [elim!]: "(xs,[]) \<in> listrel r"
--- a/src/HOL/Map.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Map.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -110,18 +110,19 @@
 
 lemma map_upd_Some_unfold:
   "((m(a\<mapsto>b)) x = Some y) = (x = a \<and> b = y \<or> x \<noteq> a \<and> m x = Some y)"
-by auto
+  by auto
 
 lemma image_map_upd [simp]: "x \<notin> A \<Longrightarrow> m(x \<mapsto> y) ` A = m ` A"
-by auto
+  by auto
 
-lemma finite_range_updI: "finite (range f) \<Longrightarrow> finite (range (f(a\<mapsto>b)))"
-unfolding image_def
-apply (simp (no_asm_use) add:full_SetCompr_eq)
-apply (rule finite_subset)
- prefer 2 apply assumption
-apply (auto)
-done
+lemma finite_range_updI:
+  assumes "finite (range f)" shows "finite (range (f(a\<mapsto>b)))"
+proof -
+  have "range (f(a\<mapsto>b)) \<subseteq> insert (Some b) (range f)"
+    by auto
+  then show ?thesis
+    by (rule finite_subset) (use assms in auto)
+qed
 
 
 subsection \<open>@{term [source] map_of}\<close>
@@ -143,21 +144,19 @@
 
 lemma map_of_eq_Some_iff [simp]:
   "distinct(map fst xys) \<Longrightarrow> (map_of xys x = Some y) = ((x,y) \<in> set xys)"
-apply (induct xys)
- apply simp
-apply (auto simp: map_of_eq_None_iff [symmetric])
-done
+proof (induct xys)
+  case (Cons xy xys)
+  then show ?case
+    by (cases xy) (auto simp flip: map_of_eq_None_iff)
+qed auto
 
 lemma Some_eq_map_of_iff [simp]:
   "distinct(map fst xys) \<Longrightarrow> (Some y = map_of xys x) = ((x,y) \<in> set xys)"
 by (auto simp del: map_of_eq_Some_iff simp: map_of_eq_Some_iff [symmetric])
 
-lemma map_of_is_SomeI [simp]: "\<lbrakk> distinct(map fst xys); (x,y) \<in> set xys \<rbrakk>
-    \<Longrightarrow> map_of xys x = Some y"
-apply (induct xys)
- apply simp
-apply force
-done
+lemma map_of_is_SomeI [simp]: 
+  "\<lbrakk>distinct(map fst xys); (x,y) \<in> set xys\<rbrakk> \<Longrightarrow> map_of xys x = Some y"
+  by simp
 
 lemma map_of_zip_is_None [simp]:
   "length xs = length ys \<Longrightarrow> (map_of (zip xs ys) x = None) = (x \<notin> set xs)"
@@ -233,12 +232,11 @@
   by (induct xs) (simp_all add: fun_eq_iff)
 
 lemma finite_range_map_of: "finite (range (map_of xys))"
-apply (induct xys)
- apply (simp_all add: image_constant)
-apply (rule finite_subset)
- prefer 2 apply assumption
-apply auto
-done
+proof (induct xys)
+  case (Cons a xys)
+  then show ?case
+    using finite_range_updI by fastforce
+qed auto
 
 lemma map_of_SomeD: "map_of xs k = Some y \<Longrightarrow> (k, y) \<in> set xs"
   by (induct xs) (auto split: if_splits)
@@ -334,23 +332,24 @@
 by (rule ext) (auto simp: map_add_def dom_def split: option.split)
 
 lemma map_of_append[simp]: "map_of (xs @ ys) = map_of ys ++ map_of xs"
-unfolding map_add_def
-apply (induct xs)
- apply simp
-apply (rule ext)
-apply (simp split: option.split)
-done
+  unfolding map_add_def
+proof (induct xs)
+  case (Cons a xs)
+  then show ?case
+    by (force split: option.split)
+qed auto
 
 lemma finite_range_map_of_map_add:
   "finite (range f) \<Longrightarrow> finite (range (f ++ map_of l))"
-apply (induct l)
- apply (auto simp del: fun_upd_apply)
-apply (erule finite_range_updI)
-done
+proof (induct l)
+case (Cons a l)
+  then show ?case
+    by (metis finite_range_updI map_add_upd map_of.simps(2))
+qed auto
 
 lemma inj_on_map_add_dom [iff]:
   "inj_on (m ++ m') (dom m') = inj_on m' (dom m')"
-by (fastforce simp: map_add_def dom_def inj_on_def split: option.splits)
+  by (fastforce simp: map_add_def dom_def inj_on_def split: option.splits)
 
 lemma map_upds_fold_map_upd:
   "m(ks[\<mapsto>]vs) = foldl (\<lambda>m (k, v). m(k \<mapsto> v)) m (zip ks vs)"
@@ -369,46 +368,46 @@
 subsection \<open>@{term [source] restrict_map}\<close>
 
 lemma restrict_map_to_empty [simp]: "m|`{} = empty"
-by (simp add: restrict_map_def)
+  by (simp add: restrict_map_def)
 
 lemma restrict_map_insert: "f |` (insert a A) = (f |` A)(a := f a)"
-by (auto simp: restrict_map_def)
+  by (auto simp: restrict_map_def)
 
 lemma restrict_map_empty [simp]: "empty|`D = empty"
-by (simp add: restrict_map_def)
+  by (simp add: restrict_map_def)
 
 lemma restrict_in [simp]: "x \<in> A \<Longrightarrow> (m|`A) x = m x"
-by (simp add: restrict_map_def)
+  by (simp add: restrict_map_def)
 
 lemma restrict_out [simp]: "x \<notin> A \<Longrightarrow> (m|`A) x = None"
-by (simp add: restrict_map_def)
+  by (simp add: restrict_map_def)
 
 lemma ran_restrictD: "y \<in> ran (m|`A) \<Longrightarrow> \<exists>x\<in>A. m x = Some y"
-by (auto simp: restrict_map_def ran_def split: if_split_asm)
+  by (auto simp: restrict_map_def ran_def split: if_split_asm)
 
 lemma dom_restrict [simp]: "dom (m|`A) = dom m \<inter> A"
-by (auto simp: restrict_map_def dom_def split: if_split_asm)
+  by (auto simp: restrict_map_def dom_def split: if_split_asm)
 
 lemma restrict_upd_same [simp]: "m(x\<mapsto>y)|`(-{x}) = m|`(-{x})"
-by (rule ext) (auto simp: restrict_map_def)
+  by (rule ext) (auto simp: restrict_map_def)
 
 lemma restrict_restrict [simp]: "m|`A|`B = m|`(A\<inter>B)"
-by (rule ext) (auto simp: restrict_map_def)
+  by (rule ext) (auto simp: restrict_map_def)
 
 lemma restrict_fun_upd [simp]:
   "m(x := y)|`D = (if x \<in> D then (m|`(D-{x}))(x := y) else m|`D)"
-by (simp add: restrict_map_def fun_eq_iff)
+  by (simp add: restrict_map_def fun_eq_iff)
 
 lemma fun_upd_None_restrict [simp]:
   "(m|`D)(x := None) = (if x \<in> D then m|`(D - {x}) else m|`D)"
-by (simp add: restrict_map_def fun_eq_iff)
+  by (simp add: restrict_map_def fun_eq_iff)
 
 lemma fun_upd_restrict: "(m|`D)(x := y) = (m|`(D-{x}))(x := y)"
-by (simp add: restrict_map_def fun_eq_iff)
+  by (simp add: restrict_map_def fun_eq_iff)
 
 lemma fun_upd_restrict_conv [simp]:
   "x \<in> D \<Longrightarrow> (m|`D)(x := y) = (m|`(D-{x}))(x := y)"
-by (simp add: restrict_map_def fun_eq_iff)
+  by (rule fun_upd_restrict)
 
 lemma map_of_map_restrict:
   "map_of (map (\<lambda>k. (k, f k)) ks) = (Some \<circ> f) |` set ks"
@@ -416,47 +415,62 @@
 
 lemma restrict_complement_singleton_eq:
   "f |` (- {x}) = f(x := None)"
-  by (simp add: restrict_map_def fun_eq_iff)
+  by auto
 
 
 subsection \<open>@{term [source] map_upds}\<close>
 
 lemma map_upds_Nil1 [simp]: "m([] [\<mapsto>] bs) = m"
-by (simp add: map_upds_def)
+  by (simp add: map_upds_def)
 
 lemma map_upds_Nil2 [simp]: "m(as [\<mapsto>] []) = m"
-by (simp add:map_upds_def)
+  by (simp add:map_upds_def)
 
 lemma map_upds_Cons [simp]: "m(a#as [\<mapsto>] b#bs) = (m(a\<mapsto>b))(as[\<mapsto>]bs)"
-by (simp add:map_upds_def)
+  by (simp add:map_upds_def)
 
-lemma map_upds_append1 [simp]: "size xs < size ys \<Longrightarrow>
-  m(xs@[x] [\<mapsto>] ys) = m(xs [\<mapsto>] ys)(x \<mapsto> ys!size xs)"
-apply(induct xs arbitrary: ys m)
- apply (clarsimp simp add: neq_Nil_conv)
-apply (case_tac ys)
- apply simp
-apply simp
-done
+lemma map_upds_append1 [simp]:
+  "size xs < size ys \<Longrightarrow> m(xs@[x] [\<mapsto>] ys) = m(xs [\<mapsto>] ys)(x \<mapsto> ys!size xs)"
+proof (induct xs arbitrary: ys m)
+  case Nil
+  then show ?case
+    by (auto simp: neq_Nil_conv)
+next
+  case (Cons a xs)
+  then show ?case
+    by (cases ys) auto
+qed
 
 lemma map_upds_list_update2_drop [simp]:
   "size xs \<le> i \<Longrightarrow> m(xs[\<mapsto>]ys[i:=y]) = m(xs[\<mapsto>]ys)"
-apply (induct xs arbitrary: m ys i)
- apply simp
-apply (case_tac ys)
- apply simp
-apply (simp split: nat.split)
-done
+proof (induct xs arbitrary: m ys i)
+  case Nil
+  then show ?case
+    by auto
+next
+  case (Cons a xs)
+  then show ?case
+    by (cases ys) (use Cons in \<open>auto split: nat.split\<close>)
+qed
 
+text \<open>Something weirdly sensitive about this proof, which needs only four lines in apply style\<close>
 lemma map_upd_upds_conv_if:
   "(f(x\<mapsto>y))(xs [\<mapsto>] ys) =
    (if x \<in> set(take (length ys) xs) then f(xs [\<mapsto>] ys)
                                     else (f(xs [\<mapsto>] ys))(x\<mapsto>y))"
-apply (induct xs arbitrary: x y ys f)
- apply simp
-apply (case_tac ys)
- apply (auto split: if_split simp: fun_upd_twist)
-done
+proof (induct xs arbitrary: x y ys f)
+  case (Cons a xs)
+  show ?case
+  proof (cases ys)
+    case (Cons z zs)
+    then show ?thesis
+      using Cons.hyps
+      apply (auto split: if_split simp: fun_upd_twist)
+      using Cons.hyps apply fastforce+
+      done
+  qed auto
+qed auto
+
 
 lemma map_upds_twist [simp]:
   "a \<notin> set as \<Longrightarrow> m(a\<mapsto>b)(as[\<mapsto>]bs) = m(as[\<mapsto>]bs)(a\<mapsto>b)"
@@ -464,39 +478,42 @@
 
 lemma map_upds_apply_nontin [simp]:
   "x \<notin> set xs \<Longrightarrow> (f(xs[\<mapsto>]ys)) x = f x"
-apply (induct xs arbitrary: ys)
- apply simp
-apply (case_tac ys)
- apply (auto simp: map_upd_upds_conv_if)
-done
+proof (induct xs arbitrary: ys)
+  case (Cons a xs)
+  then show ?case
+    by (cases ys) (auto simp: map_upd_upds_conv_if)
+qed auto
 
 lemma fun_upds_append_drop [simp]:
   "size xs = size ys \<Longrightarrow> m(xs@zs[\<mapsto>]ys) = m(xs[\<mapsto>]ys)"
-apply (induct xs arbitrary: m ys)
- apply simp
-apply (case_tac ys)
- apply simp_all
-done
+proof (induct xs arbitrary: ys)
+  case (Cons a xs)
+  then show ?case
+    by (cases ys) (auto simp: map_upd_upds_conv_if)
+qed auto
 
 lemma fun_upds_append2_drop [simp]:
   "size xs = size ys \<Longrightarrow> m(xs[\<mapsto>]ys@zs) = m(xs[\<mapsto>]ys)"
-apply (induct xs arbitrary: m ys)
- apply simp
-apply (case_tac ys)
- apply simp_all
-done
-
+proof (induct xs arbitrary: ys)
+  case (Cons a xs)
+  then show ?case
+    by (cases ys) (auto simp: map_upd_upds_conv_if)
+qed auto
 
 lemma restrict_map_upds[simp]:
   "\<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk>
     \<Longrightarrow> m(xs [\<mapsto>] ys)|`D = (m|`(D - set xs))(xs [\<mapsto>] ys)"
-apply (induct xs arbitrary: m ys)
- apply simp
-apply (case_tac ys)
- apply simp
-apply (simp add: Diff_insert [symmetric] insert_absorb)
-apply (simp add: map_upd_upds_conv_if)
-done
+proof (induct xs arbitrary: m ys)
+  case (Cons a xs)
+  then show ?case
+  proof (cases ys)
+    case (Cons z zs)
+    with Cons.hyps Cons.prems show ?thesis
+      apply (simp add: insert_absorb flip: Diff_insert)
+      apply (auto simp add: map_upd_upds_conv_if)
+      done
+  qed auto
+qed auto
 
 
 subsection \<open>@{term [source] dom}\<close>
@@ -537,11 +554,12 @@
 
 lemma dom_map_upds [simp]:
   "dom(m(xs[\<mapsto>]ys)) = set(take (length ys) xs) \<union> dom m"
-apply (induct xs arbitrary: m ys)
- apply simp
-apply (case_tac ys)
- apply auto
-done
+proof (induct xs arbitrary: ys)
+  case (Cons a xs)
+  then show ?case
+    by (cases ys) (auto simp: map_upd_upds_conv_if)
+qed auto
+
 
 lemma dom_map_add [simp]: "dom (m ++ n) = dom n \<union> dom m"
   by (auto simp: dom_def)
@@ -638,12 +656,9 @@
 lemma ran_empty [simp]: "ran empty = {}"
   by (auto simp: ran_def)
 
-lemma ran_map_upd [simp]: "m a = None \<Longrightarrow> ran(m(a\<mapsto>b)) = insert b (ran m)"
+lemma ran_map_upd [simp]:  "m a = None \<Longrightarrow> ran(m(a\<mapsto>b)) = insert b (ran m)"
   unfolding ran_def
-apply auto
-apply (subgoal_tac "aa \<noteq> a")
- apply auto
-done
+  by force
 
 lemma ran_map_add:
   assumes "dom m1 \<inter> dom m2 = {}"
@@ -712,11 +727,11 @@
 
 lemma map_le_upds [simp]:
   "f \<subseteq>\<^sub>m g \<Longrightarrow> f(as [\<mapsto>] bs) \<subseteq>\<^sub>m g(as [\<mapsto>] bs)"
-apply (induct as arbitrary: f g bs)
- apply simp
-apply (case_tac bs)
- apply auto
-done
+proof (induct as arbitrary: f g bs)
+  case (Cons a as)
+  then show ?case
+    by (cases bs) (use Cons in auto)
+qed auto
 
 lemma map_le_implies_dom_le: "(f \<subseteq>\<^sub>m g) \<Longrightarrow> (dom f \<subseteq> dom g)"
   by (fastforce simp add: map_le_def dom_def)
@@ -728,11 +743,8 @@
   by (auto simp add: map_le_def dom_def)
 
 lemma map_le_antisym: "\<lbrakk> f \<subseteq>\<^sub>m g; g \<subseteq>\<^sub>m f \<rbrakk> \<Longrightarrow> f = g"
-unfolding map_le_def
-apply (rule ext)
-apply (case_tac "x \<in> dom f", simp)
-apply (case_tac "x \<in> dom g", simp, fastforce)
-done
+  unfolding map_le_def
+  by (metis ext domIff)
 
 lemma map_le_map_add [simp]: "f \<subseteq>\<^sub>m g ++ f"
   by (fastforce simp: map_le_def)
--- a/src/HOL/Nat.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Nat.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -65,15 +65,15 @@
   by (rule iffI, rule Suc_Rep_inject) simp_all
 
 lemma nat_induct0:
-  assumes "P 0"
-    and "\<And>n. P n \<Longrightarrow> P (Suc n)"
+  assumes "P 0" and "\<And>n. P n \<Longrightarrow> P (Suc n)"
   shows "P n"
-  using assms
-  apply (unfold Zero_nat_def Suc_def)
-  apply (rule Rep_Nat_inverse [THEN subst]) \<comment> \<open>types force good instantiation\<close>
-  apply (erule Nat_Rep_Nat [THEN Nat.induct])
-  apply (iprover elim: Nat_Abs_Nat_inverse [THEN subst])
-  done
+proof -
+  have "P (Abs_Nat (Rep_Nat n))"
+    using assms unfolding Zero_nat_def Suc_def
+    by (iprover intro:  Nat_Rep_Nat [THEN Nat.induct] elim: Nat_Abs_Nat_inverse [THEN subst])
+  then show ?thesis
+    by (simp add: Rep_Nat_inverse)
+qed
 
 free_constructors case_nat for "0 :: nat" | Suc pred
   where "pred (0 :: nat) = (0 :: nat)"
@@ -87,11 +87,7 @@
 setup \<open>Sign.mandatory_path "old"\<close>
 
 old_rep_datatype "0 :: nat" Suc
-    apply (erule nat_induct0)
-    apply assumption
-   apply (rule nat.inject)
-  apply (rule nat.distinct(1))
-  done
+  by (erule nat_induct0) auto
 
 setup \<open>Sign.parent_path\<close>
 
@@ -356,18 +352,11 @@
 qed
 
 lemma one_eq_mult_iff [simp]: "Suc 0 = m * n \<longleftrightarrow> m = Suc 0 \<and> n = Suc 0"
-  apply (rule trans)
-   apply (rule_tac [2] mult_eq_1_iff)
-  apply fastforce
-  done
-
-lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \<longleftrightarrow> m = 1 \<and> n = 1"
-  for m n :: nat
-  unfolding One_nat_def by (rule mult_eq_1_iff)
-
-lemma nat_1_eq_mult_iff [simp]: "1 = m * n \<longleftrightarrow> m = 1 \<and> n = 1"
-  for m n :: nat
-  unfolding One_nat_def by (rule one_eq_mult_iff)
+  by (simp add: eq_commute flip: mult_eq_1_iff)
+
+lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \<longleftrightarrow> m = 1 \<and> n = 1" 
+  and nat_1_eq_mult_iff [simp]: "1 = m * n \<longleftrightarrow> m = 1 \<and> n = 1" for m n :: nat
+  by auto
 
 lemma mult_cancel1 [simp]: "k * m = k * n \<longleftrightarrow> m = n \<or> k = 0"
   for k m n :: nat
@@ -584,22 +573,23 @@
     and less: "m < n \<Longrightarrow> P"
     and eq: "m = n \<Longrightarrow> P"
   shows P
-  apply (rule major [THEN lessE])
-   apply (rule eq)
-   apply blast
-  apply (rule less)
-  apply blast
-  done
+proof (rule major [THEN lessE])
+  show "Suc n = Suc m \<Longrightarrow> P"
+    using eq by blast
+  show "\<And>j. \<lbrakk>m < j; Suc n = Suc j\<rbrakk> \<Longrightarrow> P"
+    by (blast intro: less)
+qed
 
 lemma Suc_lessE:
   assumes major: "Suc i < k"
     and minor: "\<And>j. i < j \<Longrightarrow> k = Suc j \<Longrightarrow> P"
   shows P
-  apply (rule major [THEN lessE])
-   apply (erule lessI [THEN minor])
-  apply (erule Suc_lessD [THEN minor])
-  apply assumption
-  done
+proof (rule major [THEN lessE])
+  show "k = Suc (Suc i) \<Longrightarrow> P"
+    using lessI minor by iprover
+  show "\<And>j. \<lbrakk>Suc i < j; k = Suc j\<rbrakk> \<Longrightarrow> P"
+    using Suc_lessD minor by iprover
+qed
 
 lemma Suc_less_SucD: "Suc m < Suc n \<Longrightarrow> m < n"
   by simp
@@ -786,12 +776,16 @@
   by (induct k) simp_all
 
 text \<open>strict, in both arguments\<close>
-lemma add_less_mono: "i < j \<Longrightarrow> k < l \<Longrightarrow> i + k < j + l"
-  for i j k l :: nat
-  apply (rule add_less_mono1 [THEN less_trans], assumption+)
-  apply (induct j)
-   apply simp_all
-  done
+lemma add_less_mono: 
+  fixes i j k l :: nat
+  assumes "i < j" "k < l" shows "i + k < j + l"
+proof -
+  have "i + k < j + k"
+    by (simp add: add_less_mono1 assms)
+  also have "...  < j + l"
+    using \<open>i < j\<close> by (induction j) (auto simp: assms)
+  finally show ?thesis .
+qed
 
 lemma less_imp_Suc_add: "m < n \<Longrightarrow> \<exists>k. n = Suc (m + k)"
 proof (induct n)
@@ -969,42 +963,55 @@
   for P :: "nat \<Rightarrow> bool"
   by (rule Least_equality[OF _ le0])
 
-lemma Least_Suc: "P n \<Longrightarrow> \<not> P 0 \<Longrightarrow> (LEAST n. P n) = Suc (LEAST m. P (Suc m))"
-  apply (cases n)
-   apply auto
-  apply (frule LeastI)
-  apply (drule_tac P = "\<lambda>x. P (Suc x)" in LeastI)
-  apply (subgoal_tac " (LEAST x. P x) \<le> Suc (LEAST x. P (Suc x))")
-   apply (erule_tac [2] Least_le)
-  apply (cases "LEAST x. P x")
-   apply auto
-  apply (drule_tac P = "\<lambda>x. P (Suc x)" in Least_le)
-  apply (blast intro: order_antisym)
-  done
+lemma Least_Suc:
+  assumes "P n" "\<not> P 0" 
+  shows "(LEAST n. P n) = Suc (LEAST m. P (Suc m))"
+proof (cases n)
+  case (Suc m)
+  show ?thesis
+  proof (rule antisym)
+    show "(LEAST x. P x) \<le> Suc (LEAST x. P (Suc x))"
+      using assms Suc by (force intro: LeastI Least_le)
+    have \<section>: "P (LEAST x. P x)"
+      by (blast intro: LeastI assms)
+    show "Suc (LEAST m. P (Suc m)) \<le> (LEAST n. P n)"
+    proof (cases "(LEAST n. P n)")
+      case 0
+      then show ?thesis
+        using \<section> by (simp add: assms)
+    next
+      case Suc
+      with \<section> show ?thesis
+        by (auto simp: Least_le)
+    qed
+  qed
+qed (use assms in auto)
 
 lemma Least_Suc2: "P n \<Longrightarrow> Q m \<Longrightarrow> \<not> P 0 \<Longrightarrow> \<forall>k. P (Suc k) = Q k \<Longrightarrow> Least P = Suc (Least Q)"
   by (erule (1) Least_Suc [THEN ssubst]) simp
 
-lemma ex_least_nat_le: "\<not> P 0 \<Longrightarrow> P n \<Longrightarrow> \<exists>k\<le>n. (\<forall>i<k. \<not> P i) \<and> P k"
-  for P :: "nat \<Rightarrow> bool"
-  apply (cases n)
-   apply blast
-  apply (rule_tac x="LEAST k. P k" in exI)
-  apply (blast intro: Least_le dest: not_less_Least intro: LeastI_ex)
-  done
-
-lemma ex_least_nat_less: "\<not> P 0 \<Longrightarrow> P n \<Longrightarrow> \<exists>k<n. (\<forall>i\<le>k. \<not> P i) \<and> P (k + 1)"
-  for P :: "nat \<Rightarrow> bool"
-  apply (cases n)
-   apply blast
-  apply (frule (1) ex_least_nat_le)
-  apply (erule exE)
-  apply (case_tac k)
-   apply simp
-  apply (rename_tac k1)
-  apply (rule_tac x=k1 in exI)
-  apply (auto simp add: less_eq_Suc_le)
-  done
+lemma ex_least_nat_le:
+  fixes P :: "nat \<Rightarrow> bool"
+  assumes "P n" "\<not> P 0" 
+  shows "\<exists>k\<le>n. (\<forall>i<k. \<not> P i) \<and> P k"
+proof (cases n)
+  case (Suc m)
+  with assms show ?thesis
+    by (blast intro: Least_le LeastI_ex dest: not_less_Least)
+qed (use assms in auto)
+
+lemma ex_least_nat_less:
+  fixes P :: "nat \<Rightarrow> bool"
+  assumes "P n" "\<not> P 0" 
+  shows "\<exists>k<n. (\<forall>i\<le>k. \<not> P i) \<and> P (Suc k)"
+proof (cases n)
+  case (Suc m)
+  then obtain k where k: "k \<le> n" "\<forall>i<k. \<not> P i" "P k"
+    using ex_least_nat_le [OF assms] by blast
+  show ?thesis 
+    by (cases k) (use assms k less_eq_Suc_le in auto)
+qed (use assms in auto)
+
 
 lemma nat_less_induct:
   fixes P :: "nat \<Rightarrow> bool"
@@ -1070,11 +1077,10 @@
   assumes "P 0"
     and "\<And>n. n > 0 \<Longrightarrow> \<not> P n \<Longrightarrow> \<exists>m. m < n \<and> \<not> P m"
   shows "P n"
-  apply (rule infinite_descent)
-  using assms
-  apply (case_tac "n > 0")
-   apply auto
-  done
+proof (rule infinite_descent)
+  show "\<And>n. \<not> P n \<Longrightarrow> \<exists>m<n. \<not> P m"
+  using assms by (case_tac "n > 0") auto
+qed
 
 text \<open>
   Infinite descent using a mapping to \<open>nat\<close>:
@@ -1178,14 +1184,11 @@
 
 lemma not_add_less1 [iff]: "\<not> i + j < i"
   for i j :: nat
-  apply (rule notI)
-  apply (drule add_lessD1)
-  apply (erule less_irrefl [THEN notE])
-  done
+  by simp
 
 lemma not_add_less2 [iff]: "\<not> j + i < i"
   for i j :: nat
-  by (simp add: add.commute)
+  by simp
 
 lemma add_leD1: "m + k \<le> n \<Longrightarrow> m \<le> n"
   for k m n :: nat
@@ -1193,9 +1196,7 @@
 
 lemma add_leD2: "m + k \<le> n \<Longrightarrow> k \<le> n"
   for k m n :: nat
-  apply (simp add: add.commute)
-  apply (erule add_leD1)
-  done
+  by (force simp add: add.commute dest: add_leD1)
 
 lemma add_leE: "m + k \<le> n \<Longrightarrow> (m \<le> n \<Longrightarrow> k \<le> n \<Longrightarrow> R) \<Longrightarrow> R"
   for k m n :: nat
@@ -1213,10 +1214,7 @@
   by (induct m n rule: diff_induct) simp_all
 
 lemma diff_less_Suc: "m - n < Suc m"
-  apply (induct m n rule: diff_induct)
-    apply (erule_tac [3] less_SucE)
-     apply (simp_all add: less_Suc_eq)
-  done
+  by (induct m n rule: diff_induct) (auto simp: less_Suc_eq)
 
 lemma diff_le_self [simp]: "m - n \<le> m"
   for m n :: nat
@@ -1345,12 +1343,26 @@
 
 lemma mult_less_cancel2 [simp]: "m * k < n * k \<longleftrightarrow> 0 < k \<and> m < n"
   for k m n :: nat
-  apply (safe intro!: mult_less_mono1)
-   apply (cases k)
-    apply auto
-  apply (simp add: linorder_not_le [symmetric])
-  apply (blast intro: mult_le_mono1)
-  done
+proof (intro iffI conjI)
+  assume m: "m * k < n * k"
+  then show "0 < k"
+    by (cases k) auto
+  show "m < n"
+  proof (cases k)
+    case 0
+    then show ?thesis
+      using m by auto
+  next
+    case (Suc k')
+    then show ?thesis
+      using m
+      by (simp flip: linorder_not_le) (blast intro: add_mono mult_le_mono1)
+  qed
+next
+  assume "0 < k \<and> m < n"
+  then show "m * k < n * k"
+    by (blast intro: mult_less_mono1)
+qed
 
 lemma mult_less_cancel1 [simp]: "k * m < k * n \<longleftrightarrow> 0 < k \<and> m < n"
   for k m n :: nat
@@ -1379,16 +1391,18 @@
   by (cases m) (auto intro: le_add1)
 
 text \<open>Lemma for \<open>gcd\<close>\<close>
-lemma mult_eq_self_implies_10: "m = m * n \<Longrightarrow> n = 1 \<or> m = 0"
-  for m n :: nat
-  apply (drule sym)
-  apply (rule disjCI)
-  apply (rule linorder_cases)
-    defer
-    apply assumption
-   apply (drule mult_less_mono2)
-    apply auto
-  done
+lemma mult_eq_self_implies_10: 
+  fixes m n :: nat
+  assumes "m = m * n" shows "n = 1 \<or> m = 0"
+proof (rule disjCI)
+  assume "m \<noteq> 0"
+  show "n = 1"
+  proof (cases n "1::nat" rule: linorder_cases)
+    case greater
+    show ?thesis
+      using assms mult_less_mono2 [OF greater, of m] \<open>m \<noteq> 0\<close> by auto
+  qed (use assms \<open>m \<noteq> 0\<close> in auto)
+qed
 
 lemma mono_times_nat:
   fixes n :: nat
@@ -1849,28 +1863,20 @@
   by (simp add: Nats_def)
 
 lemma Nats_0 [simp]: "0 \<in> \<nat>"
-  apply (simp add: Nats_def)
-  apply (rule range_eqI)
-  apply (rule of_nat_0 [symmetric])
-  done
+  using of_nat_0 [symmetric] unfolding Nats_def
+  by (rule range_eqI)
 
 lemma Nats_1 [simp]: "1 \<in> \<nat>"
-  apply (simp add: Nats_def)
-  apply (rule range_eqI)
-  apply (rule of_nat_1 [symmetric])
-  done
+  using of_nat_1 [symmetric] unfolding Nats_def
+  by (rule range_eqI)
 
 lemma Nats_add [simp]: "a \<in> \<nat> \<Longrightarrow> b \<in> \<nat> \<Longrightarrow> a + b \<in> \<nat>"
-  apply (auto simp add: Nats_def)
-  apply (rule range_eqI)
-  apply (rule of_nat_add [symmetric])
-  done
+  unfolding Nats_def using of_nat_add [symmetric]
+  by (blast intro: range_eqI)
 
 lemma Nats_mult [simp]: "a \<in> \<nat> \<Longrightarrow> b \<in> \<nat> \<Longrightarrow> a * b \<in> \<nat>"
-  apply (auto simp add: Nats_def)
-  apply (rule range_eqI)
-  apply (rule of_nat_mult [symmetric])
-  done
+  unfolding Nats_def using of_nat_mult [symmetric]
+  by (blast intro: range_eqI)
 
 lemma Nats_cases [cases set: Nats]:
   assumes "x \<in> \<nat>"
@@ -2303,21 +2309,21 @@
   qed
 qed
 
-lemma GreatestI_nat:
-  "\<lbrakk> P(k::nat); \<forall>y. P y \<longrightarrow> y \<le> b \<rbrakk> \<Longrightarrow> P (Greatest P)"
-apply(drule (1) ex_has_greatest_nat)
-using GreatestI2_order by auto
-
-lemma Greatest_le_nat:
-  "\<lbrakk> P(k::nat);  \<forall>y. P y \<longrightarrow> y \<le> b \<rbrakk> \<Longrightarrow> k \<le> (Greatest P)"
-apply(frule (1) ex_has_greatest_nat)
-using GreatestI2_order[where P=P and Q=\<open>\<lambda>x. k \<le> x\<close>] by auto
+lemma
+  fixes k::nat
+  assumes "P k" and minor: "\<And>y. P y \<Longrightarrow> y \<le> b"
+  shows GreatestI_nat: "P (Greatest P)" 
+    and Greatest_le_nat: "k \<le> Greatest P"
+proof -
+  obtain x where "P x" "\<And>y. P y \<Longrightarrow> y \<le> x"
+    using assms ex_has_greatest_nat by blast
+  with \<open>P k\<close> show "P (Greatest P)" "k \<le> Greatest P"
+    using GreatestI2_order by blast+
+qed
 
 lemma GreatestI_ex_nat:
-  "\<lbrakk> \<exists>k::nat. P k;  \<forall>y. P y \<longrightarrow> y \<le> b \<rbrakk> \<Longrightarrow> P (Greatest P)"
-apply (erule exE)
-apply (erule (1) GreatestI_nat)
-done
+  "\<lbrakk> \<exists>k::nat. P k;  \<And>y. P y \<Longrightarrow> y \<le> b \<rbrakk> \<Longrightarrow> P (Greatest P)"
+  by (blast intro: GreatestI_nat)
 
 
 subsection \<open>Monotonicity of \<open>funpow\<close>\<close>
@@ -2363,11 +2369,16 @@
   for k m n :: nat
   unfolding dvd_def by (blast intro: right_diff_distrib' [symmetric])
 
-lemma dvd_diffD: "k dvd m - n \<Longrightarrow> k dvd n \<Longrightarrow> n \<le> m \<Longrightarrow> k dvd m"
-  for k m n :: nat
-  apply (erule linorder_not_less [THEN iffD2, THEN add_diff_inverse, THEN subst])
-  apply (blast intro: dvd_add)
-  done
+lemma dvd_diffD: 
+  fixes k m n :: nat
+  assumes "k dvd m - n" "k dvd n" "n \<le> m"
+  shows "k dvd m"
+proof -
+  have "k dvd n + (m - n)"
+    using assms by (blast intro: dvd_add)
+  with assms show ?thesis
+    by simp
+qed
 
 lemma dvd_diffD1: "k dvd m - n \<Longrightarrow> k dvd m \<Longrightarrow> n \<le> m \<Longrightarrow> k dvd n"
   for k m n :: nat
@@ -2384,13 +2395,19 @@
   then show ?thesis ..
 qed
 
-lemma dvd_mult_cancel1: "0 < m \<Longrightarrow> m * n dvd m \<longleftrightarrow> n = 1"
-  for m n :: nat
-  apply auto
-  apply (subgoal_tac "m * n dvd m * 1")
-   apply (drule dvd_mult_cancel)
-    apply auto
-  done
+lemma dvd_mult_cancel1:
+  fixes m n :: nat
+  assumes "0 < m"
+  shows "m * n dvd m \<longleftrightarrow> n = 1"
+proof 
+  assume "m * n dvd m"
+  then have "m * n dvd m * 1"
+    by simp
+  then have "n dvd 1"
+    by (iprover intro: assms dvd_mult_cancel)
+  then show "n = 1"
+    by auto
+qed auto
 
 lemma dvd_mult_cancel2: "0 < m \<Longrightarrow> n * m dvd m \<longleftrightarrow> n = 1"
   for m n :: nat
@@ -2442,14 +2459,6 @@
   for n :: nat
   by (fact mult_1_right)
 
-lemma nat_add_left_cancel: "k + m = k + n \<longleftrightarrow> m = n"
-  for k m n :: nat
-  by (fact add_left_cancel)
-
-lemma nat_add_right_cancel: "m + k = n + k \<longleftrightarrow> m = n"
-  for k m n :: nat
-  by (fact add_right_cancel)
-
 lemma diff_mult_distrib: "(m - n) * k = (m * k) - (n * k)"
   for k m n :: nat
   by (fact left_diff_distrib')
@@ -2458,13 +2467,10 @@
   for k m n :: nat
   by (fact right_diff_distrib')
 
-lemma le_add_diff: "k \<le> n \<Longrightarrow> m \<le> n + m - k"
-  for k m n :: nat
-  by (fact le_add_diff)  (* FIXME delete *)
-
+(*Used in AUTO2 and Groups.le_diff_conv2 (with variables renamed) doesn't work for some reason*)
 lemma le_diff_conv2: "k \<le> j \<Longrightarrow> (i \<le> j - k) = (i + k \<le> j)"
   for i j k :: nat
-  by (fact le_diff_conv2) (* FIXME delete *)
+  by (fact le_diff_conv2) 
 
 lemma diff_self_eq_0 [simp]: "m - m = 0"
   for m :: nat
--- a/src/HOL/Quotient.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Quotient.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -253,27 +253,29 @@
   fixes P::"'a \<Rightarrow> bool"
   and x::"'a"
   assumes a: "equivp R2"
-  shows   "(Ball (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = All (\<lambda>f. P (f x)))"
-  apply(rule iffI)
-  apply(rule allI)
-  apply(drule_tac x="\<lambda>y. f x" in bspec)
-  apply(simp add: in_respects rel_fun_def)
-  apply(rule impI)
-  using a equivp_reflp_symp_transp[of "R2"]
-  apply (auto elim: equivpE reflpE)
-  done
+  shows "(Ball (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = All (\<lambda>f. P (f x)))"
+proof (intro allI iffI)
+  fix f
+  assume "\<forall>f \<in> Respects (R1 ===> R2). P (f x)"
+  moreover have "(\<lambda>y. f x) \<in> Respects (R1 ===> R2)"
+    using a equivp_reflp_symp_transp[of "R2"]
+    by(auto simp add: in_respects rel_fun_def elim: equivpE reflpE)
+  ultimately show "P (f x)"
+    by auto
+qed auto
 
 lemma bex_reg_eqv_range:
   assumes a: "equivp R2"
   shows   "(Bex (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = Ex (\<lambda>f. P (f x)))"
-  apply(auto)
-  apply(rule_tac x="\<lambda>y. f x" in bexI)
-  apply(simp)
-  apply(simp add: Respects_def in_respects rel_fun_def)
-  apply(rule impI)
-  using a equivp_reflp_symp_transp[of "R2"]
-  apply (auto elim: equivpE reflpE)
-  done
+proof -
+  { fix f
+    assume "P (f x)"
+    have "(\<lambda>y. f x) \<in> Respects (R1 ===> R2)"
+      using a equivp_reflp_symp_transp[of "R2"]
+      by (auto simp add: Respects_def in_respects rel_fun_def elim: equivpE reflpE) }
+  then show ?thesis
+    by auto
+qed
 
 (* Next four lemmas are unused *)
 lemma all_reg:
@@ -322,35 +324,42 @@
   assumes q: "Quotient3 R1 Abs1 Rep1"
   and     a: "(R1 ===> R2) f g"
   shows      "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)"
-  apply (auto simp add: Babs_def in_respects rel_fun_def)
-  apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1")
-  using a apply (simp add: Babs_def rel_fun_def)
-  apply (simp add: in_respects rel_fun_def)
-  using Quotient3_rel[OF q]
-  by metis
+proof (clarsimp simp add: Babs_def in_respects rel_fun_def)
+  fix x y
+  assume "R1 x y"
+  then have "x \<in> Respects R1 \<and> y \<in> Respects R1"
+    unfolding in_respects rel_fun_def using Quotient3_rel[OF q]by metis
+  then show "R2 (Babs (Respects R1) f x) (Babs (Respects R1) g y)"
+  using \<open>R1 x y\<close> a by (simp add: Babs_def rel_fun_def)
+qed
 
 lemma babs_prs:
   assumes q1: "Quotient3 R1 Abs1 Rep1"
   and     q2: "Quotient3 R2 Abs2 Rep2"
-  shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f"
-  apply (rule ext)
-  apply (simp add:)
-  apply (subgoal_tac "Rep1 x \<in> Respects R1")
-  apply (simp add: Babs_def Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2])
-  apply (simp add: in_respects Quotient3_rel_rep[OF q1])
-  done
+shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f"
+proof -
+  { fix x
+    have "Rep1 x \<in> Respects R1"
+      by (simp add: in_respects Quotient3_rel_rep[OF q1])
+    then have "Abs2 (Babs (Respects R1) ((Abs1 ---> Rep2) f) (Rep1 x)) = f x" 
+      by (simp add: Babs_def Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2])
+  }
+  then show ?thesis
+    by force
+qed
 
 lemma babs_simp:
   assumes q: "Quotient3 R1 Abs Rep"
   shows "((R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)) = ((R1 ===> R2) f g)"
-  apply(rule iffI)
-  apply(simp_all only: babs_rsp[OF q])
-  apply(auto simp add: Babs_def rel_fun_def)
-  apply(metis Babs_def in_respects  Quotient3_rel[OF q])
-  done
+         (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    unfolding rel_fun_def by (metis Babs_def in_respects  Quotient3_rel[OF q])
+qed (simp add: babs_rsp[OF q])
 
-(* If a user proves that a particular functional relation
-   is an equivalence this may be useful in regularising *)
+text \<open>If a user proves that a particular functional relation
+   is an equivalence, this may be useful in regularising\<close>
 lemma babs_reg_eqv:
   shows "equivp R \<Longrightarrow> Babs (Respects R) P = P"
   by (simp add: fun_eq_iff Babs_def in_respects equivp_reflp)
@@ -372,7 +381,8 @@
   shows "Ex1 (\<lambda>x. x \<in> Respects R \<and> f x) = Ex1 (\<lambda>x. x \<in> Respects R \<and> g x)"
   using a by (auto elim: rel_funE simp add: Ex1_def in_respects)
 
-(* 2 lemmas needed for cleaning of quantifiers *)
+text \<open>Two lemmas needed for cleaning of quantifiers\<close>
+
 lemma all_prs:
   assumes a: "Quotient3 R absf repf"
   shows "Ball (Respects R) ((absf ---> id) f) = All f"
@@ -410,10 +420,10 @@
 lemma ex1_prs:
   assumes "Quotient3 R absf repf"
   shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f"
-  apply (auto simp: Bex1_rel_def Respects_def)
-  apply (metis Quotient3_def assms)
-  apply (metis (full_types) Quotient3_def assms)
-  by (meson Quotient3_rel assms)
+         (is "?lhs = ?rhs")
+  using assms
+  apply (auto simp add: Bex1_rel_def Respects_def)
+  by (metis (full_types) Quotient3_def)+
 
 lemma bex1_bexeq_reg:
   shows "(\<exists>!x\<in>Respects R. P x) \<longrightarrow> (Bex1_rel R (\<lambda>x. P x))"
@@ -543,7 +553,6 @@
 
 subsection \<open>Quotient composition\<close>
 
-
 lemma OOO_quotient3:
   fixes R1 :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
   fixes Abs1 :: "'a \<Rightarrow> 'b" and Rep1 :: "'b \<Rightarrow> 'a"
@@ -558,29 +567,35 @@
 proof -
   have *: "(R1 OOO R2') r r \<and> (R1 OOO R2') s s \<and> (Abs2 \<circ> Abs1) r = (Abs2 \<circ> Abs1) s
            \<longleftrightarrow> (R1 OOO R2') r s" for r s
-    apply safe
-    subgoal for a b c d
-      apply simp
-      apply (rule_tac b="Rep1 (Abs1 r)" in relcomppI)
-      using Quotient3_refl1 R1 rep_abs_rsp apply fastforce
-      apply (rule_tac b="Rep1 (Abs1 s)" in relcomppI)
-       apply (metis (full_types) Rep1 Abs1 Quotient3_rel R2  Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1] Quotient3_rel_abs [OF R1])
-      by (metis Quotient3_rel R1 rep_abs_rsp_left)
-    subgoal for x y
-      apply (drule Abs1)
-        apply (erule Quotient3_refl2 [OF R1])
-       apply (erule Quotient3_refl1 [OF R1])
-      apply (drule Quotient3_refl1 [OF R2], drule Rep1)
-      by (metis (full_types) Quotient3_def R1 relcompp.relcompI)
-    subgoal for x y
-      apply (drule Abs1)
-        apply (erule Quotient3_refl2 [OF R1])
-       apply (erule Quotient3_refl1 [OF R1])
-      apply (drule Quotient3_refl2 [OF R2], drule Rep1)
-      by (metis (full_types) Quotient3_def R1 relcompp.relcompI)
-    subgoal for x y
-      by simp (metis (full_types) Abs1 Quotient3_rel R1 R2)
-    done
+  proof (intro iffI conjI; clarify) 
+    show "(R1 OOO R2') r s"
+      if r: "R1 r a" "R2' a b" "R1 b r" and eq: "(Abs2 \<circ> Abs1) r = (Abs2 \<circ> Abs1) s" 
+        and s: "R1 s c" "R2' c d" "R1 d s" for a b c d
+    proof -
+      have "R1 r (Rep1 (Abs1 r))"
+        using r Quotient3_refl1 R1 rep_abs_rsp by fastforce
+      moreover have "R2' (Rep1 (Abs1 r)) (Rep1 (Abs1 s))"
+        using that
+        apply (simp add: )
+        apply (metis (full_types) Rep1 Abs1 Quotient3_rel R2 Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1] Quotient3_rel_abs [OF R1])
+        done
+      moreover have "R1 (Rep1 (Abs1 s)) s"
+        by (metis s Quotient3_rel R1 rep_abs_rsp_left)
+      ultimately show ?thesis
+        by (metis relcomppI)
+    qed
+  next
+    fix x y
+    assume xy: "R1 r x" "R2' x y" "R1 y s"
+    then have "R2 (Abs1 x) (Abs1 y)"
+      by (iprover dest: Abs1 elim: Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1])
+    then have "R2' (Rep1 (Abs1 x)) (Rep1 (Abs1 x))" "R2' (Rep1 (Abs1 y)) (Rep1 (Abs1 y))"
+      by (simp_all add: Quotient3_refl1 [OF R2] Quotient3_refl2 [OF R2] Rep1)
+    with \<open>R1 r x\<close> \<open>R1 y s\<close> show "(R1 OOO R2') r r" "(R1 OOO R2') s s"
+      by (metis (full_types) Quotient3_def R1 relcompp.relcompI)+
+    show "(Abs2 \<circ> Abs1) r = (Abs2 \<circ> Abs1) s"
+      using xy by simp (metis (full_types) Abs1 Quotient3_rel R1 R2)
+  qed
   show ?thesis
     apply (rule Quotient3I)
     using * apply (simp_all add: o_def Quotient3_abs_rep [OF R2] Quotient3_abs_rep [OF R1])
--- a/src/HOL/Rat.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Rat.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -463,16 +463,12 @@
 
 lemma positive_add: "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x + y)"
   apply transfer
-  apply (simp add: zero_less_mult_iff)
-  apply (elim disjE)
-     apply (simp_all add: add_pos_pos add_neg_neg mult_pos_neg mult_neg_pos mult_neg_neg)
+     apply (auto simp add: zero_less_mult_iff add_pos_pos add_neg_neg mult_pos_neg mult_neg_pos mult_neg_neg)
   done
 
 lemma positive_mult: "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x * y)"
   apply transfer
-  apply (drule (1) mult_pos_pos)
-  apply (simp add: ac_simps)
-  done
+  by (metis fst_conv mult.commute mult_pos_neg2 snd_conv zero_less_mult_iff)
 
 lemma positive_minus: "\<not> positive x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> positive (- x)"
   by transfer (auto simp: neq_iff zero_less_mult_iff mult_less_0_iff)
@@ -495,24 +491,15 @@
     by (rule abs_rat_def)
   show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
     unfolding less_eq_rat_def less_rat_def
-    apply auto
-     apply (drule (1) positive_add)
-     apply (simp_all add: positive_zero)
-    done
+    using positive_add positive_zero by force
   show "a \<le> a"
     unfolding less_eq_rat_def by simp
   show "a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> a \<le> c"
     unfolding less_eq_rat_def less_rat_def
-    apply auto
-    apply (drule (1) positive_add)
-    apply (simp add: algebra_simps)
-    done
+    using positive_add by fastforce
   show "a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> a = b"
     unfolding less_eq_rat_def less_rat_def
-    apply auto
-    apply (drule (1) positive_add)
-    apply (simp add: positive_zero)
-    done
+    using positive_add positive_zero by fastforce
   show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
     unfolding less_eq_rat_def less_rat_def by auto
   show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
@@ -522,9 +509,7 @@
     by (auto dest!: positive_minus)
   show "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
     unfolding less_rat_def
-    apply (drule (1) positive_mult)
-    apply (simp add: algebra_simps)
-    done
+    by (metis diff_zero positive_mult right_diff_distrib')
 qed
 
 end
@@ -712,8 +697,7 @@
 
 lemma of_rat_eq_iff [simp]: "of_rat a = of_rat b \<longleftrightarrow> a = b"
   apply transfer
-  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
-  apply (simp only: of_int_mult [symmetric] of_int_eq_iff)
+  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq flip: of_int_mult)
   done
 
 lemma of_rat_eq_0_iff [simp]: "of_rat a = 0 \<longleftrightarrow> a = 0"
@@ -856,46 +840,28 @@
   unfolding Rats_def by (rule range_eqI) (rule of_rat_1 [symmetric])
 
 lemma Rats_add [simp]: "a \<in> \<rat> \<Longrightarrow> b \<in> \<rat> \<Longrightarrow> a + b \<in> \<rat>"
-  apply (auto simp add: Rats_def)
-  apply (rule range_eqI)
-  apply (rule of_rat_add [symmetric])
-  done
+  by (metis Rats_cases Rats_of_rat of_rat_add)
 
 lemma Rats_minus_iff [simp]: "- a \<in> \<rat> \<longleftrightarrow> a \<in> \<rat>"
-by (metis Rats_cases Rats_of_rat add.inverse_inverse of_rat_minus)
+  by (metis Rats_cases Rats_of_rat add.inverse_inverse of_rat_minus)
 
 lemma Rats_diff [simp]: "a \<in> \<rat> \<Longrightarrow> b \<in> \<rat> \<Longrightarrow> a - b \<in> \<rat>"
-  apply (auto simp add: Rats_def)
-  apply (rule range_eqI)
-  apply (rule of_rat_diff [symmetric])
-  done
+  by (metis Rats_add Rats_minus_iff diff_conv_add_uminus)
 
 lemma Rats_mult [simp]: "a \<in> \<rat> \<Longrightarrow> b \<in> \<rat> \<Longrightarrow> a * b \<in> \<rat>"
-  apply (auto simp add: Rats_def)
-  apply (rule range_eqI)
-  apply (rule of_rat_mult [symmetric])
-  done
+  by (metis Rats_cases Rats_of_rat of_rat_mult)
 
 lemma Rats_inverse [simp]: "a \<in> \<rat> \<Longrightarrow> inverse a \<in> \<rat>"
   for a :: "'a::field_char_0"
-  apply (auto simp add: Rats_def)
-  apply (rule range_eqI)
-  apply (rule of_rat_inverse [symmetric])
-  done
+  by (metis Rats_cases Rats_of_rat of_rat_inverse)
 
 lemma Rats_divide [simp]: "a \<in> \<rat> \<Longrightarrow> b \<in> \<rat> \<Longrightarrow> a / b \<in> \<rat>"
   for a b :: "'a::field_char_0"
-  apply (auto simp add: Rats_def)
-  apply (rule range_eqI)
-  apply (rule of_rat_divide [symmetric])
-  done
+  by (simp add: divide_inverse)
 
 lemma Rats_power [simp]: "a \<in> \<rat> \<Longrightarrow> a ^ n \<in> \<rat>"
   for a :: "'a::field_char_0"
-  apply (auto simp add: Rats_def)
-  apply (rule range_eqI)
-  apply (rule of_rat_power [symmetric])
-  done
+  by (metis Rats_cases Rats_of_rat of_rat_power)
 
 lemma Rats_induct [case_names of_rat, induct set: Rats]: "q \<in> \<rat> \<Longrightarrow> (\<And>r. P (of_rat r)) \<Longrightarrow> P q"
   by (rule Rats_cases) auto
--- a/src/HOL/Real_Vector_Spaces.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Real_Vector_Spaces.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -62,14 +62,10 @@
     and span_raw_def: span = real_vector.span
     and extend_basis_raw_def: extend_basis = real_vector.extend_basis
     and dim_raw_def: dim = real_vector.dim
-    apply unfold_locales
-       apply (rule scaleR_add_right)
-      apply (rule scaleR_add_left)
-     apply (rule scaleR_scaleR)
-    apply (rule scaleR_one)
-   apply (force simp: linear_def)
-  apply (force simp: linear_def real_scaleR_def[abs_def])
-  done
+proof unfold_locales
+  show "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" "Vector_Spaces.linear (*) (*\<^sub>R) = linear"
+    by (force simp: linear_def real_scaleR_def[abs_def])+
+qed (use scaleR_add_right scaleR_add_left scaleR_scaleR scaleR_one in auto)
 
 hide_const (open)\<comment> \<open>locale constants\<close>
   real_vector.dependent
@@ -86,9 +82,10 @@
   rewrites  "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear"
     and "Vector_Spaces.linear (*) (*\<^sub>R) = linear"
   defines construct_raw_def: construct = real_vector.construct
-  apply unfold_locales
-  unfolding linear_def real_scaleR_def
-  by (rule refl)+
+proof unfold_locales
+  show "Vector_Spaces.linear (*) (*\<^sub>R) = linear"
+  unfolding linear_def real_scaleR_def by auto
+qed (auto simp: linear_def)
 
 hide_const (open)\<comment> \<open>locale constants\<close>
   real_vector.construct
@@ -390,8 +387,7 @@
   by (auto simp: Reals_def)
 
 lemma Reals_minus_iff [simp]: "- a \<in> \<real> \<longleftrightarrow> a \<in> \<real>"
-  apply (auto simp: Reals_def)
-  by (metis add.inverse_inverse of_real_minus rangeI)
+  using Reals_minus by fastforce
 
 lemma Reals_diff [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a - b \<in> \<real>"
   by (metis Reals_add Reals_minus_iff add_uminus_conv_diff)
@@ -514,10 +510,8 @@
   using that by (simp add: less_le_not_le pos_le_minus_divideR_eq pos_minus_divideR_le_eq) 
 
 lemma scaleR_image_atLeastAtMost: "c > 0 \<Longrightarrow> scaleR c ` {x..y} = {c *\<^sub>R x..c *\<^sub>R y}"
-  apply (auto intro!: scaleR_left_mono)
-  apply (rule_tac x = "inverse c *\<^sub>R xa" in image_eqI)
-   apply (simp_all add: pos_le_divideR_eq[symmetric] scaleR_scaleR scaleR_one)
-  done
+  apply (auto intro!: scaleR_left_mono simp: image_iff Bex_def)
+  by (meson local.eq_iff pos_divideR_le_eq pos_le_divideR_eq)
 
 end
 
@@ -1273,13 +1267,12 @@
     assume "open S"
     then obtain f where "\<forall>x\<in>S. 0 < f x \<and> (\<forall>y. dist y x < f x \<longrightarrow> y \<in> S)"
       unfolding open_dist bchoice_iff ..
-    then have *: "S = (\<Union>x\<in>S. {x - f x <..} \<inter> {..< x + f x})"
+    then have *: "(\<Union>x\<in>S. {x - f x <..} \<inter> {..< x + f x}) = S" (is "?S = S")
       by (fastforce simp: dist_real_def)
-    show "generate_topology (range lessThan \<union> range greaterThan) S"
-      apply (subst *)
-      apply (intro generate_topology_Union generate_topology.Int)
-       apply (auto intro: generate_topology.Basis)
-      done
+    moreover have "generate_topology (range lessThan \<union> range greaterThan) ?S"
+      by (force intro: generate_topology.Basis generate_topology_Union generate_topology.Int)
+    ultimately show "generate_topology (range lessThan \<union> range greaterThan) S"
+      by simp
   next
     fix S :: "real set"
     assume "generate_topology (range lessThan \<union> range greaterThan) S"
@@ -1548,9 +1541,10 @@
   by (simp add: diff_left diff_right)
 
 lemma flip: "bounded_bilinear (\<lambda>x y. y ** x)"
-  apply standard
-      apply (simp_all add: add_right add_left scaleR_right scaleR_left)
-  by (metis bounded mult.commute)
+proof
+  show "\<exists>K. \<forall>a b. norm (b ** a) \<le> norm a * norm b * K"
+    by (metis bounded mult.commute)
+qed (simp_all add: add_right add_left scaleR_right scaleR_left)
 
 lemma comp1:
   assumes "bounded_linear g"
@@ -1653,11 +1647,10 @@
 qed
 
 lemma bounded_bilinear_mult: "bounded_bilinear ((*) :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra)"
-  apply (rule bounded_bilinear.intro)
-      apply (auto simp: algebra_simps)
-  apply (rule_tac x=1 in exI)
-  apply (simp add: norm_mult_ineq)
-  done
+proof (rule bounded_bilinear.intro)
+  show "\<exists>K. \<forall>a b::'a. norm (a * b) \<le> norm a * norm b * K"
+    by (rule_tac x=1 in exI) (simp add: norm_mult_ineq)
+qed (auto simp: algebra_simps)
 
 lemma bounded_linear_mult_left: "bounded_linear (\<lambda>x::'a::real_normed_algebra. x * y)"
   using bounded_bilinear_mult
@@ -1678,10 +1671,10 @@
   unfolding divide_inverse by (rule bounded_linear_mult_left)
 
 lemma bounded_bilinear_scaleR: "bounded_bilinear scaleR"
-  apply (rule bounded_bilinear.intro)
-      apply (auto simp: algebra_simps)
-  apply (rule_tac x=1 in exI, simp)
-  done
+proof (rule bounded_bilinear.intro)
+  show "\<exists>K. \<forall>a b. norm (a *\<^sub>R b) \<le> norm a * norm b * K"
+    using less_eq_real_def by auto
+qed (auto simp: algebra_simps)
 
 lemma bounded_linear_scaleR_left: "bounded_linear (\<lambda>r. scaleR r x)"
   using bounded_bilinear_scaleR
@@ -1716,11 +1709,11 @@
 
 instance real_normed_algebra_1 \<subseteq> perfect_space
 proof
-  show "\<not> open {x}" for x :: 'a
-    apply (clarsimp simp: open_dist dist_norm)
-    apply (rule_tac x = "x + of_real (e/2)" in exI)
-    apply simp
-    done
+  fix x::'a
+  have "\<And>e. 0 < e \<Longrightarrow> \<exists>y. norm (y - x) < e \<and> y \<noteq> x"
+    by (rule_tac x = "x + of_real (e/2)" in exI) auto
+  then show "\<not> open {x}" 
+    by (clarsimp simp: open_dist dist_norm)
 qed
 
 
@@ -1793,9 +1786,11 @@
 qed
 
 lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top"
-  apply (clarsimp simp: filterlim_at_top)
-  apply (rule_tac c="nat \<lceil>Z + 1\<rceil>" in eventually_sequentiallyI, linarith)
-  done
+proof (clarsimp simp: filterlim_at_top)
+  fix Z
+  show "\<forall>\<^sub>F x in sequentially. Z \<le> real x"
+    by (meson eventually_sequentiallyI nat_ceiling_le_eq)
+qed
 
 lemma filterlim_nat_sequentially: "filterlim nat sequentially at_top"
 proof -
@@ -1814,17 +1809,20 @@
 qed
 
 lemma filterlim_sequentially_iff_filterlim_real:
-  "filterlim f sequentially F \<longleftrightarrow> filterlim (\<lambda>x. real (f x)) at_top F"
-  apply (rule iffI)
-  subgoal using filterlim_compose filterlim_real_sequentially by blast
-  subgoal premises prems
+  "filterlim f sequentially F \<longleftrightarrow> filterlim (\<lambda>x. real (f x)) at_top F" (is "?lhs = ?rhs")
+proof
+  assume ?lhs then show ?rhs
+    using filterlim_compose filterlim_real_sequentially by blast
+next
+  assume R: ?rhs
+  show ?lhs
   proof -
     have "filterlim (\<lambda>x. nat (floor (real (f x)))) sequentially F"
       by (intro filterlim_compose[OF filterlim_nat_sequentially]
-          filterlim_compose[OF filterlim_floor_sequentially] prems)
+          filterlim_compose[OF filterlim_floor_sequentially] R)
     then show ?thesis by simp
   qed
-  done
+qed
 
 
 subsubsection \<open>Limits of Sequences\<close>
@@ -1899,10 +1897,8 @@
   shows "f \<midarrow>a\<rightarrow> l"
 proof -
   have "\<And>S. \<lbrakk>open S; l \<in> S; \<forall>\<^sub>F x in at a. g x \<in> S\<rbrakk> \<Longrightarrow> \<forall>\<^sub>F x in at a. f x \<in> S"
-    apply (clarsimp simp add: eventually_at)
-    apply (rule_tac x="min d R" in exI)
-     apply (auto simp: assms)
-    done
+    apply (simp add: eventually_at)
+    by (metis assms(2) assms(3) dual_order.strict_trans linorder_neqE_linordered_idom)
   then show ?thesis
     using assms by (simp add: tendsto_def)
 qed
@@ -2171,16 +2167,19 @@
       have "dist (u n) l \<le> dist (u n) ((u \<circ> r) n) + dist ((u \<circ> r) n) l"
         by (rule dist_triangle)
       also have "\<dots> < e/2 + e/2"
-        apply (intro add_strict_mono)
-        using N1[of n "r n"] N2[of n] that unfolding comp_def
-        by (auto simp: less_imp_le) (meson assms(2) less_imp_le order.trans seq_suble)
+      proof (intro add_strict_mono)
+        show "dist (u n) ((u \<circ> r) n) < e / 2"
+          using N1[of n "r n"] N2[of n] that unfolding comp_def
+          by (meson assms(2) le_trans max.bounded_iff strict_mono_imp_increasing)
+        show "dist ((u \<circ> r) n) l < e / 2"
+          using N2 that by auto
+      qed
       finally show ?thesis by simp
     qed 
     then show ?thesis unfolding eventually_sequentially by blast
   qed
   have "(\<lambda>n. dist (u n) l) \<longlonglongrightarrow> 0"
-    apply (rule order_tendstoI)
-    using * by auto (meson eventually_sequentiallyI less_le_trans zero_le_dist)
+    by (simp add: less_le_trans * order_tendstoI)
   then show ?thesis using tendsto_dist_iff by auto
 qed
 
--- a/src/HOL/Rings.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Rings.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1975,31 +1975,39 @@
 lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0"
   by (drule mult_strict_right_mono [of b 0]) auto
 
-lemma zero_less_mult_pos: "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
-  apply (cases "b \<le> 0")
-   apply (auto simp add: le_less not_less)
-  apply (drule_tac mult_pos_neg [of a b])
-   apply (auto dest: less_not_sym)
-  done
-
-lemma zero_less_mult_pos2: "0 < b * a \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
-  apply (cases "b \<le> 0")
-   apply (auto simp add: le_less not_less)
-  apply (drule_tac mult_pos_neg2 [of a b])
-   apply (auto dest: less_not_sym)
-  done
+lemma zero_less_mult_pos: 
+  assumes "0 < a * b" "0 < a" shows "0 < b"
+proof (cases "b \<le> 0")
+  case True
+  then show ?thesis
+    using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg [of a b])
+qed (auto simp add: le_less not_less)
+
+
+lemma zero_less_mult_pos2: 
+  assumes "0 < b * a" "0 < a" shows "0 < b"
+proof (cases "b \<le> 0")
+  case True
+  then show ?thesis
+    using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg2 [of a b])
+qed (auto simp add: le_less not_less)
 
 text \<open>Strict monotonicity in both arguments\<close>
 lemma mult_strict_mono:
-  assumes "a < b" and "c < d" and "0 < b" and "0 \<le> c"
+  assumes "a < b" "c < d" "0 < b" "0 \<le> c"
   shows "a * c < b * d"
-  using assms
-  apply (cases "c = 0")
-   apply simp
-  apply (erule mult_strict_right_mono [THEN less_trans])
-   apply (auto simp add: le_less)
-  apply (erule (1) mult_strict_left_mono)
-  done
+proof (cases "c = 0")
+  case True
+  with assms show ?thesis
+    by simp
+next
+  case False
+  with assms have "a*c < b*c"
+    by (simp add: mult_strict_right_mono [OF \<open>a < b\<close>])
+  also have "\<dots> < b*d"
+    by (simp add: assms mult_strict_left_mono)
+  finally show ?thesis .
+qed
 
 text \<open>This weaker variant has more natural premises\<close>
 lemma mult_strict_mono':
@@ -2010,24 +2018,24 @@
 lemma mult_less_le_imp_less:
   assumes "a < b" and "c \<le> d" and "0 \<le> a" and "0 < c"
   shows "a * c < b * d"
-  using assms
-  apply (subgoal_tac "a * c < b * c")
-   apply (erule less_le_trans)
-   apply (erule mult_left_mono)
-   apply simp
-  apply (erule (1) mult_strict_right_mono)
-  done
+proof -
+  have "a * c < b * c"
+    by (simp add: assms mult_strict_right_mono)
+  also have "... \<le> b * d"
+    by (intro mult_left_mono) (use assms in auto)
+  finally show ?thesis .
+qed
 
 lemma mult_le_less_imp_less:
   assumes "a \<le> b" and "c < d" and "0 < a" and "0 \<le> c"
   shows "a * c < b * d"
-  using assms
-  apply (subgoal_tac "a * c \<le> b * c")
-   apply (erule le_less_trans)
-   apply (erule mult_strict_left_mono)
-   apply simp
-  apply (erule (1) mult_right_mono)
-  done
+proof -
+  have "a * c \<le> b * c"
+    by (simp add: assms mult_right_mono)
+  also have "... < b * d"
+    by (intro mult_strict_left_mono) (use assms in auto)
+  finally show ?thesis .
+qed
 
 end
 
@@ -2114,14 +2122,10 @@
   by (simp add: algebra_simps)
 
 lemma mult_left_mono_neg: "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c * a \<le> c * b"
-  apply (drule mult_left_mono [of _ _ "- c"])
-  apply simp_all
-  done
+  by (auto dest: mult_left_mono [of _ _ "- c"])
 
 lemma mult_right_mono_neg: "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a * c \<le> b * c"
-  apply (drule mult_right_mono [of _ _ "- c"])
-  apply simp_all
-  done
+  by (auto dest: mult_right_mono [of _ _ "- c"])
 
 lemma mult_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
   using mult_right_mono_neg [of a 0 b] by simp
@@ -2251,21 +2255,39 @@
   an assumption, but effectively four when the comparison is a goal.
 \<close>
 
-lemma mult_less_cancel_right_disj: "a * c < b * c \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and>  b < a"
-  apply (cases "c = 0")
-   apply (auto simp add: neq_iff mult_strict_right_mono mult_strict_right_mono_neg)
-     apply (auto simp add: not_less not_le [symmetric, of "a*c"] not_le [symmetric, of a])
-     apply (erule_tac [!] notE)
-     apply (auto simp add: less_imp_le mult_right_mono mult_right_mono_neg)
-  done
-
-lemma mult_less_cancel_left_disj: "c * a < c * b \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and>  b < a"
-  apply (cases "c = 0")
-   apply (auto simp add: neq_iff mult_strict_left_mono mult_strict_left_mono_neg)
-     apply (auto simp add: not_less not_le [symmetric, of "c * a"] not_le [symmetric, of a])
-     apply (erule_tac [!] notE)
-     apply (auto simp add: less_imp_le mult_left_mono mult_left_mono_neg)
-  done
+lemma mult_less_cancel_right_disj: "a * c < b * c \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
+proof (cases "c = 0")
+  case False
+  show ?thesis (is "?lhs \<longleftrightarrow> ?rhs")
+  proof
+    assume ?lhs
+    then have "c < 0 \<Longrightarrow> b < a" "c > 0 \<Longrightarrow> b > a"
+      by (auto simp flip: not_le intro: mult_right_mono mult_right_mono_neg)
+    with False show ?rhs 
+      by (auto simp add: neq_iff)
+  next
+    assume ?rhs
+    with False show ?lhs 
+      by (auto simp add: mult_strict_right_mono mult_strict_right_mono_neg)
+  qed
+qed auto
+
+lemma mult_less_cancel_left_disj: "c * a < c * b \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
+proof (cases "c = 0")
+  case False
+  show ?thesis (is "?lhs \<longleftrightarrow> ?rhs")
+  proof
+    assume ?lhs
+    then have "c < 0 \<Longrightarrow> b < a" "c > 0 \<Longrightarrow> b > a"
+      by (auto simp flip: not_le intro: mult_left_mono mult_left_mono_neg)
+    with False show ?rhs 
+      by (auto simp add: neq_iff)
+  next
+    assume ?rhs
+    with False show ?lhs 
+      by (auto simp add: mult_strict_left_mono mult_strict_left_mono_neg)
+  qed
+qed auto
 
 text \<open>
   The ``conjunction of implication'' lemmas produce two cases when the
@@ -2364,29 +2386,29 @@
 lemma add_diff_inverse: "\<not> a < b \<Longrightarrow> b + (a - b) = a"
   by simp
 
-lemma add_le_imp_le_diff: "i + k \<le> n \<Longrightarrow> i \<le> n - k"
-  apply (subst add_le_cancel_right [where c=k, symmetric])
-  apply (frule le_add_diff_inverse2)
-  apply (simp only: add.assoc [symmetric])
-  using add_implies_diff
-  apply fastforce
-  done
+lemma add_le_imp_le_diff: 
+  assumes "i + k \<le> n" shows "i \<le> n - k"
+proof -
+  have "n - (i + k) + i + k = n"
+    by (simp add: assms add.assoc)
+  with assms add_implies_diff have "i + k \<le> n - k + k"
+    by fastforce
+  then show ?thesis
+    by simp
+qed
 
 lemma add_le_add_imp_diff_le:
   assumes 1: "i + k \<le> n"
     and 2: "n \<le> j + k"
   shows "i + k \<le> n \<Longrightarrow> n \<le> j + k \<Longrightarrow> n - k \<le> j"
 proof -
-  have "n - (i + k) + (i + k) = n"
-    using 1 by simp
+  have "n - (i + k) + i + k = n"
+    using 1 by (simp add: add.assoc)
   moreover have "n - k = n - k - i + i"
     using 1 by (simp add: add_le_imp_le_diff)
   ultimately show ?thesis
-    using 2
-    apply (simp add: add.assoc [symmetric])
-    apply (rule add_le_imp_le_diff [of _ k "j + k", simplified add_diff_cancel_right'])
-    apply (simp add: add.commute diff_diff_add)
-    done
+    using 2 add_le_imp_le_diff [of "n-k" k "j + k"]
+    by (simp add: add.commute diff_diff_add)
 qed
 
 lemma less_1_mult: "1 < m \<Longrightarrow> 1 < n \<Longrightarrow> 1 < m * n"
--- a/src/HOL/Set_Interval.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Set_Interval.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -812,10 +812,10 @@
     greaterThanLessThan_def)
 
 lemma atLeastAtMostSuc_conv: "m \<le> Suc n \<Longrightarrow> {m..Suc n} = insert (Suc n) {m..n}"
-by (auto simp add: atLeastAtMost_def)
+  by auto
 
 lemma atLeastAtMost_insertL: "m \<le> n \<Longrightarrow> insert m {Suc m..n} = {m ..n}"
-by auto
+  by auto
 
 text \<open>The analogous result is useful on \<^typ>\<open>int\<close>:\<close>
 (* here, because we don't have an own int section *)
--- a/src/HOL/Tools/BNF/bnf_lift.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Tools/BNF/bnf_lift.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -2012,14 +2012,10 @@
               else bad_input [thm])
       | NONE => (case Lifting_Info.lookup_quotients lthy absT_name of
             SOME {quot_thm = qthm, ...} =>
-              let
-                val qthm = Thm.transfer' lthy qthm;
-              in
-                case [qthm] RL @{thms Quotient_eq_onp_typedef Quotient_eq_onp_type_copy} of
-                  thm :: _ => (thm, Typedef)
-                | _ => (qthm RS @{thm Quotient_Quotient3},
-                   Quotient (find_equiv_thm_via_Quotient qthm))
-              end
+              (case [qthm] RL @{thms Quotient_eq_onp_typedef Quotient_eq_onp_type_copy} of
+                thm :: _ => (thm, Typedef)
+              | _ => (qthm RS @{thm Quotient_Quotient3},
+                 Quotient (find_equiv_thm_via_Quotient qthm)))
           | NONE => (Typedef.get_info lthy absT_name |> hd |> snd |> #type_definition, Typedef))
       | SOME thms => bad_input thms);
     val wits = (Option.map o map) (prepare_term lthy) raw_wits;
--- a/src/HOL/Tools/Sledgehammer/async_manager_legacy.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Tools/Sledgehammer/async_manager_legacy.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -24,7 +24,7 @@
       Runtime.debugging NONE body () handle exn =>
         if Exn.is_interrupt exn then ()
         else writeln ("## INTERNAL ERROR ##\n" ^ Runtime.exn_message exn),
-      Standard_Thread.attributes
+      Isabelle_Thread.attributes
         {name = "async_manager", stack_limit = NONE, interrupts = interrupts});
 
 fun implode_message (workers, work) =
@@ -108,7 +108,7 @@
               NONE
             else
               let
-                val _ = List.app (Standard_Thread.interrupt_unsynchronized o #1) canceling
+                val _ = List.app (Isabelle_Thread.interrupt_unsynchronized o #1) canceling
                 val canceling' = filter (Thread.isActive o #1) canceling
                 val state' = make_state manager timeout_heap' active canceling' messages
               in SOME (map #2 timeout_threads, state') end
--- a/src/HOL/Transcendental.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Transcendental.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -618,10 +618,6 @@
   for r :: "'a::ring_1"
   by (simp add: sum_subtractf)
 
-lemma lemma_realpow_rev_sumr:
-  "(\<Sum>p<Suc n. (x ^ p) * (y ^ (n - p))) = (\<Sum>p<Suc n. (x ^ (n - p)) * (y ^ p))"
-  by (subst sum.nat_diff_reindex[symmetric]) simp
-
 lemma lemma_termdiff2:
   fixes h :: "'a::field"
   assumes h: "h \<noteq> 0"
@@ -629,26 +625,26 @@
          h * (\<Sum>p< n - Suc 0. \<Sum>q< n - Suc 0 - p. (z + h) ^ q * z ^ (n - 2 - q))"
     (is "?lhs = ?rhs")
 proof (cases n)
-  case (Suc n)
+  case (Suc m)
   have 0: "\<And>x k. (\<Sum>n<Suc k. h * (z ^ x * (z ^ (k - n) * (h + z) ^ n))) =
                  (\<Sum>j<Suc k.  h * ((h + z) ^ j * z ^ (x + k - j)))"
-    apply (rule sum.cong [OF refl])
-    by (simp add: power_add [symmetric] mult.commute)
-  have *: "(\<Sum>i<n. z ^ i * ((z + h) ^ (n - i) - z ^ (n - i))) =
-           (\<Sum>i<n. \<Sum>j<n - i. h * ((z + h) ^ j * z ^ (n - Suc j)))"
-    apply (rule sum.cong [OF refl])
-    apply (clarsimp simp add: less_iff_Suc_add sum_distrib_left diff_power_eq_sum ac_simps 0
-        simp del: sum.lessThan_Suc power_Suc)
-    done
-  have "h * ?lhs = h * ?rhs"
-    apply (simp add: right_diff_distrib diff_divide_distrib h mult.assoc [symmetric])
-    using Suc
-    apply (simp add: diff_power_eq_sum h right_diff_distrib [symmetric] mult.assoc
+    by (auto simp add: power_add [symmetric] mult.commute intro: sum.cong)
+  have *: "(\<Sum>i<m. z ^ i * ((z + h) ^ (m - i) - z ^ (m - i))) =
+           (\<Sum>i<m. \<Sum>j<m - i. h * ((z + h) ^ j * z ^ (m - Suc j)))"
+    by (force simp add: less_iff_Suc_add sum_distrib_left diff_power_eq_sum ac_simps 0
+        simp del: sum.lessThan_Suc power_Suc intro: sum.cong)
+  have "h * ?lhs = (z + h) ^ n - z ^ n - h * of_nat n * z ^ (n - Suc 0)"
+    by (simp add: right_diff_distrib diff_divide_distrib h mult.assoc [symmetric])
+  also have "... = h * ((\<Sum>p<Suc m. (z + h) ^ p * z ^ (m - p)) - of_nat (Suc m) * z ^ m)"
+    by (simp add: Suc diff_power_eq_sum h right_diff_distrib [symmetric] mult.assoc
         del: power_Suc sum.lessThan_Suc of_nat_Suc)
-    apply (subst lemma_realpow_rev_sumr)
-    apply (subst sumr_diff_mult_const2)
-    apply (simp add: lemma_termdiff1 sum_distrib_left *)
-    done
+  also have "... = h * ((\<Sum>p<Suc m. (z + h) ^ (m - p) * z ^ p) - of_nat (Suc m) * z ^ m)"
+    by (subst sum.nat_diff_reindex[symmetric]) simp
+  also have "... = h * (\<Sum>i<Suc m. (z + h) ^ (m - i) * z ^ i - z ^ m)"
+    by (simp add: sum_subtractf)
+  also have "... = h * ?rhs"
+    by (simp add: lemma_termdiff1 sum_distrib_left Suc *)
+  finally have "h * ?lhs = h * ?rhs" .
   then show ?thesis
     by (simp add: h)
 qed auto
@@ -656,12 +652,15 @@
 
 lemma real_sum_nat_ivl_bounded2:
   fixes K :: "'a::linordered_semidom"
-  assumes f: "\<And>p::nat. p < n \<Longrightarrow> f p \<le> K"
-    and K: "0 \<le> K"
+  assumes f: "\<And>p::nat. p < n \<Longrightarrow> f p \<le> K" and K: "0 \<le> K"
   shows "sum f {..<n-k} \<le> of_nat n * K"
-  apply (rule order_trans [OF sum_mono [OF f]])
-  apply (auto simp: mult_right_mono K)
-  done
+proof -
+  have "sum f {..<n-k} \<le> (\<Sum>i<n - k. K)"
+    by (rule sum_mono [OF f]) auto
+  also have "... \<le> of_nat n * K"
+    by (auto simp: mult_right_mono K)
+  finally show ?thesis .
+qed
 
 lemma lemma_termdiff3:
   fixes h z :: "'a::real_normed_field"
@@ -678,21 +677,23 @@
   proof (rule mult_right_mono [OF _ norm_ge_zero])
     from norm_ge_zero 2 have K: "0 \<le> K"
       by (rule order_trans)
-    have le_Kn: "\<And>i j n. i + j = n \<Longrightarrow> norm ((z + h) ^ i * z ^ j) \<le> K ^ n"
-      apply (erule subst)
-      apply (simp only: norm_mult norm_power power_add)
-      apply (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K)
-      done
+    have le_Kn: "norm ((z + h) ^ i * z ^ j) \<le> K ^ n" if "i + j = n" for i j n
+    proof -
+      have "norm (z + h) ^ i * norm z ^ j \<le> K ^ i * K ^ j"
+        by (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K)
+      also have "... = K^n"
+        by (metis power_add that)
+      finally show ?thesis
+        by (simp add: norm_mult norm_power) 
+    qed
+    then have "\<And>p q.
+       \<lbrakk>p < n; q < n - Suc 0\<rbrakk> \<Longrightarrow> norm ((z + h) ^ q * z ^ (n - 2 - q)) \<le> K ^ (n - 2)"
+      by simp
+    then
     show "norm (\<Sum>p<n - Suc 0. \<Sum>q<n - Suc 0 - p. (z + h) ^ q * z ^ (n - 2 - q)) \<le>
         of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))"
-      apply (intro
-          order_trans [OF norm_sum]
-          real_sum_nat_ivl_bounded2
-          mult_nonneg_nonneg
-          of_nat_0_le_iff
-          zero_le_power K)
-      apply (rule le_Kn, simp)
-      done
+      by (intro order_trans [OF norm_sum]
+          real_sum_nat_ivl_bounded2 mult_nonneg_nonneg of_nat_0_le_iff zero_le_power K)
   qed
   also have "\<dots> = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h"
     by (simp only: mult.assoc)
@@ -775,39 +776,30 @@
     then have "summable (\<lambda>n. of_nat n * diffs (\<lambda>n. norm (c n)) n * r ^ (n - Suc 0))"
       by (rule diffs_equiv [THEN sums_summable])
     also have "(\<lambda>n. of_nat n * diffs (\<lambda>n. norm (c n)) n * r ^ (n - Suc 0)) =
-      (\<lambda>n. diffs (\<lambda>m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))"
-      apply (rule ext)
-      apply (case_tac n)
-      apply (simp_all add: diffs_def r_neq_0)
-      done
+               (\<lambda>n. diffs (\<lambda>m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))"
+      by (simp add: diffs_def r_neq_0 fun_eq_iff split: nat_diff_split)
     finally have "summable
       (\<lambda>n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))"
       by (rule diffs_equiv [THEN sums_summable])
     also have
       "(\<lambda>n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0)) =
        (\<lambda>n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))"
-      apply (rule ext)
-      apply (case_tac n, simp)
-      apply (rename_tac nat)
-      apply (case_tac nat, simp)
-      apply (simp add: r_neq_0)
-      done
+      by (rule ext) (simp add: r_neq_0 split: nat_diff_split)
     finally show "summable (\<lambda>n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" .
   next
-    fix h :: 'a
-    fix n :: nat
+    fix h :: 'a and n
     assume h: "h \<noteq> 0"
     assume "norm h < r - norm x"
     then have "norm x + norm h < r" by simp
-    with norm_triangle_ineq have xh: "norm (x + h) < r"
+    with norm_triangle_ineq 
+    have xh: "norm (x + h) < r"
       by (rule order_le_less_trans)
-    show "norm (c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \<le>
+    have "norm (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0))
+    \<le> real n * (real (n - Suc 0) * (r ^ (n - 2) * norm h))"
+      by (metis (mono_tags, lifting) h mult.assoc lemma_termdiff3 less_eq_real_def r1 xh)
+    then show "norm (c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \<le>
       norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h"
-      apply (simp only: norm_mult mult.assoc)
-      apply (rule mult_left_mono [OF _ norm_ge_zero])
-      apply (simp add: mult.assoc [symmetric])
-      apply (metis h lemma_termdiff3 less_eq_real_def r1 xh)
-      done
+      by (simp only: norm_mult mult.assoc mult_left_mono [OF _ norm_ge_zero])
   qed
 qed
 
@@ -900,12 +892,10 @@
     and K: "norm x < norm K"
   shows "DERIV (\<lambda>x. \<Sum>n. c n * x^n) x :> (\<Sum>n. diffs c n * x^n)"
 proof -
-  have K2: "norm ((of_real (norm K) + of_real (norm x)) / 2 :: 'a) < norm K"
-    using K
-    apply (auto simp: norm_divide field_simps)
-    apply (rule le_less_trans [of _ "of_real (norm K) + of_real (norm x)"])
-     apply (auto simp: mult_2_right norm_triangle_mono)
-    done
+  have "norm K + norm x < norm K + norm K"
+    using K by force
+  then have K2: "norm ((of_real (norm K) + of_real (norm x)) / 2 :: 'a) < norm K"
+    by (auto simp: norm_triangle_lt norm_divide field_simps)
   then have [simp]: "norm ((of_real (norm K) + of_real (norm x)) :: 'a) < norm K * 2"
     by simp
   have "summable (\<lambda>n. c n * (of_real (norm x + norm K) / 2) ^ n)"
@@ -915,11 +905,8 @@
   moreover have "\<And>x. norm x < norm K \<Longrightarrow> summable (\<lambda>n. diffs(diffs c) n * x ^ n)"
     by (blast intro: sm termdiff_converges powser_inside)
   ultimately show ?thesis
-    apply (rule termdiffs [where K = "of_real (norm x + norm K) / 2"])
-    using K
-      apply (auto simp: field_simps)
-    apply (simp flip: of_real_add)
-    done
+    by (rule termdiffs [where K = "of_real (norm x + norm K) / 2"])
+       (use K in \<open>auto simp: field_simps simp flip: of_real_add\<close>)
 qed
 
 lemma termdiffs_strong_converges_everywhere:
@@ -999,11 +986,12 @@
     by (blast intro: DERIV_continuous)
   then have "((\<lambda>x. \<Sum>n. a n * x ^ n) \<longlongrightarrow> a 0) (at 0)"
     by (simp add: continuous_within)
-  then show ?thesis
-    apply (rule Lim_transform)
+  moreover have "(\<lambda>x. f x - (\<Sum>n. a n * x ^ n)) \<midarrow>0\<rightarrow> 0"
     apply (clarsimp simp: LIM_eq)
     apply (rule_tac x=s in exI)
     using s sm sums_unique by fastforce
+  ultimately show ?thesis
+    by (rule Lim_transform)
 qed
 
 lemma powser_limit_0_strong:
@@ -1015,9 +1003,7 @@
   have *: "((\<lambda>x. if x = 0 then a 0 else f x) \<longlongrightarrow> a 0) (at 0)"
     by (rule powser_limit_0 [OF s]) (auto simp: powser_sums_zero sm)
   show ?thesis
-    apply (subst LIM_equal [where g = "(\<lambda>x. if x = 0 then a 0 else f x)"])
-     apply (simp_all add: *)
-    done
+    by (simp add: * LIM_equal [where g = "(\<lambda>x. if x = 0 then a 0 else f x)"])
 qed
 
 
@@ -1591,10 +1577,10 @@
   have "1 + x \<le> (\<Sum>n<2. inverse (fact n) * x^n)"
     by (auto simp: numeral_2_eq_2)
   also have "\<dots> \<le> (\<Sum>n. inverse (fact n) * x^n)"
-    apply (rule sum_le_suminf [OF summable_exp])
-    using \<open>0 < x\<close>
-    apply (auto  simp add: zero_le_mult_iff)
-    done
+  proof (rule sum_le_suminf [OF summable_exp])
+    show "\<forall>m\<in>- {..<2}. 0 \<le> inverse (fact m) * x ^ m"
+      using \<open>0 < x\<close> by (auto  simp add: zero_le_mult_iff)
+  qed auto
   finally show "1 + x \<le> exp x"
     by (simp add: exp_def)
 qed auto
@@ -2049,9 +2035,7 @@
 proof (cases "0 \<le> x \<or> x \<le> -1")
   case True
   then show ?thesis
-    apply (rule disjE)
-     apply (simp add: exp_ge_add_one_self_aux)
-    using exp_ge_zero order_trans real_add_le_0_iff by blast
+    by (meson exp_ge_add_one_self_aux exp_ge_zero order.trans real_add_le_0_iff)
 next
   case False
   then have ln1: "ln (1 + x) \<le> x"
@@ -2463,11 +2447,12 @@
 lemma powr_non_neg[simp]: "\<not>a powr x < 0" for a x::real
   using powr_ge_pzero[of a x] by arith
 
+lemma inverse_powr: "\<And>y::real. 0 \<le> y \<Longrightarrow> inverse y powr a = inverse (y powr a)"
+    by (simp add: exp_minus ln_inverse powr_def)
+
 lemma powr_divide: "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> (x / y) powr a = (x powr a) / (y powr a)"
   for a b x :: real
-  apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult)
-  apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse)
-  done
+    by (simp add: divide_inverse powr_mult inverse_powr)
 
 lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)"
   for a b x :: "'a::{ln,real_normed_field}"
@@ -3198,17 +3183,20 @@
 
 lemma summable_norm_sin: "summable (\<lambda>n. norm (sin_coeff n *\<^sub>R x^n))"
   for x :: "'a::{real_normed_algebra_1,banach}"
-  unfolding sin_coeff_def
-  apply (rule summable_comparison_test [OF _ summable_norm_exp [where x=x]])
-  apply (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff)
-  done
+proof (rule summable_comparison_test [OF _ summable_norm_exp])
+  show "\<exists>N. \<forall>n\<ge>N. norm (norm (sin_coeff n *\<^sub>R x ^ n)) \<le> norm (x ^ n /\<^sub>R fact n)"
+    unfolding sin_coeff_def
+    by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff)
+qed
 
 lemma summable_norm_cos: "summable (\<lambda>n. norm (cos_coeff n *\<^sub>R x^n))"
   for x :: "'a::{real_normed_algebra_1,banach}"
-  unfolding cos_coeff_def
-  apply (rule summable_comparison_test [OF _ summable_norm_exp [where x=x]])
-  apply (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff)
-  done
+proof (rule summable_comparison_test [OF _ summable_norm_exp])
+  show "\<exists>N. \<forall>n\<ge>N. norm (norm (cos_coeff n *\<^sub>R x ^ n)) \<le> norm (x ^ n /\<^sub>R fact n)"
+    unfolding cos_coeff_def
+    by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff)
+qed
+
 
 lemma sin_converges: "(\<lambda>n. sin_coeff n *\<^sub>R x^n) sums sin x"
   unfolding sin_def
@@ -3230,8 +3218,7 @@
     by (rule sin_converges)
   finally have "(\<lambda>n. of_real (sin_coeff n *\<^sub>R x^n)) sums (sin (of_real x))" .
   then show ?thesis
-    using sums_unique2 sums_of_real [OF sin_converges]
-    by blast
+    using sums_unique2 sums_of_real [OF sin_converges] by blast
 qed
 
 corollary sin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> sin z \<in> \<real>"
@@ -3317,44 +3304,41 @@
 lemma continuous_on_cos_real: "continuous_on {a..b} cos" for a::real
   using continuous_at_imp_continuous_on isCont_cos by blast
 
+
+context
+  fixes f :: "'a::t2_space \<Rightarrow> 'b::{real_normed_field,banach}"
+begin
+
 lemma isCont_sin' [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. sin (f x)) a"
-  for f :: "_ \<Rightarrow> 'a::{real_normed_field,banach}"
   by (rule isCont_o2 [OF _ isCont_sin])
 
-(* FIXME a context for f would be better *)
-
 lemma isCont_cos' [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. cos (f x)) a"
-  for f :: "_ \<Rightarrow> 'a::{real_normed_field,banach}"
   by (rule isCont_o2 [OF _ isCont_cos])
 
 lemma tendsto_sin [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. sin (f x)) \<longlongrightarrow> sin a) F"
-  for f :: "_ \<Rightarrow> 'a::{real_normed_field,banach}"
   by (rule isCont_tendsto_compose [OF isCont_sin])
 
 lemma tendsto_cos [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. cos (f x)) \<longlongrightarrow> cos a) F"
-  for f :: "_ \<Rightarrow> 'a::{real_normed_field,banach}"
   by (rule isCont_tendsto_compose [OF isCont_cos])
 
 lemma continuous_sin [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. sin (f x))"
-  for f :: "_ \<Rightarrow> 'a::{real_normed_field,banach}"
   unfolding continuous_def by (rule tendsto_sin)
 
 lemma continuous_on_sin [continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. sin (f x))"
-  for f :: "_ \<Rightarrow> 'a::{real_normed_field,banach}"
   unfolding continuous_on_def by (auto intro: tendsto_sin)
 
+lemma continuous_cos [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. cos (f x))"
+  unfolding continuous_def by (rule tendsto_cos)
+
+lemma continuous_on_cos [continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. cos (f x))"
+  unfolding continuous_on_def by (auto intro: tendsto_cos)
+
+end
+
 lemma continuous_within_sin: "continuous (at z within s) sin"     
   for z :: "'a::{real_normed_field,banach}"
   by (simp add: continuous_within tendsto_sin)
 
-lemma continuous_cos [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. cos (f x))"
-  for f :: "_ \<Rightarrow> 'a::{real_normed_field,banach}"
-  unfolding continuous_def by (rule tendsto_cos)
-
-lemma continuous_on_cos [continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. cos (f x))"
-  for f :: "_ \<Rightarrow> 'a::{real_normed_field,banach}"
-  unfolding continuous_on_def by (auto intro: tendsto_cos)
-
 lemma continuous_within_cos: "continuous (at z within s) cos"
   for z :: "'a::{real_normed_field,banach}"
   by (simp add: continuous_within tendsto_cos)
@@ -3369,10 +3353,10 @@
   by (simp add: cos_def cos_coeff_def scaleR_conv_of_real)
 
 lemma DERIV_fun_sin: "DERIV g x :> m \<Longrightarrow> DERIV (\<lambda>x. sin (g x)) x :> cos (g x) * m"
-  by (auto intro!: derivative_intros)
+  by (fact derivative_intros)
 
 lemma DERIV_fun_cos: "DERIV g x :> m \<Longrightarrow> DERIV (\<lambda>x. cos(g x)) x :> - sin (g x) * m"
-  by (auto intro!: derivative_eq_intros)
+  by (fact derivative_intros)
 
 
 subsection \<open>Deriving the Addition Formulas\<close>
@@ -3428,15 +3412,16 @@
     have "(-1) ^ ((n - Suc 0) div 2) * (-1) ^ ((p - Suc n) div 2) = - ((-1 :: real) ^ (p div 2))"
       if np: "odd n" "even p"
     proof -
+      have "p > 0"
+        using \<open>n \<le> p\<close> neq0_conv that(1) by blast
+      then have \<section>: "(- 1::real) ^ (p div 2 - Suc 0) = - ((- 1) ^ (p div 2))"
+        using \<open>even p\<close> by (auto simp add: dvd_def power_eq_if)
       from \<open>n \<le> p\<close> np have *: "n - Suc 0 + (p - Suc n) = p - Suc (Suc 0)" "Suc (Suc 0) \<le> p"
         by arith+
       have "(p - Suc (Suc 0)) div 2 = p div 2 - Suc 0"
         by simp
-      with \<open>n \<le> p\<close> np * show ?thesis
-        apply (simp add: power_add [symmetric] div_add [symmetric] del: div_add)
-        apply (metis (no_types) One_nat_def Suc_1 le_div_geq minus_minus
-            mult.left_neutral mult_minus_left power.simps(2) zero_less_Suc)
-        done
+      with \<open>n \<le> p\<close> np  \<section> * show ?thesis
+        by (simp add: flip: div_add power_add)
     qed
     then show ?thesis
       using \<open>n\<le>p\<close> by (auto simp: algebra_simps sin_coeff_def binomial_fact)
@@ -3884,22 +3869,31 @@
 lemma sin_two_pi [simp]: "sin (2 * pi) = 0"
   by (simp add: sin_double)
 
+context
+  fixes w :: "'a::{real_normed_field,banach}"
+
+begin
+
 lemma sin_times_sin: "sin w * sin z = (cos (w - z) - cos (w + z)) / 2"
-  for w :: "'a::{real_normed_field,banach}"
   by (simp add: cos_diff cos_add)
 
 lemma sin_times_cos: "sin w * cos z = (sin (w + z) + sin (w - z)) / 2"
-  for w :: "'a::{real_normed_field,banach}"
   by (simp add: sin_diff sin_add)
 
 lemma cos_times_sin: "cos w * sin z = (sin (w + z) - sin (w - z)) / 2"
-  for w :: "'a::{real_normed_field,banach}"
   by (simp add: sin_diff sin_add)
 
 lemma cos_times_cos: "cos w * cos z = (cos (w - z) + cos (w + z)) / 2"
-  for w :: "'a::{real_normed_field,banach}"
   by (simp add: cos_diff cos_add)
 
+lemma cos_double_cos: "cos (2 * w) = 2 * cos w ^ 2 - 1"
+  by (simp add: cos_double sin_squared_eq)
+
+lemma cos_double_sin: "cos (2 * w) = 1 - 2 * sin w ^ 2"
+  by (simp add: cos_double sin_squared_eq)
+
+end
+
 lemma sin_plus_sin: "sin w + sin z = 2 * sin ((w + z) / 2) * cos ((w - z) / 2)"
   for w :: "'a::{real_normed_field,banach}" 
   apply (simp add: mult.assoc sin_times_cos)
@@ -3924,14 +3918,6 @@
   apply (simp add: field_simps)
   done
 
-lemma cos_double_cos: "cos (2 * z) = 2 * cos z ^ 2 - 1"
-  for z :: "'a::{real_normed_field,banach}"
-  by (simp add: cos_double sin_squared_eq)
-
-lemma cos_double_sin: "cos (2 * z) = 1 - 2 * sin z ^ 2"
-  for z :: "'a::{real_normed_field,banach}"
-  by (simp add: cos_double sin_squared_eq)
-
 lemma sin_pi_minus [simp]: "sin (pi - x) = sin x"
   by (metis sin_minus sin_periodic_pi minus_minus uminus_add_conv_diff)
 
@@ -4074,7 +4060,7 @@
 
 lemma cos_zero_lemma:
   assumes "0 \<le> x" "cos x = 0"
-  shows "\<exists>n. odd n \<and> x = of_nat n * (pi/2) \<and> n > 0"
+  shows "\<exists>n. odd n \<and> x = of_nat n * (pi/2)"
 proof -
   have xle: "x < (1 + real_of_int \<lfloor>x/pi\<rfloor>) * pi"
     using floor_correct [of "x/pi"]
@@ -4101,12 +4087,17 @@
     by (rule_tac x = "Suc (2 * n)" in exI) (simp add: algebra_simps)
 qed
 
-lemma sin_zero_lemma: "0 \<le> x \<Longrightarrow> sin x = 0 \<Longrightarrow> \<exists>n::nat. even n \<and> x = real n * (pi/2)"
-  using cos_zero_lemma [of "x + pi/2"]
-  apply (clarsimp simp add: cos_add)
-  apply (rule_tac x = "n - 1" in exI)
-  apply (simp add: algebra_simps of_nat_diff)
-  done
+lemma sin_zero_lemma:
+  assumes "0 \<le> x" "sin x = 0"
+  shows "\<exists>n::nat. even n \<and> x = real n * (pi/2)"
+proof -
+  obtain n where "odd n" and n: "x + pi/2 = of_nat n * (pi/2)" "n > 0"
+    using cos_zero_lemma [of "x + pi/2"] assms by (auto simp add: cos_add)
+  then have "x = real (n - 1) * (pi / 2)"
+    by (simp add: algebra_simps of_nat_diff)
+  then show ?thesis
+    by (simp add: \<open>odd n\<close>)
+qed
 
 lemma cos_zero_iff:
   "cos x = 0 \<longleftrightarrow> ((\<exists>n. odd n \<and> x = real n * (pi/2)) \<or> (\<exists>n. odd n \<and> x = - (real n * (pi/2))))"
@@ -4146,7 +4137,7 @@
     using that assms by (auto simp: sin_zero_iff)
 qed auto
 
-lemma cos_zero_iff_int: "cos x = 0 \<longleftrightarrow> (\<exists>n. odd n \<and> x = of_int n * (pi/2))"
+lemma cos_zero_iff_int: "cos x = 0 \<longleftrightarrow> (\<exists>i. odd i \<and> x = of_int i * (pi/2))"
 proof -
   have 1: "\<And>n. odd n \<Longrightarrow> \<exists>i. odd i \<and> real n = real_of_int i"
     by (metis even_of_nat of_int_of_nat_eq)
@@ -4159,15 +4150,21 @@
     by (force simp: cos_zero_iff intro!: 1 2 3)
 qed
 
-lemma sin_zero_iff_int: "sin x = 0 \<longleftrightarrow> (\<exists>n. even n \<and> x = of_int n * (pi/2))"
+lemma sin_zero_iff_int: "sin x = 0 \<longleftrightarrow> (\<exists>i. even i \<and> x = of_int i * (pi/2))" (is "?lhs = ?rhs")
 proof safe
-  assume "sin x = 0"
+  assume ?lhs
+  then consider (plus) n where "even n" "x = real n * (pi/2)" | (minus) n where "even n"  "x = - (real n * (pi/2))"
+    using sin_zero_iff by auto
   then show "\<exists>n. even n \<and> x = of_int n * (pi/2)"
-    apply (simp add: sin_zero_iff, safe)
-     apply (metis even_of_nat of_int_of_nat_eq)
-    apply (rule_tac x="- (int n)" in exI)
-    apply simp
-    done
+  proof cases
+    case plus
+    then show ?rhs
+      by (metis even_of_nat of_int_of_nat_eq)
+  next
+    case minus
+    then show ?thesis
+      by (rule_tac x="- (int n)" in exI) simp
+  qed
 next
   fix i :: int
   assume "even i"
@@ -4175,12 +4172,16 @@
     by (cases i rule: int_cases2, simp_all add: sin_zero_iff)
 qed
 
-lemma sin_zero_iff_int2: "sin x = 0 \<longleftrightarrow> (\<exists>n::int. x = of_int n * pi)"
-  apply (simp only: sin_zero_iff_int)
-  apply (safe elim!: evenE)
-   apply (simp_all add: field_simps)
-  using dvd_triv_left apply fastforce
-  done
+lemma sin_zero_iff_int2: "sin x = 0 \<longleftrightarrow> (\<exists>i::int. x = of_int i * pi)"
+proof -
+  have "sin x = 0 \<longleftrightarrow> (\<exists>i. even i \<and> x = real_of_int i * (pi / 2))"
+    by (auto simp: sin_zero_iff_int)
+  also have "... = (\<exists>j. x = real_of_int (2*j) * (pi / 2))"
+    using dvd_triv_left by blast
+  also have "... = (\<exists>i::int. x = of_int i * pi)"
+    by auto
+  finally show ?thesis .
+qed
 
 lemma sin_npi_int [simp]: "sin (pi * of_int n) = 0"
   by (simp add: sin_zero_iff_int2)
@@ -4252,15 +4253,14 @@
 
 lemma sin_x_le_x:
   fixes x :: real
-  assumes x: "x \<ge> 0"
+  assumes "x \<ge> 0"
   shows "sin x \<le> x"
 proof -
   let ?f = "\<lambda>x. x - sin x"
-  from x have "?f x \<ge> ?f 0"
-    apply (rule DERIV_nonneg_imp_nondecreasing)
-    apply (intro allI impI exI[of _ "1 - cos x" for x])
-    apply (auto intro!: derivative_eq_intros simp: field_simps)
-    done
+  have "\<And>u. \<lbrakk>0 \<le> u; u \<le> x\<rbrakk> \<Longrightarrow> \<exists>y. (?f has_real_derivative 1 - cos u) (at u)"
+    by (auto intro!: derivative_eq_intros simp: field_simps)
+  then have "?f x \<ge> ?f 0"
+    by (metis cos_le_one diff_ge_0_iff_ge DERIV_nonneg_imp_nondecreasing [OF assms])
   then show "sin x \<le> x" by simp
 qed
 
@@ -4270,11 +4270,10 @@
   shows "sin x \<ge> - x"
 proof -
   let ?f = "\<lambda>x. x + sin x"
-  from x have "?f x \<ge> ?f 0"
-    apply (rule DERIV_nonneg_imp_nondecreasing)
-    apply (intro allI impI exI[of _ "1 + cos x" for x])
-    apply (auto intro!: derivative_eq_intros simp: field_simps real_0_le_add_iff)
-    done
+  have \<section>: "\<And>u. \<lbrakk>0 \<le> u; u \<le> x\<rbrakk> \<Longrightarrow> \<exists>y. (?f has_real_derivative 1 + cos u) (at u)"
+    by (auto intro!: derivative_eq_intros simp: field_simps)
+  have "?f x \<ge> ?f 0"
+    by (rule DERIV_nonneg_imp_nondecreasing [OF assms]) (use \<section> real_0_le_add_iff in force)
   then show "sin x \<ge> -x" by simp
 qed
 
@@ -4409,9 +4408,8 @@
   have "cos(3 * x) = cos(2*x + x)"
     by simp
   also have "\<dots> = 4 * cos x ^ 3 - 3 * cos x"
-    apply (simp only: cos_add cos_double sin_double)
-    apply (simp add: * field_simps power2_eq_square power3_eq_cube)
-    done
+    unfolding cos_add cos_double sin_double
+    by (simp add: * field_simps power2_eq_square power3_eq_cube)
   finally show ?thesis .
 qed
 
@@ -4482,12 +4480,18 @@
   by (metis Ints_of_int sin_integer_2pi)
 
 lemma sincos_principal_value: "\<exists>y. (- pi < y \<and> y \<le> pi) \<and> (sin y = sin x \<and> cos y = cos x)"
-  apply (rule exI [where x="pi - (2 * pi) * frac ((pi - x) / (2 * pi))"])
-  apply (auto simp: field_simps frac_lt_1)
-   apply (simp_all add: frac_def field_simps)
-   apply (simp_all add: add_divide_distrib diff_divide_distrib)
-   apply (simp_all add: sin_add cos_add mult.assoc [symmetric])
-  done
+proof -
+  define y where "y \<equiv> pi - (2 * pi) * frac ((pi - x) / (2 * pi))"
+  have "-pi < y"" y \<le> pi"
+    by (auto simp: field_simps frac_lt_1 y_def)
+  moreover
+  have "sin y = sin x" "cos y = cos x"
+    unfolding y_def
+     apply (simp_all add: frac_def divide_simps sin_add cos_add)
+    by (metis sin_int_2pin cos_int_2pin diff_zero add.right_neutral mult.commute mult.left_neutral mult_zero_left)+
+  ultimately
+  show ?thesis by metis
+qed
 
 
 subsection \<open>Tangent\<close>
@@ -5238,10 +5242,11 @@
 qed (use assms isCont_arccos in auto)
 
 lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\<^sup>2)"
-proof (rule DERIV_inverse_function [where f=tan and a="x - 1" and b="x + 1"])
-  show "(tan has_real_derivative 1 + x\<^sup>2) (at (arctan x))"
-    apply (rule derivative_eq_intros | simp)+
+proof (rule DERIV_inverse_function)
+  have "inverse ((cos (arctan x))\<^sup>2) = 1 + x\<^sup>2"
     by (metis arctan cos_arctan_not_zero power_inverse tan_sec)
+  then show "(tan has_real_derivative 1 + x\<^sup>2) (at (arctan x))"
+    by (auto intro!: derivative_eq_intros)
   show "\<And>y. \<lbrakk>x - 1 < y; y < x + 1\<rbrakk> \<Longrightarrow> tan (arctan y) = y"
     using tan_arctan by blast
   show "1 + x\<^sup>2 \<noteq> 0"
@@ -5999,19 +6004,15 @@
   assumes "x \<noteq> 0"
   shows "arctan (1 / x) = sgn x * pi/2 - arctan x"
 proof (rule arctan_unique)
+  have \<section>: "x > 0 \<Longrightarrow> arctan x < pi"
+    using arctan_bounded [of x] by linarith 
   show "- (pi/2) < sgn x * pi/2 - arctan x"
-    using arctan_bounded [of x] assms
-    unfolding sgn_real_def
-    apply (auto simp: arctan algebra_simps)
-    apply (drule zero_less_arctan_iff [THEN iffD2], arith)
-    done
+    using assms by (auto simp: sgn_real_def arctan algebra_simps \<section>)
   show "sgn x * pi/2 - arctan x < pi/2"
     using arctan_bounded [of "- x"] assms
-    unfolding sgn_real_def arctan_minus
-    by (auto simp: algebra_simps)
+    by (auto simp: algebra_simps sgn_real_def arctan_minus)
   show "tan (sgn x * pi/2 - arctan x) = 1 / x"
-    unfolding tan_inverse [of "arctan x", unfolded tan_arctan]
-    unfolding sgn_real_def
+    unfolding tan_inverse [of "arctan x", unfolded tan_arctan] sgn_real_def
     by (simp add: tan_def cos_arctan sin_arctan sin_diff cos_diff)
 qed
 
@@ -6032,18 +6033,18 @@
   by (rule power2_le_imp_le [OF _ zero_le_one])
     (simp add: power_divide divide_le_eq not_sum_power2_lt_zero)
 
-lemmas cos_arccos_lemma1 = cos_arccos_abs [OF cos_x_y_le_one]
-
-lemmas sin_arccos_lemma1 = sin_arccos_abs [OF cos_x_y_le_one]
-
 lemma polar_Ex: "\<exists>r::real. \<exists>a. x = r * cos a \<and> y = r * sin a"
 proof -
-  have polar_ex1: "0 < y \<Longrightarrow> \<exists>r a. x = r * cos a \<and> y = r * sin a" for y
-    apply (rule exI [where x = "sqrt (x\<^sup>2 + y\<^sup>2)"])
-    apply (rule exI [where x = "arccos (x / sqrt (x\<^sup>2 + y\<^sup>2))"])
-    apply (simp add: cos_arccos_lemma1 sin_arccos_lemma1 power_divide
-        real_sqrt_mult [symmetric] right_diff_distrib)
-    done
+  have polar_ex1: "\<exists>r a. x = r * cos a \<and> y = r * sin a" if "0 < y" for y
+  proof -
+    have "x = sqrt (x\<^sup>2 + y\<^sup>2) * cos (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))"
+      by (simp add: cos_arccos_abs [OF cos_x_y_le_one])
+    moreover have "y = sqrt (x\<^sup>2 + y\<^sup>2) * sin (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))"
+      using that
+      by (simp add: sin_arccos_abs [OF cos_x_y_le_one] power_divide right_diff_distrib flip: real_sqrt_mult)
+    ultimately show ?thesis
+      by blast
+  qed
   show ?thesis
   proof (cases "0::real" y rule: linorder_cases)
     case less
@@ -6083,24 +6084,24 @@
   assumes m: "\<And>i. i > m \<Longrightarrow> a i = 0"
     and n: "\<And>j. j > n \<Longrightarrow> b j = 0"
   shows "(\<Sum>i\<le>m. (a i) * x ^ i) * (\<Sum>j\<le>n. (b j) * x ^ j) =
-    (\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
-proof -
+         (\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
+proof -
+  have "\<And>i j. \<lbrakk>m + n - i < j; a i \<noteq> 0\<rbrakk> \<Longrightarrow> b j = 0"
+    by (meson le_add_diff leI le_less_trans m n)
+  then have \<section>: "(\<Sum>(i,j)\<in>(SIGMA i:{..m+n}. {m+n - i<..m+n}). a i * x ^ i * (b j * x ^ j)) = 0"
+    by (clarsimp simp add: sum_Un Sigma_interval_disjoint intro!: sum.neutral)
   have "(\<Sum>i\<le>m. (a i) * x ^ i) * (\<Sum>j\<le>n. (b j) * x ^ j) = (\<Sum>i\<le>m. \<Sum>j\<le>n. (a i * x ^ i) * (b j * x ^ j))"
     by (rule sum_product)
   also have "\<dots> = (\<Sum>i\<le>m + n. \<Sum>j\<le>n + m. a i * x ^ i * (b j * x ^ j))"
     using assms by (auto simp: sum_up_index_split)
   also have "\<dots> = (\<Sum>r\<le>m + n. \<Sum>j\<le>m + n - r. a r * x ^ r * (b j * x ^ j))"
-    apply (simp add: add_ac sum.Sigma product_atMost_eq_Un)
-    apply (clarsimp simp add: sum_Un Sigma_interval_disjoint intro!: sum.neutral)
-    apply (metis add_diff_assoc2 add.commute add_lessD1 leD m n nat_le_linear neqE)
-    done
+    by (simp add: add_ac sum.Sigma product_atMost_eq_Un sum_Un Sigma_interval_disjoint \<section>)
   also have "\<dots> = (\<Sum>(i,j)\<in>{(i,j). i+j \<le> m+n}. (a i * x ^ i) * (b j * x ^ j))"
     by (auto simp: pairs_le_eq_Sigma sum.Sigma)
+  also have "... = (\<Sum>k\<le>m + n. \<Sum>i\<le>k. a i * x ^ i * (b (k - i) * x ^ (k - i)))"
+    by (rule sum.triangle_reindex_eq)
   also have "\<dots> = (\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
-    apply (subst sum.triangle_reindex_eq)
-    apply (auto simp: algebra_simps sum_distrib_left intro!: sum.cong)
-    apply (metis le_add_diff_inverse power_add)
-    done
+    by (auto simp: algebra_simps sum_distrib_left simp flip: power_add intro!: sum.cong)
   finally show ?thesis .
 qed
 
@@ -6109,7 +6110,7 @@
   assumes m: "\<And>i. i > m \<Longrightarrow> a i = 0"
     and n: "\<And>j. j > n \<Longrightarrow> b j = 0"
   shows "(\<Sum>i\<le>m. (a i) * x ^ i) * (\<Sum>j\<le>n. (b j) * x ^ j) =
-    (\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
+         (\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
   using polynomial_product [of m a n b x] assms
   by (simp only: of_nat_mult [symmetric] of_nat_power [symmetric]
       of_nat_eq_iff Int.int_sum [symmetric])
@@ -6148,10 +6149,10 @@
   have "(\<Sum>i=Suc j..n. a i * y^(i - j - 1)) = (\<Sum>k<n-j. a(j+k+1) * y^k)"
     if "j < n" for j :: nat
   proof -
-    have h: "bij_betw (\<lambda>i. i - (j + 1)) {Suc j..n} (lessThan (n-j))"
-      apply (auto simp: bij_betw_def inj_on_def)
-      apply (rule_tac x="x + Suc j" in image_eqI, auto)
-      done
+    have "\<And>k. k < n - j \<Longrightarrow> k \<in> (\<lambda>i. i - Suc j) ` {Suc j..n}"
+      by (rule_tac x="k + Suc j" in image_eqI, auto)
+    then have h: "bij_betw (\<lambda>i. i - (j + 1)) {Suc j..n} (lessThan (n-j))"
+      by (auto simp: bij_betw_def inj_on_def)
     then show ?thesis
       by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp)
   qed
--- a/src/HOL/Transfer.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Transfer.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -162,31 +162,33 @@
 using assms by(auto simp add: right_total_def)
 
 lemma right_total_alt_def2:
-  "right_total R \<longleftrightarrow> ((R ===> (\<longrightarrow>)) ===> (\<longrightarrow>)) All All"
-  unfolding right_total_def rel_fun_def
-  apply (rule iffI, fast)
-  apply (rule allI)
-  apply (drule_tac x="\<lambda>x. True" in spec)
-  apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec)
-  apply fast
-  done
+  "right_total R \<longleftrightarrow> ((R ===> (\<longrightarrow>)) ===> (\<longrightarrow>)) All All" (is "?lhs = ?rhs")
+proof
+  assume ?lhs then show ?rhs
+    unfolding right_total_def rel_fun_def by blast
+next
+  assume \<section>: ?rhs
+  show ?lhs
+    using \<section> [unfolded rel_fun_def, rule_format, of "\<lambda>x. True" "\<lambda>y. \<exists>x. R x y"]
+    unfolding right_total_def by blast
+qed
 
 lemma right_unique_alt_def2:
   "right_unique R \<longleftrightarrow> (R ===> R ===> (\<longrightarrow>)) (=) (=)"
   unfolding right_unique_def rel_fun_def by auto
 
 lemma bi_total_alt_def2:
-  "bi_total R \<longleftrightarrow> ((R ===> (=)) ===> (=)) All All"
-  unfolding bi_total_def rel_fun_def
-  apply (rule iffI, fast)
-  apply safe
-  apply (drule_tac x="\<lambda>x. \<exists>y. R x y" in spec)
-  apply (drule_tac x="\<lambda>y. True" in spec)
-  apply fast
-  apply (drule_tac x="\<lambda>x. True" in spec)
-  apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec)
-  apply fast
-  done
+  "bi_total R \<longleftrightarrow> ((R ===> (=)) ===> (=)) All All" (is "?lhs = ?rhs")
+proof
+  assume ?lhs then show ?rhs
+    unfolding bi_total_def rel_fun_def by blast
+next
+  assume \<section>: ?rhs
+  show ?lhs
+    using \<section> [unfolded rel_fun_def, rule_format, of "\<lambda>x. \<exists>y. R x y" "\<lambda>y. True"]
+    using \<section> [unfolded rel_fun_def, rule_format, of "\<lambda>x. True" "\<lambda>y. \<exists>x. R x y"]
+    by (auto simp: bi_total_def)
+qed
 
 lemma bi_unique_alt_def2:
   "bi_unique R \<longleftrightarrow> (R ===> R ===> (=)) (=) (=)"
@@ -194,19 +196,19 @@
 
 lemma [simp]:
   shows left_unique_conversep: "left_unique A\<inverse>\<inverse> \<longleftrightarrow> right_unique A"
-  and right_unique_conversep: "right_unique A\<inverse>\<inverse> \<longleftrightarrow> left_unique A"
-by(auto simp add: left_unique_def right_unique_def)
+    and right_unique_conversep: "right_unique A\<inverse>\<inverse> \<longleftrightarrow> left_unique A"
+  by(auto simp add: left_unique_def right_unique_def)
 
 lemma [simp]:
   shows left_total_conversep: "left_total A\<inverse>\<inverse> \<longleftrightarrow> right_total A"
-  and right_total_conversep: "right_total A\<inverse>\<inverse> \<longleftrightarrow> left_total A"
-by(simp_all add: left_total_def right_total_def)
+    and right_total_conversep: "right_total A\<inverse>\<inverse> \<longleftrightarrow> left_total A"
+  by(simp_all add: left_total_def right_total_def)
 
 lemma bi_unique_conversep [simp]: "bi_unique R\<inverse>\<inverse> = bi_unique R"
-by(auto simp add: bi_unique_def)
+  by(auto simp add: bi_unique_def)
 
 lemma bi_total_conversep [simp]: "bi_total R\<inverse>\<inverse> = bi_total R"
-by(auto simp add: bi_total_def)
+  by(auto simp add: bi_total_def)
 
 lemma right_unique_alt_def: "right_unique R = (conversep R OO R \<le> (=))" unfolding right_unique_def by blast
 lemma left_unique_alt_def: "left_unique R = (R OO (conversep R) \<le> (=))" unfolding left_unique_def by blast
@@ -230,21 +232,21 @@
 
 
 lemma is_equality_lemma: "(\<And>R. is_equality R \<Longrightarrow> PROP (P R)) \<equiv> PROP (P (=))"
-  apply (unfold is_equality_def)
-  apply (rule equal_intr_rule)
-   apply (drule meta_spec)
-   apply (erule meta_mp)
-   apply (rule refl)
-  apply simp
-  done
+  unfolding is_equality_def
+proof (rule equal_intr_rule)
+  show "(\<And>R. R = (=) \<Longrightarrow> PROP P R) \<Longrightarrow> PROP P (=)"
+    apply (drule meta_spec)
+    apply (erule meta_mp [OF _ refl])
+    done
+qed simp
 
 lemma Domainp_lemma: "(\<And>R. Domainp T = R \<Longrightarrow> PROP (P R)) \<equiv> PROP (P (Domainp T))"
-  apply (rule equal_intr_rule)
-   apply (drule meta_spec)
-   apply (erule meta_mp)
-   apply (rule refl)
-  apply simp
-  done
+proof (rule equal_intr_rule)
+  show "(\<And>R. Domainp T = R \<Longrightarrow> PROP P R) \<Longrightarrow> PROP P (Domainp T)"
+    apply (drule meta_spec)
+    apply (erule meta_mp [OF _ refl])
+    done
+qed simp
 
 ML_file \<open>Tools/Transfer/transfer.ML\<close>
 declare refl [transfer_rule]
@@ -266,14 +268,21 @@
 
 lemma Domainp_pred_fun_eq[relator_domain]:
   assumes "left_unique T"
-  shows "Domainp (T ===> S) = pred_fun (Domainp T) (Domainp S)"
-  using assms unfolding rel_fun_def Domainp_iff[abs_def] left_unique_def fun_eq_iff pred_fun_def
-  apply safe
-   apply blast
-  apply (subst all_comm)
-  apply (rule choice)
-  apply blast
-  done
+  shows "Domainp (T ===> S) = pred_fun (Domainp T) (Domainp S)"   (is "?lhs = ?rhs")
+proof (intro ext iffI)
+  fix x
+  assume "?lhs x"
+  then show "?rhs x"
+    using assms unfolding rel_fun_def pred_fun_def by blast
+next
+  fix x
+  assume "?rhs x"
+  then have "\<exists>g. \<forall>y xa. T xa y \<longrightarrow> S (x xa) (g y)"
+    using assms unfolding Domainp_iff left_unique_def  pred_fun_def
+    by (intro choice) blast
+  then show "?lhs x"
+    by blast
+qed
 
 text \<open>Properties are preserved by relation composition.\<close>
 
@@ -295,10 +304,10 @@
   unfolding right_unique_def OO_def by fast
 
 lemma left_total_OO: "left_total R \<Longrightarrow> left_total S \<Longrightarrow> left_total (R OO S)"
-unfolding left_total_def OO_def by fast
+  unfolding left_total_def OO_def by fast
 
 lemma left_unique_OO: "left_unique R \<Longrightarrow> left_unique S \<Longrightarrow> left_unique (R OO S)"
-unfolding left_unique_def OO_def by blast
+  unfolding left_unique_def OO_def by blast
 
 
 subsection \<open>Properties of relators\<close>
@@ -322,18 +331,22 @@
   unfolding bi_unique_def by simp
 
 lemma left_total_fun[transfer_rule]:
-  "\<lbrakk>left_unique A; left_total B\<rbrakk> \<Longrightarrow> left_total (A ===> B)"
-  unfolding left_total_def rel_fun_def
-  apply (rule allI, rename_tac f)
-  apply (rule_tac x="\<lambda>y. SOME z. B (f (THE x. A x y)) z" in exI)
-  apply clarify
-  apply (subgoal_tac "(THE x. A x y) = x", simp)
-  apply (rule someI_ex)
-  apply (simp)
-  apply (rule the_equality)
-  apply assumption
-  apply (simp add: left_unique_def)
-  done
+  assumes "left_unique A" "left_total B"
+  shows "left_total (A ===> B)"
+  unfolding left_total_def 
+proof
+  fix f
+  show "Ex ((A ===> B) f)"
+    unfolding rel_fun_def
+  proof (intro exI strip)
+    fix x y
+    assume A: "A x y"
+    have "(THE x. A x y) = x"
+      using A assms by (simp add: left_unique_def the_equality)
+    then show "B (f x) (SOME z. B (f (THE x. A x y)) z)"
+      using assms by (force simp: left_total_def intro: someI_ex)
+  qed
+qed
 
 lemma left_unique_fun[transfer_rule]:
   "\<lbrakk>left_total A; left_unique B\<rbrakk> \<Longrightarrow> left_unique (A ===> B)"
@@ -341,18 +354,22 @@
   by (clarify, rule ext, fast)
 
 lemma right_total_fun [transfer_rule]:
-  "\<lbrakk>right_unique A; right_total B\<rbrakk> \<Longrightarrow> right_total (A ===> B)"
-  unfolding right_total_def rel_fun_def
-  apply (rule allI, rename_tac g)
-  apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI)
-  apply clarify
-  apply (subgoal_tac "(THE y. A x y) = y", simp)
-  apply (rule someI_ex)
-  apply (simp)
-  apply (rule the_equality)
-  apply assumption
-  apply (simp add: right_unique_def)
-  done
+  assumes "right_unique A" "right_total B"
+  shows "right_total (A ===> B)"
+  unfolding right_total_def 
+proof
+  fix g
+  show "\<exists>x. (A ===> B) x g"
+    unfolding rel_fun_def
+  proof (intro exI strip)
+    fix x y
+    assume A: "A x y"
+    have "(THE y. A x y) = y"
+      using A assms by (simp add: right_unique_def the_equality)
+    then show "B (SOME z. B z (g (THE y. A x y))) (g y)"
+      using assms by (force simp: right_total_def intro: someI_ex)
+  qed
+qed
 
 lemma right_unique_fun [transfer_rule]:
   "\<lbrakk>right_total A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A ===> B)"
@@ -435,8 +452,8 @@
   assumes "right_total B"
   assumes "bi_unique A"
   shows "((A ===> B ===> (=)) ===> implies) left_unique left_unique"
-using assms unfolding left_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def
-by metis
+  using assms unfolding left_unique_def right_total_def bi_unique_def rel_fun_def
+  by metis
 
 lemma eq_transfer [transfer_rule]:
   assumes "bi_unique A"
@@ -446,14 +463,14 @@
 lemma right_total_Ex_transfer[transfer_rule]:
   assumes "right_total A"
   shows "((A ===> (=)) ===> (=)) (Bex (Collect (Domainp A))) Ex"
-using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff[abs_def]
-by fast
+  using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff
+  by fast
 
 lemma right_total_All_transfer[transfer_rule]:
   assumes "right_total A"
   shows "((A ===> (=)) ===> (=)) (Ball (Collect (Domainp A))) All"
-using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff[abs_def]
-by fast
+  using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff
+  by fast
 
 context
   includes lifting_syntax
@@ -480,7 +497,7 @@
 lemma Ex1_parametric [transfer_rule]:
   assumes [transfer_rule]: "bi_unique A" "bi_total A"
   shows "((A ===> (=)) ===> (=)) Ex1 Ex1"
-unfolding Ex1_def[abs_def] by transfer_prover
+unfolding Ex1_def by transfer_prover
 
 declare If_transfer [transfer_rule]
 
@@ -498,7 +515,7 @@
 lemma fun_upd_transfer [transfer_rule]:
   assumes [transfer_rule]: "bi_unique A"
   shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd"
-  unfolding fun_upd_def [abs_def] by transfer_prover
+  unfolding fun_upd_def  by transfer_prover
 
 lemma case_nat_transfer [transfer_rule]:
   "(A ===> ((=) ===> A) ===> (=) ===> A) case_nat case_nat"
@@ -517,18 +534,18 @@
   assumes [transfer_rule]: "(A ===> A ===> (=)) (\<le>) (\<le>)"
   assumes [transfer_rule]: "(B ===> B ===> (=)) (\<le>) (\<le>)"
   shows "((A ===> B) ===> (=)) mono mono"
-unfolding mono_def[abs_def] by transfer_prover
+unfolding mono_def by transfer_prover
 
 lemma right_total_relcompp_transfer[transfer_rule]:
   assumes [transfer_rule]: "right_total B"
   shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=))
     (\<lambda>R S x z. \<exists>y\<in>Collect (Domainp B). R x y \<and> S y z) (OO)"
-unfolding OO_def[abs_def] by transfer_prover
+unfolding OO_def by transfer_prover
 
 lemma relcompp_transfer[transfer_rule]:
   assumes [transfer_rule]: "bi_total B"
   shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=)) (OO) (OO)"
-unfolding OO_def[abs_def] by transfer_prover
+unfolding OO_def by transfer_prover
 
 lemma right_total_Domainp_transfer[transfer_rule]:
   assumes [transfer_rule]: "right_total B"
@@ -538,7 +555,7 @@
 lemma Domainp_transfer[transfer_rule]:
   assumes [transfer_rule]: "bi_total B"
   shows "((A ===> B ===> (=)) ===> A ===> (=)) Domainp Domainp"
-unfolding Domainp_iff[abs_def] by transfer_prover
+unfolding Domainp_iff by transfer_prover
 
 lemma reflp_transfer[transfer_rule]:
   "bi_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> (=)) reflp reflp"
@@ -546,38 +563,38 @@
   "right_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> implies) reflp reflp"
   "bi_total A \<Longrightarrow> ((A ===> A ===> rev_implies) ===> rev_implies) reflp reflp"
   "bi_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> rev_implies) reflp reflp"
-unfolding reflp_def[abs_def] rev_implies_def bi_total_def right_total_def rel_fun_def
+unfolding reflp_def rev_implies_def bi_total_def right_total_def rel_fun_def
 by fast+
 
 lemma right_unique_transfer [transfer_rule]:
   "\<lbrakk> right_total A; right_total B; bi_unique B \<rbrakk>
   \<Longrightarrow> ((A ===> B ===> (=)) ===> implies) right_unique right_unique"
-unfolding right_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def
+unfolding right_unique_def right_total_def bi_unique_def rel_fun_def
 by metis
 
 lemma left_total_parametric [transfer_rule]:
   assumes [transfer_rule]: "bi_total A" "bi_total B"
   shows "((A ===> B ===> (=)) ===> (=)) left_total left_total"
-unfolding left_total_def[abs_def] by transfer_prover
+unfolding left_total_def by transfer_prover
 
 lemma right_total_parametric [transfer_rule]:
   assumes [transfer_rule]: "bi_total A" "bi_total B"
   shows "((A ===> B ===> (=)) ===> (=)) right_total right_total"
-unfolding right_total_def[abs_def] by transfer_prover
+unfolding right_total_def by transfer_prover
 
 lemma left_unique_parametric [transfer_rule]:
   assumes [transfer_rule]: "bi_unique A" "bi_total A" "bi_total B"
   shows "((A ===> B ===> (=)) ===> (=)) left_unique left_unique"
-unfolding left_unique_def[abs_def] by transfer_prover
+unfolding left_unique_def by transfer_prover
 
 lemma prod_pred_parametric [transfer_rule]:
   "((A ===> (=)) ===> (B ===> (=)) ===> rel_prod A B ===> (=)) pred_prod pred_prod"
-unfolding prod.pred_set[abs_def] Basic_BNFs.fsts_def Basic_BNFs.snds_def fstsp.simps sndsp.simps 
+unfolding prod.pred_set Basic_BNFs.fsts_def Basic_BNFs.snds_def fstsp.simps sndsp.simps 
 by simp transfer_prover
 
 lemma apfst_parametric [transfer_rule]:
   "((A ===> B) ===> rel_prod A C ===> rel_prod B C) apfst apfst"
-unfolding apfst_def[abs_def] by transfer_prover
+unfolding apfst_def by transfer_prover
 
 lemma rel_fun_eq_eq_onp: "((=) ===> eq_onp P) = eq_onp (\<lambda>f. \<forall>x. P(f x))"
 unfolding eq_onp_def rel_fun_def by auto
@@ -589,7 +606,7 @@
 lemma eq_onp_transfer [transfer_rule]:
   assumes [transfer_rule]: "bi_unique A"
   shows "((A ===> (=)) ===> A ===> A ===> (=)) eq_onp eq_onp"
-unfolding eq_onp_def[abs_def] by transfer_prover
+unfolding eq_onp_def by transfer_prover
 
 lemma rtranclp_parametric [transfer_rule]:
   assumes "bi_unique A" "bi_total A"
@@ -633,11 +650,11 @@
 lemma right_unique_parametric [transfer_rule]:
   assumes [transfer_rule]: "bi_total A" "bi_unique B" "bi_total B"
   shows "((A ===> B ===> (=)) ===> (=)) right_unique right_unique"
-unfolding right_unique_def[abs_def] by transfer_prover
+  unfolding right_unique_def by transfer_prover
 
 lemma map_fun_parametric [transfer_rule]:
   "((A ===> B) ===> (C ===> D) ===> (B ===> C) ===> A ===> D) map_fun map_fun"
-unfolding map_fun_def[abs_def] by transfer_prover
+  unfolding map_fun_def by transfer_prover
 
 end
 
@@ -652,14 +669,14 @@
   \<open>((\<longleftrightarrow>) ===> (\<cong>)) of_bool of_bool\<close>
     if [transfer_rule]: \<open>0 \<cong> 0\<close> \<open>1 \<cong> 1\<close>
     for R :: \<open>'a::zero_neq_one \<Rightarrow> 'b::zero_neq_one \<Rightarrow> bool\<close>  (infix \<open>\<cong>\<close> 50)
-  by (unfold of_bool_def [abs_def]) transfer_prover
+  unfolding of_bool_def by transfer_prover
 
 lemma transfer_rule_of_nat:
   "((=) ===> (\<cong>)) of_nat of_nat"
     if [transfer_rule]: \<open>0 \<cong> 0\<close> \<open>1 \<cong> 1\<close>
     \<open>((\<cong>) ===> (\<cong>) ===> (\<cong>)) (+) (+)\<close>
   for R :: \<open>'a::semiring_1 \<Rightarrow> 'b::semiring_1 \<Rightarrow> bool\<close>  (infix \<open>\<cong>\<close> 50)
-  by (unfold of_nat_def [abs_def]) transfer_prover
+  unfolding of_nat_def by transfer_prover
 
 end
 
--- a/src/HOL/Transitive_Closure.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/HOL/Transitive_Closure.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -87,9 +87,7 @@
 
 lemma r_into_rtrancl [intro]: "\<And>p. p \<in> r \<Longrightarrow> p \<in> r\<^sup>*"
   \<comment> \<open>\<open>rtrancl\<close> of \<open>r\<close> contains \<open>r\<close>\<close>
-  apply (simp only: split_tupled_all)
-  apply (erule rtrancl_refl [THEN rtrancl_into_rtrancl])
-  done
+  by (simp add: split_tupled_all rtrancl_refl [THEN rtrancl_into_rtrancl])
 
 lemma r_into_rtranclp [intro]: "r x y \<Longrightarrow> r\<^sup>*\<^sup>* x y"
   \<comment> \<open>\<open>rtrancl\<close> of \<open>r\<close> contains \<open>r\<close>\<close>
@@ -97,10 +95,11 @@
 
 lemma rtranclp_mono: "r \<le> s \<Longrightarrow> r\<^sup>*\<^sup>* \<le> s\<^sup>*\<^sup>*"
   \<comment> \<open>monotonicity of \<open>rtrancl\<close>\<close>
-  apply (rule predicate2I)
-  apply (erule rtranclp.induct)
-   apply (rule_tac [2] rtranclp.rtrancl_into_rtrancl, blast+)
-  done
+proof (rule predicate2I)
+  show "s\<^sup>*\<^sup>* x y" if "r \<le> s" "r\<^sup>*\<^sup>* x y" for x y
+    using \<open>r\<^sup>*\<^sup>* x y\<close> \<open>r \<le> s\<close>
+    by (induction rule: rtranclp.induct) (blast intro: rtranclp.rtrancl_into_rtrancl)+
+qed
 
 lemma mono_rtranclp[mono]: "(\<And>a b. x a b \<longrightarrow> y a b) \<Longrightarrow> x\<^sup>*\<^sup>* a b \<longrightarrow> y\<^sup>*\<^sup>* a b"
    using rtranclp_mono[of x y] by auto
@@ -164,9 +163,7 @@
 qed
 
 lemma rtrancl_Int_subset: "Id \<subseteq> s \<Longrightarrow> (r\<^sup>* \<inter> s) O r \<subseteq> s \<Longrightarrow> r\<^sup>* \<subseteq> s"
-  apply clarify
-  apply (erule rtrancl_induct, auto)
-  done
+  by (fastforce elim: rtrancl_induct)
 
 lemma converse_rtranclp_into_rtranclp: "r a b \<Longrightarrow> r\<^sup>*\<^sup>* b c \<Longrightarrow> r\<^sup>*\<^sup>* a c"
   by (rule rtranclp_trans) iprover+
@@ -176,27 +173,23 @@
 text \<open>\<^medskip> More \<^term>\<open>r\<^sup>*\<close> equations and inclusions.\<close>
 
 lemma rtranclp_idemp [simp]: "(r\<^sup>*\<^sup>*)\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*"
-  apply (auto intro!: order_antisym)
-  apply (erule rtranclp_induct)
-   apply (rule rtranclp.rtrancl_refl)
-  apply (blast intro: rtranclp_trans)
-  done
+proof -
+  have "r\<^sup>*\<^sup>*\<^sup>*\<^sup>* x y \<Longrightarrow> r\<^sup>*\<^sup>* x y" for x y
+    by (induction rule: rtranclp_induct) (blast intro: rtranclp_trans)+
+  then show ?thesis
+    by (auto intro!: order_antisym)
+qed
 
 lemmas rtrancl_idemp [simp] = rtranclp_idemp [to_set]
 
 lemma rtrancl_idemp_self_comp [simp]: "R\<^sup>* O R\<^sup>* = R\<^sup>*"
-  apply (rule set_eqI)
-  apply (simp only: split_tupled_all)
-  apply (blast intro: rtrancl_trans)
-  done
+  by (force intro: rtrancl_trans)
 
 lemma rtrancl_subset_rtrancl: "r \<subseteq> s\<^sup>* \<Longrightarrow> r\<^sup>* \<subseteq> s\<^sup>*"
-by (drule rtrancl_mono, simp)
+  by (drule rtrancl_mono, simp)
 
 lemma rtranclp_subset: "R \<le> S \<Longrightarrow> S \<le> R\<^sup>*\<^sup>* \<Longrightarrow> S\<^sup>*\<^sup>* = R\<^sup>*\<^sup>*"
-  apply (drule rtranclp_mono)
-  apply (drule rtranclp_mono, simp)
-  done
+  by (fastforce dest: rtranclp_mono)
 
 lemmas rtrancl_subset = rtranclp_subset [to_set]
 
@@ -319,11 +312,15 @@
 
 subsection \<open>Transitive closure\<close>
 
-lemma trancl_mono: "\<And>p. p \<in> r\<^sup>+ \<Longrightarrow> r \<subseteq> s \<Longrightarrow> p \<in> s\<^sup>+"
-  apply (simp add: split_tupled_all)
-  apply (erule trancl.induct)
-   apply (iprover dest: subsetD)+
-  done
+lemma trancl_mono:
+  assumes "p \<in> r\<^sup>+" "r \<subseteq> s"
+  shows "p \<in> s\<^sup>+"
+proof -
+  have "\<lbrakk>(a, b) \<in> r\<^sup>+; r \<subseteq> s\<rbrakk> \<Longrightarrow> (a, b) \<in> s\<^sup>+" for a b
+    by (induction rule: trancl.induct) (iprover dest: subsetD)+
+  with assms show ?thesis
+    by (cases p) force
+qed
 
 lemma r_into_trancl': "\<And>p. p \<in> r \<Longrightarrow> p \<in> r\<^sup>+"
   by (simp only: split_tupled_all) (erule r_into_trancl)
@@ -342,12 +339,19 @@
 
 lemmas rtrancl_into_trancl1 = rtranclp_into_tranclp1 [to_set]
 
-lemma rtranclp_into_tranclp2: "r a b \<Longrightarrow> r\<^sup>*\<^sup>* b c \<Longrightarrow> r\<^sup>+\<^sup>+ a c"
+lemma rtranclp_into_tranclp2:
+  assumes "r a b" "r\<^sup>*\<^sup>* b c" shows "r\<^sup>+\<^sup>+ a c"
   \<comment> \<open>intro rule from \<open>r\<close> and \<open>rtrancl\<close>\<close>
-  apply (erule rtranclp.cases, iprover)
-  apply (rule rtranclp_trans [THEN rtranclp_into_tranclp1])
-    apply (simp | rule r_into_rtranclp)+
-  done
+  using \<open>r\<^sup>*\<^sup>* b c\<close>
+proof (cases rule: rtranclp.cases)
+  case rtrancl_refl
+  with assms show ?thesis
+    by iprover
+next
+  case rtrancl_into_rtrancl
+  with assms show ?thesis
+    by (auto intro: rtranclp_trans [THEN rtranclp_into_tranclp1])
+qed
 
 lemmas rtrancl_into_trancl2 = rtranclp_into_tranclp2 [to_set]
 
@@ -384,9 +388,7 @@
   using assms by cases simp_all
 
 lemma trancl_Int_subset: "r \<subseteq> s \<Longrightarrow> (r\<^sup>+ \<inter> s) O r \<subseteq> s \<Longrightarrow> r\<^sup>+ \<subseteq> s"
-  apply clarify
-  apply (erule trancl_induct, auto)
-  done
+  by (fastforce simp add: elim: trancl_induct)
 
 lemma trancl_unfold: "r\<^sup>+ = r \<union> r\<^sup>+ O r"
   by (auto intro: trancl_into_trancl elim: tranclE)
@@ -418,10 +420,7 @@
   using assms(2,1) by induct iprover+
 
 lemma trancl_id [simp]: "trans r \<Longrightarrow> r\<^sup>+ = r"
-  apply auto
-  apply (erule trancl_induct, assumption)
-  apply (unfold trans_def, blast)
-  done
+  unfolding trans_def by (fastforce simp add: elim: trancl_induct)
 
 lemma rtranclp_tranclp_tranclp:
   assumes "r\<^sup>*\<^sup>* x y"
@@ -435,19 +434,30 @@
 
 lemmas trancl_into_trancl2 = tranclp_into_tranclp2 [to_set]
 
-lemma tranclp_converseI: "(r\<^sup>+\<^sup>+)\<inverse>\<inverse> x y \<Longrightarrow> (r\<inverse>\<inverse>)\<^sup>+\<^sup>+ x y"
-  apply (drule conversepD)
-  apply (erule tranclp_induct)
-   apply (iprover intro: conversepI tranclp_trans)+
-  done
+lemma tranclp_converseI:
+  assumes "(r\<^sup>+\<^sup>+)\<inverse>\<inverse> x y" shows "(r\<inverse>\<inverse>)\<^sup>+\<^sup>+ x y"
+  using conversepD [OF assms]
+proof (induction rule: tranclp_induct)
+  case (base y)
+  then show ?case 
+    by (iprover intro: conversepI)
+next
+  case (step y z)
+  then show ?case
+    by (iprover intro: conversepI tranclp_trans)
+qed
 
 lemmas trancl_converseI = tranclp_converseI [to_set]
 
-lemma tranclp_converseD: "(r\<inverse>\<inverse>)\<^sup>+\<^sup>+ x y \<Longrightarrow> (r\<^sup>+\<^sup>+)\<inverse>\<inverse> x y"
-  apply (rule conversepI)
-  apply (erule tranclp_induct)
-   apply (iprover dest: conversepD intro: tranclp_trans)+
-  done
+lemma tranclp_converseD:
+  assumes "(r\<inverse>\<inverse>)\<^sup>+\<^sup>+ x y" shows "(r\<^sup>+\<^sup>+)\<inverse>\<inverse> x y"
+proof -
+  have "r\<^sup>+\<^sup>+ y x"
+    using assms
+    by (induction rule: tranclp_induct) (iprover dest: conversepD intro: tranclp_trans)+
+  then show ?thesis
+    by (rule conversepI)
+qed
 
 lemmas trancl_converseD = tranclp_converseD [to_set]
 
@@ -463,17 +473,21 @@
   assumes major: "r\<^sup>+\<^sup>+ a b"
     and cases: "\<And>y. r y b \<Longrightarrow> P y" "\<And>y z. r y z \<Longrightarrow> r\<^sup>+\<^sup>+ z b \<Longrightarrow> P z \<Longrightarrow> P y"
   shows "P a"
-  apply (rule tranclp_induct [OF tranclp_converseI, OF conversepI, OF major])
-   apply (blast intro: cases)
-  apply (blast intro: assms dest!: tranclp_converseD)
-  done
+proof -
+  have "r\<inverse>\<inverse>\<^sup>+\<^sup>+ b a"
+    by (intro tranclp_converseI conversepI major)
+  then show ?thesis
+    by (induction rule: tranclp_induct) (blast intro: cases dest: tranclp_converseD)+
+qed
 
 lemmas converse_trancl_induct = converse_tranclp_induct [to_set]
 
 lemma tranclpD: "R\<^sup>+\<^sup>+ x y \<Longrightarrow> \<exists>z. R x z \<and> R\<^sup>*\<^sup>* z y"
-  apply (erule converse_tranclp_induct, auto)
-  apply (blast intro: rtranclp_trans)
-  done
+proof (induction rule: converse_tranclp_induct)
+  case (step u v)
+  then show ?case
+    by (blast intro: rtranclp_trans)
+qed auto
 
 lemmas tranclD = tranclpD [to_set]
 
@@ -492,7 +506,7 @@
       by iprover
   next
     case rtrancl_into_rtrancl
-    from this have "tranclp r y z"
+    then have "tranclp r y z"
       by (iprover intro: rtranclp_into_tranclp1)
     with \<open>r x y\<close> step show P
       by iprover
@@ -513,17 +527,15 @@
 lemma trancl_subset_Sigma_aux: "(a, b) \<in> r\<^sup>* \<Longrightarrow> r \<subseteq> A \<times> A \<Longrightarrow> a = b \<or> a \<in> A"
   by (induct rule: rtrancl_induct) auto
 
-lemma trancl_subset_Sigma: "r \<subseteq> A \<times> A \<Longrightarrow> r\<^sup>+ \<subseteq> A \<times> A"
-  apply (clarsimp simp:)
-  apply (erule tranclE)
-   apply (blast dest!: trancl_into_rtrancl trancl_subset_Sigma_aux)+
-  done
+lemma trancl_subset_Sigma:
+  assumes "r \<subseteq> A \<times> A" shows "r\<^sup>+ \<subseteq> A \<times> A"
+proof (rule trancl_Int_subset [OF assms])
+  show "(r\<^sup>+ \<inter> A \<times> A) O r \<subseteq> A \<times> A"
+    using assms by auto
+qed
 
 lemma reflclp_tranclp [simp]: "(r\<^sup>+\<^sup>+)\<^sup>=\<^sup>= = r\<^sup>*\<^sup>*"
-  apply (safe intro!: order_antisym)
-   apply (erule tranclp_into_rtranclp)
-  apply (blast elim: rtranclp.cases dest: rtranclp_into_tranclp1)
-  done
+  by (fast elim: rtranclp.cases tranclp_into_rtranclp dest: rtranclp_into_tranclp1)
 
 lemmas reflcl_trancl [simp] = reflclp_tranclp [to_set]
 
@@ -629,19 +641,14 @@
 qed
 
 lemma trancl_subset_Field2: "r\<^sup>+ \<subseteq> Field r \<times> Field r"
-  apply clarify
-  apply (erule trancl_induct)
-   apply (auto simp: Field_def)
-  done
+  by (rule trancl_Int_subset) (auto simp: Field_def)
 
 lemma finite_trancl[simp]: "finite (r\<^sup>+) = finite r"
 proof
   show "finite (r\<^sup>+) \<Longrightarrow> finite r"
     by (blast intro: r_into_trancl' finite_subset)
   show "finite r \<Longrightarrow> finite (r\<^sup>+)"
-   apply (rule trancl_subset_Field2 [THEN finite_subset])
-   apply (auto simp: finite_Field)
-  done
+    by (auto simp: finite_Field trancl_subset_Field2 [THEN finite_subset])
 qed
 
 lemma finite_rtrancl_Image[simp]: assumes "finite R" "finite A" shows "finite (R\<^sup>* `` A)"
@@ -665,9 +672,7 @@
 next
   case (step y z)
   with xz \<open>single_valued r\<close> show ?case
-    apply (auto simp: elim: converse_rtranclE dest: single_valuedD)
-    apply (blast intro: rtrancl_trans)
-    done
+    by (auto elim: converse_rtranclE dest: single_valuedD intro: rtrancl_trans)
 qed
 
 lemma r_r_into_trancl: "(a, b) \<in> R \<Longrightarrow> (b, c) \<in> R \<Longrightarrow> (a, c) \<in> R\<^sup>+"
@@ -676,12 +681,14 @@
 lemma trancl_into_trancl: "(a, b) \<in> r\<^sup>+ \<Longrightarrow> (b, c) \<in> r \<Longrightarrow> (a, c) \<in> r\<^sup>+"
   by (induct rule: trancl_induct) (fast intro: r_r_into_trancl trancl_trans)+
 
-lemma tranclp_rtranclp_tranclp: "r\<^sup>+\<^sup>+ a b \<Longrightarrow> r\<^sup>*\<^sup>* b c \<Longrightarrow> r\<^sup>+\<^sup>+ a c"
-  apply (drule tranclpD)
-  apply (elim exE conjE)
-  apply (drule rtranclp_trans, assumption)
-  apply (drule (2) rtranclp_into_tranclp2)
-  done
+lemma tranclp_rtranclp_tranclp:
+  assumes "r\<^sup>+\<^sup>+ a b" "r\<^sup>*\<^sup>* b c" shows "r\<^sup>+\<^sup>+ a c"
+proof -
+  obtain z where "r a z" "r\<^sup>*\<^sup>* z c"
+    using assms by (iprover dest: tranclpD rtranclp_trans)
+  then show ?thesis
+    by (blast dest: rtranclp_into_tranclp2)
+qed
 
 lemma rtranclp_conversep: "r\<inverse>\<inverse>\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*\<inverse>\<inverse>"
   by(auto simp add: fun_eq_iff intro: rtranclp_converseI rtranclp_converseD)
@@ -717,13 +724,13 @@
 
 lemma symclpI [simp, intro?]:
   shows symclpI1: "r x y \<Longrightarrow> symclp r x y"
-  and symclpI2: "r y x \<Longrightarrow> symclp r x y"
-by(simp_all add: symclp_def)
+    and symclpI2: "r y x \<Longrightarrow> symclp r x y"
+  by(simp_all add: symclp_def)
 
 lemma symclpE [consumes 1, cases pred]:
   assumes "symclp r x y"
   obtains (base) "r x y" | (sym) "r y x"
-using assms by(auto simp add: symclp_def)
+  using assms by(auto simp add: symclp_def)
 
 lemma symclp_pointfree: "symclp r = sup r r\<inverse>\<inverse>"
   by(auto simp add: symclp_def fun_eq_iff)
@@ -962,12 +969,10 @@
 
 lemma trancl_power: "p \<in> R\<^sup>+ \<longleftrightarrow> (\<exists>n > 0. p \<in> R ^^ n)"
 proof -
-  have "((a, b) \<in> R\<^sup>+) = (\<exists>n>0. (a, b) \<in> R ^^ n)" for a b
+  have "(a, b) \<in> R\<^sup>+ \<longleftrightarrow> (\<exists>n>0. (a, b) \<in> R ^^ n)" for a b
   proof safe
     show "(a, b) \<in> R\<^sup>+ \<Longrightarrow> \<exists>n>0. (a, b) \<in> R ^^ n"
-      apply (drule tranclD2)
-      apply (fastforce simp: rtrancl_is_UN_relpow relcomp_unfold)
-      done
+      by (fastforce simp: rtrancl_is_UN_relpow relcomp_unfold dest: tranclD2)
     show "(a, b) \<in> R\<^sup>+" if "n > 0" "(a, b) \<in> R ^^ n" for n
     proof (cases n)
       case (Suc m)
@@ -1117,17 +1122,23 @@
   fixes R :: "('a \<times> 'a) set"
   assumes "finite R"
   shows "R^^k \<subseteq> (\<Union>n\<in>{n. n \<le> card R}. R^^n)"
-  apply (cases k, force)
-  apply (use relpow_finite_bounded1[OF assms, of k] in auto)
-  done
+proof (cases k)
+  case (Suc k')
+  then show ?thesis
+    using relpow_finite_bounded1[OF assms, of k] by auto
+qed force
 
 lemma rtrancl_finite_eq_relpow: "finite R \<Longrightarrow> R\<^sup>* = (\<Union>n\<in>{n. n \<le> card R}. R^^n)"
   by (fastforce simp: rtrancl_power dest: relpow_finite_bounded)
 
-lemma trancl_finite_eq_relpow: "finite R \<Longrightarrow> R\<^sup>+ = (\<Union>n\<in>{n. 0 < n \<and> n \<le> card R}. R^^n)"
-  apply (auto simp: trancl_power)
-  apply (auto dest: relpow_finite_bounded1)
-  done
+lemma trancl_finite_eq_relpow:
+  assumes "finite R" shows "R\<^sup>+ = (\<Union>n\<in>{n. 0 < n \<and> n \<le> card R}. R^^n)"
+proof -
+  have "\<And>a b n. \<lbrakk>0 < n; (a, b) \<in> R ^^ n\<rbrakk> \<Longrightarrow> \<exists>x>0. x \<le> card R \<and> (a, b) \<in> R ^^ x"
+    using assms by (auto dest: relpow_finite_bounded1)
+  then show ?thesis
+    by (auto simp: trancl_power)
+qed
 
 lemma finite_relcomp[simp,intro]:
   assumes "finite R" and "finite S"
@@ -1189,7 +1200,7 @@
     show ?thesis
     proof (cases "i = 1")
       case True
-      from this \<open>(a, b) \<in> R ^^ i\<close> show ?thesis
+      with \<open>(a, b) \<in> R ^^ i\<close> show ?thesis
         by (auto simp: ntrancl_def)
     next
       case False
--- a/src/Pure/Admin/afp.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/afp.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -166,7 +166,7 @@
   /* dependency graph */
 
   private def sessions_deps(entry: AFP.Entry): List[String] =
-    entry.sessions.flatMap(sessions_structure.imports_graph.imm_preds(_)).distinct.sorted
+    entry.sessions.flatMap(sessions_structure.imports_graph.imm_preds).distinct.sorted
 
   lazy val entries_graph: Graph[String, Unit] =
   {
@@ -214,7 +214,7 @@
       case 1 | 2 =>
         val graph = sessions_structure.build_graph.restrict(sessions.toSet)
         val force_part1 =
-          graph.all_preds(graph.all_succs(AFP.force_partition1.filter(graph.defined(_)))).toSet
+          graph.all_preds(graph.all_succs(AFP.force_partition1.filter(graph.defined))).toSet
         val (part1, part2) = graph.keys.partition(a => force_part1(a) || graph.is_isolated(a))
         if (n == 1) part1 else part2
       case _ => error("Bad AFP partition: " + n + " (should be 0, 1, 2)")
--- a/src/Pure/Admin/build_doc.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/build_doc.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -16,7 +16,7 @@
 
   def build_doc(
     options: Options,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     all_docs: Boolean = false,
     max_jobs: Int = 1,
     docs: List[String] = Nil): Int =
--- a/src/Pure/Admin/build_fonts.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/build_fonts.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -216,7 +216,7 @@
     target_prefix: String = "Isabelle",
     target_version: String = "",
     target_dir: Path = default_target_dir,
-    progress: Progress = No_Progress)
+    progress: Progress = new Progress)
   {
     progress.echo("Directory " + target_dir)
     hinting.foreach(hinted => Isabelle_System.mkdirs(target_dir + hinted_path(hinted)))
@@ -284,7 +284,7 @@
       val domain =
         (for ((name, index) <- targets if index == 0)
           yield Fontforge.font_domain(target_dir + hinted_path(false) + name))
-        .flatten.toSet.toList.sorted
+        .flatten.distinct.sorted
 
       Fontforge.execute(
         Fontforge.commands(
--- a/src/Pure/Admin/build_history.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/build_history.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -18,7 +18,6 @@
 
   val engine = "build_history"
   val log_prefix = engine + "_"
-  val META_INFO_MARKER = "\fmeta_info = "
 
 
   /* augment settings */
@@ -107,7 +106,7 @@
     options: Options,
     root: Path,
     user_home: Path = default_user_home,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     rev: String = default_rev,
     afp_rev: Option[String] = None,
     afp_partition: Int = 0,
@@ -282,8 +281,7 @@
                 val theory_timings =
                   try {
                     store.read_theory_timings(db, session_name).map(ps =>
-                      Build_Log.Log_File.print_props(Build_Log.THEORY_TIMING_MARKER,
-                        (Build_Log.SESSION_NAME, session_name) :: ps))
+                      Protocol.Theory_Timing_Marker((Build_Log.SESSION_NAME, session_name) :: ps))
                   }
                   catch { case ERROR(_) => Nil }
 
@@ -357,10 +355,10 @@
       build_out_progress.echo("Writing log file " + log_path.ext("xz") + " ...")
       File.write_xz(log_path.ext("xz"),
         terminate_lines(
-          Build_Log.Log_File.print_props(META_INFO_MARKER, meta_info) :: build_result.out_lines :::
+          Protocol.Meta_Info_Marker(meta_info) :: build_result.out_lines :::
           session_build_info :::
-          ml_statistics.map(Build_Log.Log_File.print_props(Build_Log.ML_STATISTICS_MARKER, _)) :::
-          session_errors.map(Build_Log.Log_File.print_props(Build_Log.ERROR_MESSAGE_MARKER, _)) :::
+          ml_statistics.map(Protocol.ML_Statistics_Marker.apply) :::
+          session_errors.map(Protocol.Error_Message_Marker.apply) :::
           heap_sizes), XZ.options(6))
 
 
@@ -395,7 +393,7 @@
 
   def main(args: Array[String])
   {
-    Command_Line.tool0 {
+    Command_Line.tool {
       var afp_rev: Option[String] = None
       var multicore_base = false
       var components_base: Path = Components.default_components_base
@@ -454,7 +452,7 @@
         "B" -> (_ => multicore_base = true),
         "C:" -> (arg => components_base = Path.explode(arg)),
         "H:" -> (arg => heap = Some(Value.Int.parse(arg))),
-        "M:" -> (arg => multicore_list = space_explode(',', arg).map(Multicore.parse(_))),
+        "M:" -> (arg => multicore_list = space_explode(',', arg).map(Multicore.parse)),
         "N:" -> (arg => isabelle_identifier = arg),
         "P:" -> (arg => afp_partition = Value.Int.parse(arg)),
         "U:" -> (arg => max_heap = Some(Value.Int.parse(arg))),
@@ -521,7 +519,7 @@
     afp_repos_source: String = AFP.repos_source,
     isabelle_identifier: String = "remote_build_history",
     self_update: Boolean = false,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     rev: String = "",
     afp_rev: Option[String] = None,
     options: String = "",
--- a/src/Pure/Admin/build_jdk.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/build_jdk.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -106,7 +106,7 @@
       }
 
       val dir_entry =
-        File.read_dir(tmp_dir).filterNot(suppress_name(_)) match {
+        File.read_dir(tmp_dir).filterNot(suppress_name) match {
           case List(s) => s
           case _ => error("Archive contains multiple directories")
         }
@@ -133,7 +133,7 @@
 
   def build_jdk(
     archives: List[Path],
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     target_dir: Path = Path.current)
   {
     if (Platform.is_windows) error("Cannot build jdk on Windows")
@@ -144,7 +144,7 @@
         val extracted = archives.map(extract_archive(dir, _))
 
         val version =
-          extracted.map(_._1).toSet.toList match {
+          extracted.map(_._1).distinct match {
             case List(version) => version
             case Nil => error("No archives")
             case versions =>
@@ -232,7 +232,7 @@
       val more_args = getopts(args)
       if (more_args.isEmpty) getopts.usage()
 
-      val archives = more_args.map(Path.explode(_))
+      val archives = more_args.map(Path.explode)
       val progress = new Console_Progress()
 
       build_jdk(archives = archives, progress = progress, target_dir = target_dir)
--- a/src/Pure/Admin/build_log.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/build_log.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -103,17 +103,17 @@
 
     def plain_name(name: String): String =
     {
-      List(".log", ".log.gz", ".log.xz", ".gz", ".xz").find(name.endsWith(_)) match {
+      List(".log", ".log.gz", ".log.xz", ".gz", ".xz").find(name.endsWith) match {
         case Some(s) => Library.try_unsuffix(s, name).get
         case None => name
       }
     }
 
     def apply(name: String, lines: List[String]): Log_File =
-      new Log_File(plain_name(name), lines)
+      new Log_File(plain_name(name), lines.map(Library.trim_line))
 
     def apply(name: String, text: String): Log_File =
-      Log_File(name, Library.trim_split_lines(text))
+      new Log_File(plain_name(name), Library.trim_split_lines(text))
 
     def apply(file: JFile): Log_File =
     {
@@ -138,8 +138,8 @@
     {
       val name = file.getName
 
-      prefixes.exists(name.startsWith(_)) &&
-      suffixes.exists(name.endsWith(_)) &&
+      prefixes.exists(name.startsWith) &&
+      suffixes.exists(name.endsWith) &&
       name != "isatest.log" &&
       name != "afp-test.log" &&
       name != "main.log"
@@ -179,20 +179,14 @@
       def tune(s: String): String =
         Word.implode(
           Word.explode(s) match {
-            case a :: "M\uFFFDr" :: bs => tune_weekday(a) :: "Mär" :: bs.map(tune_timezone(_))
-            case a :: bs => tune_weekday(a) :: bs.map(tune_timezone(_))
+            case a :: "M\uFFFDr" :: bs => tune_weekday(a) :: "Mär" :: bs.map(tune_timezone)
+            case a :: bs => tune_weekday(a) :: bs.map(tune_timezone)
             case Nil => Nil
           }
         )
 
       Date.Format.make(fmts, tune)
     }
-
-
-    /* inlined content */
-
-    def print_props(marker: String, props: Properties.T): String =
-      marker + YXML.string_of_body(XML.Encode.properties(Properties.encode_lines(props)))
   }
 
   class Log_File private(val name: String, val lines: List[String])
@@ -217,13 +211,13 @@
     }
 
 
-    /* inlined content */
+    /* inlined text */
 
-    def find[A](f: String => Option[A]): Option[A] =
-      lines.iterator.map(f).find(_.isDefined).map(_.get)
+    def filter(Marker: Protocol_Message.Marker): List[String] =
+      for (Marker(text) <- lines) yield text
 
-    def find_line(marker: String): Option[String] =
-      find(Library.try_unprefix(marker, _))
+    def find(Marker: Protocol_Message.Marker): Option[String] =
+      lines.collectFirst({ case Marker(text) => text })
 
     def find_match(regexes: List[Regex]): Option[String] =
       regexes match {
@@ -249,25 +243,17 @@
 
     /* properties (YXML) */
 
-    val xml_cache = XML.make_cache()
+    val xml_cache: XML.Cache = XML.make_cache()
 
     def parse_props(text: String): Properties.T =
-      try {
-        xml_cache.props(Properties.decode_lines(XML.Decode.properties(YXML.parse_body(text))))
-      }
+      try { xml_cache.props(XML.Decode.properties(YXML.parse_body(text))) }
       catch { case _: XML.Error => log_file.err("malformed properties") }
 
-    def filter_lines(marker: String): List[String] =
-      for (line <- lines; s <- Library.try_unprefix(marker, line)) yield s
-
-    def filter_props(marker: String): List[Properties.T] =
-      for (s <- filter_lines(marker) if YXML.detect(s)) yield parse_props(s)
+    def filter_props(marker: Protocol_Message.Marker): List[Properties.T] =
+      for (text <- filter(marker) if YXML.detect(text)) yield parse_props(text)
 
-    def find_props(marker: String): Option[Properties.T] =
-      find_line(marker) match {
-        case Some(text) if YXML.detect(text) => Some(parse_props(text))
-        case _ => None
-      }
+    def find_props(marker: Protocol_Message.Marker): Option[Properties.T] =
+      for (text <- find(marker) if YXML.detect(text)) yield parse_props(text)
 
 
     /* parse various formats */
@@ -304,7 +290,7 @@
       Properties.get(settings, c.name)
 
     def get_date(c: SQL.Column): Option[Date] =
-      get(c).map(Log_File.Date_Format.parse(_))
+      get(c).map(Log_File.Date_Format.parse)
   }
 
   object Identify
@@ -401,9 +387,8 @@
     }
 
     log_file.lines match {
-      case line :: _ if line.startsWith(Build_History.META_INFO_MARKER) =>
-        Meta_Info(log_file.find_props(Build_History.META_INFO_MARKER).get,
-          log_file.get_all_settings)
+      case line :: _ if Protocol.Meta_Info_Marker.test_yxml(line) =>
+        Meta_Info(log_file.find_props(Protocol.Meta_Info_Marker).get, log_file.get_all_settings)
 
       case Identify.Start(log_file.Strict_Date(start)) :: _ =>
         parse(Identify.engine(log_file), "", start, Identify.No_End,
@@ -446,9 +431,6 @@
 
   /** build info: toplevel output of isabelle build or Admin/build_history **/
 
-  val THEORY_TIMING_MARKER = "\ftheory_timing = "
-  val ML_STATISTICS_MARKER = "\fML_statistics = "
-  val ERROR_MESSAGE_MARKER = "\ferror_message = "
   val SESSION_NAME = "session_name"
 
   object Session_Status extends Enumeration
@@ -514,9 +496,8 @@
     object Theory_Timing
     {
       def unapply(line: String): Option[(String, (String, Timing))] =
-      {
-        val line1 = line.replace('~', '-')
-        Library.try_unprefix(THEORY_TIMING_MARKER, line1).map(log_file.parse_props(_)) match {
+        Protocol.Theory_Timing_Marker.unapply(line.replace('~', '-')).map(log_file.parse_props)
+        match {
           case Some((SESSION_NAME, name) :: props) =>
             (props, props) match {
               case (Markup.Name(thy), Markup.Timing_Properties(t)) => Some((name, thy -> t))
@@ -524,7 +505,6 @@
             }
           case _ => None
         }
-      }
     }
 
     var chapter = Map.empty[String, String]
@@ -586,24 +566,24 @@
         case Heap(name, Value.Long(size)) =>
           heap_sizes += (name -> size)
 
-        case _ if line.startsWith(THEORY_TIMING_MARKER) && YXML.detect(line) =>
+        case _ if Protocol.Theory_Timing_Marker.test_yxml(line) =>
           line match {
             case Theory_Timing(name, theory_timing) =>
               theory_timings += (name -> (theory_timings.getOrElse(name, Map.empty) + theory_timing))
             case _ => log_file.err("malformed theory_timing " + quote(line))
           }
 
-        case _ if parse_ml_statistics && line.startsWith(ML_STATISTICS_MARKER) && YXML.detect(line) =>
-          Library.try_unprefix(ML_STATISTICS_MARKER, line).map(log_file.parse_props(_)) match {
+        case _ if parse_ml_statistics && Protocol.ML_Statistics_Marker.test_yxml(line) =>
+          Protocol.ML_Statistics_Marker.unapply(line).map(log_file.parse_props) match {
             case Some((SESSION_NAME, name) :: props) =>
               ml_statistics += (name -> (props :: ml_statistics.getOrElse(name, Nil)))
             case _ => log_file.err("malformed ML_statistics " + quote(line))
           }
 
-        case _ if line.startsWith(ERROR_MESSAGE_MARKER) && YXML.detect(line) =>
-          Library.try_unprefix(ERROR_MESSAGE_MARKER, line).map(log_file.parse_props(_)) match {
+        case _ if Protocol.Error_Message_Marker.test_yxml(line) =>
+          Protocol.Error_Message_Marker.unapply(line).map(log_file.parse_props) match {
             case Some(List((SESSION_NAME, name), (Markup.CONTENT, msg))) =>
-              errors += (name -> (Library.decode_lines(msg) :: errors.getOrElse(name, Nil)))
+              errors += (name -> (msg :: errors.getOrElse(name, Nil)))
             case _ => log_file.err("malformed error message " + quote(line))
           }
 
@@ -663,12 +643,16 @@
     task_statistics: Boolean): Session_Info =
   {
     Session_Info(
-      session_timing = log_file.find_props("\fTiming = ") getOrElse Nil,
-      command_timings = if (command_timings) log_file.filter_props("\fcommand_timing = ") else Nil,
-      theory_timings = if (theory_timings) log_file.filter_props(THEORY_TIMING_MARKER) else Nil,
-      ml_statistics = if (ml_statistics) log_file.filter_props(ML_STATISTICS_MARKER) else Nil,
-      task_statistics = if (task_statistics) log_file.filter_props("\ftask_statistics = ") else Nil,
-      errors = log_file.filter_lines(ERROR_MESSAGE_MARKER).map(Library.decode_lines(_)))
+      session_timing = log_file.find_props(Protocol.Timing_Marker) getOrElse Nil,
+      command_timings =
+        if (command_timings) log_file.filter_props(Protocol.Command_Timing_Marker) else Nil,
+      theory_timings =
+        if (theory_timings) log_file.filter_props(Protocol.Theory_Timing_Marker) else Nil,
+      ml_statistics =
+        if (ml_statistics) log_file.filter_props(Protocol.ML_Statistics_Marker) else Nil,
+      task_statistics =
+        if (task_statistics) log_file.filter_props(Protocol.Task_Statistics_Marker) else Nil,
+      errors = log_file.filter(Protocol.Error_Message_Marker))
   }
 
   def compress_errors(errors: List[String], cache: XZ.Cache = XZ.cache()): Option[Bytes] =
@@ -681,7 +665,7 @@
   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(cache = cache).text))
+      XML.Decode.list(YXML.string_of_body)(YXML.parse_body(bytes.uncompress(cache = cache).text))
     }
 
 
@@ -754,7 +738,7 @@
 
     /* earliest pull date for repository version (PostgreSQL queries) */
 
-    def pull_date(afp: Boolean = false) =
+    def pull_date(afp: Boolean = false): SQL.Column =
       if (afp) SQL.Column.date("afp_pull_date")
       else SQL.Column.date("pull_date")
 
@@ -1110,7 +1094,7 @@
             files.filter(file => status.exists(_.required(file))).
               grouped(options.int("build_log_transaction_size") max 1))
       {
-        val log_files = Par_List.map[JFile, Log_File](Log_File.apply _, file_group)
+        val log_files = Par_List.map[JFile, Log_File](Log_File.apply, file_group)
         db.transaction { log_files.foreach(log_file => status.foreach(_.update(log_file))) }
       }
     }
@@ -1182,7 +1166,7 @@
                   res.timing(Data.ml_timing_elapsed, Data.ml_timing_cpu, Data.ml_timing_gc),
                 sources = res.get_string(Data.sources),
                 heap_size = res.get_long(Data.heap_size),
-                status = res.get_string(Data.status).map(Session_Status.withName(_)),
+                status = res.get_string(Data.status).map(Session_Status.withName),
                 errors = uncompress_errors(res.bytes(Data.errors), cache = xz_cache),
                 ml_statistics =
                   if (ml_statistics) {
--- a/src/Pure/Admin/build_polyml.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/build_polyml.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -49,7 +49,7 @@
   def build_polyml(
     root: Path,
     sha1_root: Option[Path] = None,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     arch_64: Boolean = false,
     options: List[String] = Nil,
     msys_root: Option[Path] = None)
@@ -256,12 +256,11 @@
   val isabelle_tool1 =
     Isabelle_Tool("build_polyml", "build Poly/ML from sources", args =>
     {
-      Command_Line.tool0 {
-        var msys_root: Option[Path] = None
-        var arch_64 = false
-        var sha1_root: Option[Path] = None
+      var msys_root: Option[Path] = None
+      var arch_64 = false
+      var sha1_root: Option[Path] = None
 
-        val getopts = Getopts("""
+      val getopts = Getopts("""
 Usage: isabelle build_polyml [OPTIONS] ROOT [CONFIGURE_OPTIONS]
 
   Options are:
@@ -272,33 +271,31 @@
   Build Poly/ML in the ROOT directory of its sources, with additional
   CONFIGURE_OPTIONS (e.g. --without-gmp).
 """,
-          "M:" -> (arg => msys_root = Some(Path.explode(arg))),
-          "m:" ->
-            {
-              case "32" | "x86_64_32" => arch_64 = false
-              case "64" | "x86_64" => arch_64 = true
-              case bad => error("Bad processor architecture: " + quote(bad))
-            },
-          "s:" -> (arg => sha1_root = Some(Path.explode(arg))))
+        "M:" -> (arg => msys_root = Some(Path.explode(arg))),
+        "m:" ->
+          {
+            case "32" | "x86_64_32" => arch_64 = false
+            case "64" | "x86_64" => arch_64 = true
+            case bad => error("Bad processor architecture: " + quote(bad))
+          },
+        "s:" -> (arg => sha1_root = Some(Path.explode(arg))))
 
-        val more_args = getopts(args)
-        val (root, options) =
-          more_args match {
-            case root :: options => (Path.explode(root), options)
-            case Nil => getopts.usage()
-          }
-        build_polyml(root, sha1_root = sha1_root, progress = new Console_Progress,
-          arch_64 = arch_64, options = options, msys_root = msys_root)
-      }
+      val more_args = getopts(args)
+      val (root, options) =
+        more_args match {
+          case root :: options => (Path.explode(root), options)
+          case Nil => getopts.usage()
+        }
+      build_polyml(root, sha1_root = sha1_root, progress = new Console_Progress,
+        arch_64 = arch_64, options = options, msys_root = msys_root)
     })
 
   val isabelle_tool2 =
     Isabelle_Tool("build_polyml_component", "make skeleton for Poly/ML component", args =>
     {
-      Command_Line.tool0 {
-        var sha1_root: Option[Path] = None
+      var sha1_root: Option[Path] = None
 
-        val getopts = Getopts("""
+      val getopts = Getopts("""
 Usage: isabelle build_polyml_component [OPTIONS] SOURCE_ARCHIVE COMPONENT_DIR
 
   Options are:
@@ -307,16 +304,15 @@
   Make skeleton for Poly/ML component, based on official source archive
   (zip or tar.gz).
 """,
-          "s:" -> (arg => sha1_root = Some(Path.explode(arg))))
+        "s:" -> (arg => sha1_root = Some(Path.explode(arg))))
 
-        val more_args = getopts(args)
+      val more_args = getopts(args)
 
-        val (source_archive, component_dir) =
-          more_args match {
-            case List(archive, dir) => (Path.explode(archive), Path.explode(dir))
-            case _ => getopts.usage()
-          }
-        build_polyml_component(source_archive, component_dir, sha1_root = sha1_root)
-      }
+      val (source_archive, component_dir) =
+        more_args match {
+          case List(archive, dir) => (Path.explode(archive), Path.explode(dir))
+          case _ => getopts.usage()
+        }
+      build_polyml_component(source_archive, component_dir, sha1_root = sha1_root)
     })
 }
--- a/src/Pure/Admin/build_release.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/build_release.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -163,7 +163,7 @@
       terminate_lines("#bundled components" ::
         (for {
           (catalog, bundled) <- catalogs.iterator
-          val path = Components.admin(dir) + Path.basic(catalog)
+          path = Components.admin(dir) + Path.basic(catalog)
           if path.is_file
           line <- split_lines(File.read(path))
           if line.nonEmpty && !line.startsWith("#") && !line.startsWith("jedit_build")
@@ -196,7 +196,7 @@
             if (Components.check_dir(Components.contrib(dir, name))) Some(contrib_name(name))
             else None
           case _ => if (Bundled.detect(line)) None else Some(line)
-        }) ::: more_names.map(contrib_name(_)))
+        }) ::: more_names.map(contrib_name))
   }
 
   def make_contrib(dir: Path)
@@ -265,7 +265,7 @@
   def build_release(base_dir: Path,
     options: Options,
     components_base: Path = Components.default_components_base,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     rev: String = "",
     afp_rev: String = "",
     official_release: Boolean = false,
@@ -428,7 +428,6 @@
     val bundle_infos = platform_families.map(release.bundle_info)
 
     for (bundle_info <- bundle_infos) {
-      val bundle_archive = release.dist_dir + bundle_info.path
       val isabelle_name = release.dist_name
       val platform = bundle_info.platform
 
@@ -776,7 +775,7 @@
 
   def main(args: Array[String])
   {
-    Command_Line.tool0 {
+    Command_Line.tool {
       var afp_rev = ""
       var components_base: Path = Components.default_components_base
       var official_release = false
--- a/src/Pure/Admin/build_status.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/build_status.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -49,7 +49,7 @@
   /* build status */
 
   def build_status(options: Options,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     profiles: List[Profile] = default_profiles,
     only_sessions: Set[String] = Set.empty,
     verbose: Boolean = false,
@@ -200,7 +200,7 @@
   }
 
   def read_data(options: Options,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     profiles: List[Profile] = default_profiles,
     only_sessions: Set[String] = Set.empty,
     ml_statistics: Boolean = false,
@@ -362,7 +362,7 @@
   /* present data */
 
   def present_data(data: Data,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     target_dir: Path = default_target_dir,
     image_size: (Int, Int) = default_image_size)
   {
--- a/src/Pure/Admin/components.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/components.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -38,14 +38,14 @@
 
   /* component collections */
 
-  val default_components_base = Path.explode("$ISABELLE_COMPONENTS_BASE")
+  val default_components_base: Path = Path.explode("$ISABELLE_COMPONENTS_BASE")
 
   def admin(dir: Path): Path = dir + Path.explode("Admin/components")
 
   def contrib(dir: Path = Path.current, name: String = ""): Path =
     dir + Path.explode("contrib") + Path.explode(name)
 
-  def unpack(dir: Path, archive: Path, progress: Progress = No_Progress): String =
+  def unpack(dir: Path, archive: Path, progress: Progress = new Progress): String =
   {
     val name = Archive.get_name(archive.file_name)
     progress.echo("Unpacking " + name)
@@ -56,7 +56,7 @@
   def resolve(base_dir: Path, names: List[String],
     target_dir: Option[Path] = None,
     copy_dir: Option[Path] = None,
-    progress: Progress = No_Progress)
+    progress: Progress = new Progress)
   {
     Isabelle_System.mkdirs(base_dir)
     for (name <- names) {
@@ -125,7 +125,7 @@
         case _ => error("Bad components.sha1 entry: " + quote(line))
       })
 
-  def write_components_sha1(entries: List[SHA1_Digest]) =
+  def write_components_sha1(entries: List[SHA1_Digest]): Unit =
     File.write(components_sha1, entries.sortBy(_.file_name).mkString("", "\n", "\n"))
 
 
@@ -135,7 +135,7 @@
   def build_components(
     options: Options,
     components: List[Path],
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     publish: Boolean = false,
     force: Boolean = false,
     update_components_sha1: Boolean = false)
--- a/src/Pure/Admin/isabelle_cronjob.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/isabelle_cronjob.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -18,14 +18,14 @@
   /* global resources: owned by main cronjob */
 
   val backup = "lxbroy10:cronjob"
-  val main_dir = Path.explode("~/cronjob")
-  val main_state_file = main_dir + Path.explode("run/main.state")
-  val current_log = main_dir + Path.explode("run/main.log")  // owned by log service
-  val cumulative_log = main_dir + Path.explode("log/main.log")  // owned by log service
+  val main_dir: Path = Path.explode("~/cronjob")
+  val main_state_file: Path = main_dir + Path.explode("run/main.state")
+  val current_log: Path = main_dir + Path.explode("run/main.log")  // owned by log service
+  val cumulative_log: Path = main_dir + Path.explode("log/main.log")  // owned by log service
 
   val isabelle_repos_source = "https://isabelle.sketis.net/repos/isabelle"
-  val isabelle_repos = main_dir + Path.explode("isabelle")
-  val afp_repos = main_dir + Path.explode("AFP")
+  val isabelle_repos: Path = main_dir + Path.explode("isabelle")
+  val afp_repos: Path = main_dir + Path.explode("AFP")
 
   val build_log_dirs =
     List(Path.explode("~/log"), Path.explode("~/afp/log"), Path.explode("~/cronjob/log"))
@@ -42,7 +42,7 @@
   def get_rev(): String = Mercurial.repository(isabelle_repos).id()
   def get_afp_rev(): String = Mercurial.repository(afp_repos).id()
 
-  val init =
+  val init: Logger_Task =
     Logger_Task("init", logger =>
       {
         Isabelle_Devel.make_index()
@@ -61,7 +61,7 @@
           Files.createSymbolicLink(Isabelle_Devel.cronjob_log.file.toPath, current_log.file.toPath)
       })
 
-  val exit =
+  val exit: Logger_Task =
     Logger_Task("exit", logger =>
       {
         Isabelle_System.bash(
@@ -72,7 +72,7 @@
 
   /* build release */
 
-  val build_release =
+  val build_release: Logger_Task =
     Logger_Task("build_release", logger =>
       {
         Isabelle_Devel.release_snapshot(logger.options, rev = get_rev(), afp_rev = get_afp_rev())
@@ -81,7 +81,7 @@
 
   /* integrity test of build_history vs. build_history_base */
 
-  val build_history_base =
+  val build_history_base: Logger_Task =
     Logger_Task("build_history_base", logger =>
       {
         using(logger.ssh_context.open_session("lxbroy10"))(ssh =>
@@ -147,9 +147,9 @@
   sealed case class Remote_Build(
     description: String,
     host: String,
+    actual_host: String = "",
     user: String = "",
     port: Int = 0,
-    ssh_host: String = "",
     proxy_host: String = "",
     proxy_user: String = "",
     proxy_port: Int = 0,
@@ -166,7 +166,7 @@
     active: Boolean = true)
   {
     def ssh_session(context: SSH.Context): SSH.Session =
-      context.open_session(host = proper_string(ssh_host) getOrElse host, user = user, port = port,
+      context.open_session(host = host, user = user, port = port, actual_host = actual_host,
         proxy_host = proxy_host, proxy_user = proxy_user, proxy_port = proxy_port,
         permissive = proxy_host.nonEmpty)
 
@@ -216,7 +216,6 @@
     List(
       Remote_Build("AFP bulky", "lrzcloud1", self_update = true,
         proxy_host = "lxbroy10", proxy_user = "i21isatest",
-        ssh_host = "10.155.208.96",
         options = "-m64 -M6 -U30000 -s10 -t AFP",
         args = "-g large -g slow",
         afp = true,
@@ -291,18 +290,18 @@
           detect = Build_Log.Prop.build_tags + " = " + SQL.string("skip_proofs"),
           history_base = "2c0f24e927dd")),
       List(
-        Remote_Build("macOS 10.12 Sierra", "macbroy30", options = "-m32 -M2", args = "-a",
+        Remote_Build("Mac OS X 10.12 Sierra", "macbroy30", options = "-m32 -M2", args = "-a",
           detect = Build_Log.Prop.build_start + " > date '2017-03-03'")),
       List(Remote_Build("Mac OS X 10.10 Yosemite", "macbroy31", options = "-m32 -M2", args = "-a")),
-      List(Remote_Build("macOS 10.13 High Sierra", "lapbroy68",
+      List(Remote_Build("Mac OS X 10.13 High Sierra", "lapbroy68",
         options = "-m32 -M1,2,4 -e ISABELLE_GHC_SETUP=true",
         self_update = true, args = "-a -d '~~/src/Benchmarks'")),
-      List(Remote_Build("macOS 10.14 Mojave", "lapnipkow3",
+      List(Remote_Build("Mac OS X 10.14 Mojave", "lapnipkow3",
         options = "-m32 -M1,2 -e ISABELLE_GHC_SETUP=true",
         self_update = true, args = "-a -d '~~/src/Benchmarks'")),
-      List(Remote_Build("macOS 10.15 Catalina", "laramac01", user = "makarius",
+      List(Remote_Build("Mac OS X 10.15 Catalina", "laramac01", user = "makarius",
         proxy_host = "laraserver", proxy_user = "makarius",
-        self_update = true, options = "-m32 -M1,2,4 -e ISABELLE_GHC_SETUP=true",
+        self_update = true, options = "-m32 -M4 -e ISABELLE_GHC_SETUP=true",
         args = "-a -d '~~/src/Benchmarks'")),
       List(
         Remote_Build("Windows", "vmnipkow9", historic = true, history = 90, self_update = true,
@@ -339,9 +338,8 @@
   val remote_builds2: List[List[Remote_Build]] =
     List(
       List(
-        Remote_Build("AFP2", "lrzcloud2", self_update = true,
+        Remote_Build("AFP2", "lrzcloud2", actual_host = "10.195.2.10", self_update = true,
           proxy_host = "lxbroy10", proxy_user = "i21isatest",
-          ssh_host = "10.195.2.10",
           options = "-m32 -M1x8 -t AFP" +
             " -e ISABELLE_GHC=ghc" +
             " -e ISABELLE_MLTON=mlton" +
@@ -350,9 +348,8 @@
           args = "-a -X large -X slow",
           afp = true,
           detect = Build_Log.Prop.build_tags + " = " + SQL.string("AFP")),
-        Remote_Build("AFP bulky2", "lrzcloud2", self_update = true,
+        Remote_Build("AFP bulky2", "lrzcloud2", actual_host = "10.195.2.10", self_update = true,
           proxy_host = "lxbroy10", proxy_user = "i21isatest",
-          ssh_host = "10.195.2.10",
           options = "-m64 -M8 -U30000 -s10 -t AFP",
           args = "-g large -g slow",
           afp = true,
@@ -397,7 +394,7 @@
 
   object Log_Service
   {
-    def apply(options: Options, progress: Progress = No_Progress): Log_Service =
+    def apply(options: Options, progress: Progress = new Progress): Log_Service =
       new Log_Service(SSH.init_context(options), progress)
   }
 
@@ -417,7 +414,7 @@
 
     def shutdown() { thread.shutdown() }
 
-    val hostname = Isabelle_System.hostname()
+    val hostname: String = Isabelle_System.hostname()
 
     def log(date: Date, task_name: String, msg: String): Unit =
       if (task_name != "")
@@ -517,7 +514,7 @@
         {
           running.partition(_.is_finished) match {
             case (Nil, Nil) =>
-            case (Nil, _ :: _) => Thread.sleep(500); join(running)
+            case (Nil, _ :: _) => Time.seconds(0.5).sleep; join(running)
             case (_ :: _, remaining) => join(remaining)
           }
         }
@@ -582,7 +579,7 @@
 
   def main(args: Array[String])
   {
-    Command_Line.tool0 {
+    Command_Line.tool {
       var force = false
       var verbose = false
       var exclude_task = Set.empty[String]
@@ -602,7 +599,7 @@
       val more_args = getopts(args)
       if (more_args.nonEmpty) getopts.usage()
 
-      val progress = if (verbose) new Console_Progress() else No_Progress
+      val progress = if (verbose) new Console_Progress() else new Progress
 
       if (force) cronjob(progress, exclude_task)
       else error("Need to apply force to do anything")
--- a/src/Pure/Admin/isabelle_devel.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/isabelle_devel.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -14,8 +14,8 @@
   val BUILD_STATUS = "build_status"
   val CRONJOB_LOG = "cronjob-main.log"
 
-  val root = Path.explode("~/html-data/devel")
-  val cronjob_log = root + Path.basic(CRONJOB_LOG)
+  val root: Path = Path.explode("~/html-data/devel")
+  val cronjob_log: Path = root + Path.basic(CRONJOB_LOG)
 
 
   /* index */
--- a/src/Pure/Admin/jenkins.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/jenkins.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -40,10 +40,10 @@
 
 
   def download_logs(
-    options: Options, job_names: List[String], dir: Path, progress: Progress = No_Progress)
+    options: Options, job_names: List[String], dir: Path, progress: Progress = new Progress)
   {
     val store = Sessions.store(options)
-    val infos = job_names.flatMap(build_job_infos(_))
+    val infos = job_names.flatMap(build_job_infos)
     Par_List.map((info: Job_Info) => info.download_log(store, dir, progress), infos)
   }
 
@@ -96,7 +96,7 @@
       }
     }
 
-    def download_log(store: Sessions.Store, dir: Path, progress: Progress = No_Progress)
+    def download_log(store: Sessions.Store, dir: Path, progress: Progress = new Progress)
     {
       val log_dir = dir + Build_Log.log_subdir(date)
       val log_path = log_dir + log_filename.ext("xz")
@@ -106,13 +106,13 @@
         Isabelle_System.mkdirs(log_dir)
 
         val ml_statistics =
-          session_logs.map(_._1).toSet.toList.sorted.flatMap(session_name =>
+          session_logs.map(_._1).distinct.sorted.flatMap(session_name =>
             read_ml_statistics(store, session_name).
               map(props => (Build_Log.SESSION_NAME -> session_name) :: props))
 
         File.write_xz(log_path,
           terminate_lines(Url.read(main_log) ::
-            ml_statistics.map(Build_Log.Log_File.print_props(Build_Log.ML_STATISTICS_MARKER, _))),
+            ml_statistics.map(Protocol.ML_Statistics_Marker.apply)),
           XZ.options(6))
       }
     }
--- a/src/Pure/Admin/other_isabelle.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Admin/other_isabelle.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -12,7 +12,7 @@
   def apply(isabelle_home: Path,
       isabelle_identifier: String = "",
       user_home: Path = Path.explode("$USER_HOME"),
-      progress: Progress = No_Progress): Other_Isabelle =
+      progress: Progress = new Progress): Other_Isabelle =
     new Other_Isabelle(isabelle_home.canonical, isabelle_identifier, user_home, progress)
 }
 
@@ -62,7 +62,7 @@
   val etc_preferences: Path = etc + Path.explode("preferences")
 
   def copy_fonts(target_dir: Path): Unit =
-    Isabelle_Fonts.make_entries(getenv = getenv(_), hidden = true).
+    Isabelle_Fonts.make_entries(getenv = getenv, hidden = true).
       foreach(entry => File.copy(entry.path, target_dir))
 
 
--- a/src/Pure/Concurrent/consumer_thread.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Concurrent/consumer_thread.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -47,7 +47,7 @@
   private var active = true
   private val mailbox = Mailbox[Option[Request]]
 
-  private val thread = Standard_Thread.fork(name, daemon) { main_loop(Nil) }
+  private val thread = Isabelle_Thread.fork(name = name, daemon = daemon) { main_loop(Nil) }
   def is_active: Boolean = active && thread.isAlive
   def check_thread: Boolean = Thread.currentThread == thread
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/delay.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -0,0 +1,66 @@
+/*  Title:      Pure/Concurrent/delay.scala
+    Author:     Makarius
+
+Delayed events.
+*/
+
+package isabelle
+
+
+object Delay
+{
+  // delayed event after first invocation
+  def first(delay: => Time, log: Logger = No_Logger, gui: Boolean = false)(event: => Unit): Delay =
+    new Delay(true, delay, log, if (gui) GUI_Thread.later { event } else event )
+
+  // delayed event after last invocation
+  def last(delay: => Time, log: Logger = No_Logger, gui: Boolean = false)(event: => Unit): Delay =
+    new Delay(false, delay, log, if (gui) GUI_Thread.later { event } else event )
+}
+
+final class Delay private(first: Boolean, delay: => Time, log: Logger, event: => Unit)
+{
+  private var running: Option[Event_Timer.Request] = None
+
+  private def run: Unit =
+  {
+    val do_run = synchronized {
+      if (running.isDefined) { running = None; true } else false
+    }
+    if (do_run) {
+      try { event }
+      catch { case exn: Throwable if !Exn.is_interrupt(exn) => log(Exn.message(exn)); throw exn }
+    }
+  }
+
+  def invoke(): Unit = synchronized
+  {
+    val new_run =
+      running match {
+        case Some(request) => if (first) false else { request.cancel; true }
+        case None => true
+      }
+    if (new_run)
+      running = Some(Event_Timer.request(Time.now() + delay)(run))
+  }
+
+  def revoke(): Unit = synchronized
+  {
+    running match {
+      case Some(request) => request.cancel; running = None
+      case None =>
+    }
+  }
+
+  def postpone(alt_delay: Time): Unit = synchronized
+  {
+    running match {
+      case Some(request) =>
+        val alt_time = Time.now() + alt_delay
+        if (request.time < alt_time && request.cancel) {
+          running = Some(Event_Timer.request(alt_time)(run))
+        }
+      case None =>
+    }
+  }
+}
--- a/src/Pure/Concurrent/event_timer.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Concurrent/event_timer.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -130,7 +130,7 @@
 fun manager_check manager =
   if is_some manager andalso Thread.isActive (the manager) then manager
   else
-    SOME (Standard_Thread.fork {name = "event_timer", stack_limit = NONE, interrupts = false}
+    SOME (Isabelle_Thread.fork {name = "event_timer", stack_limit = NONE, interrupts = false}
       manager_loop);
 
 fun shutdown () =
--- a/src/Pure/Concurrent/future.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Concurrent/future.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -202,14 +202,14 @@
   let
     val running = Task_Queue.cancel (! queue) group;
     val _ = running |> List.app (fn thread =>
-      if Standard_Thread.is_self thread then ()
-      else Standard_Thread.interrupt_unsynchronized thread);
+      if Isabelle_Thread.is_self thread then ()
+      else Isabelle_Thread.interrupt_unsynchronized thread);
   in running end;
 
 fun cancel_all () = (*requires SYNCHRONIZED*)
   let
     val (groups, threads) = Task_Queue.cancel_all (! queue);
-    val _ = List.app Standard_Thread.interrupt_unsynchronized threads;
+    val _ = List.app Isabelle_Thread.interrupt_unsynchronized threads;
   in groups end;
 
 fun cancel_later group = (*requires SYNCHRONIZED*)
@@ -280,7 +280,7 @@
       Real.floor (Options.default_real "threads_stack_limit" * 1024.0 * 1024.0 * 1024.0);
     val stack_limit = if threads_stack_limit <= 0 then NONE else SOME threads_stack_limit;
     val worker =
-      Standard_Thread.fork {name = "worker", stack_limit = stack_limit, interrupts = false}
+      Isabelle_Thread.fork {name = "worker", stack_limit = stack_limit, interrupts = false}
         (fn () => worker_loop name);
   in Unsynchronized.change workers (cons (worker, Unsynchronized.ref Working)) end
   handle Fail msg => Multithreading.tracing 0 (fn () => "SCHEDULER: " ^ msg);
@@ -383,7 +383,7 @@
   if scheduler_active () then ()
   else
     scheduler :=
-      SOME (Standard_Thread.fork {name = "scheduler", stack_limit = NONE, interrupts = false}
+      SOME (Isabelle_Thread.fork {name = "scheduler", stack_limit = NONE, interrupts = false}
         scheduler_loop));
 
 
--- a/src/Pure/Concurrent/future.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Concurrent/future.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -17,8 +17,17 @@
   def value[A](x: A): Future[A] = new Value_Future(x)
   def fork[A](body: => A): Future[A] = new Task_Future[A](body)
   def promise[A]: Promise[A] = new Promise_Future[A]
-  def thread[A](name: String = "", daemon: Boolean = false)(body: => A): Future[A] =
-    new Thread_Future[A](name, daemon, body)
+
+  def thread[A](
+    name: String = "",
+    group: ThreadGroup = Isabelle_Thread.current_thread_group,
+    pri: Int = Thread.NORM_PRIORITY,
+    daemon: Boolean = false,
+    inherit_locals: Boolean = false,
+    uninterruptible: Boolean = false)(body: => A): Future[A] =
+    {
+      new Thread_Future[A](name, group, pri, daemon, inherit_locals, uninterruptible, body)
+    }
 }
 
 trait Future[A]
@@ -87,7 +96,7 @@
       status.change(_ => Finished(if (Thread.interrupted) Exn.Exn(Exn.Interrupt()) else result))
     }
   }
-  private val task = Standard_Thread.pool.submit(new Callable[Unit] { def call = try_run() })
+  private val task = Isabelle_Thread.pool.submit(new Callable[Unit] { def call = try_run() })
 
   def join_result: Exn.Result[A] =
   {
@@ -131,11 +140,20 @@
 
 /* thread future */
 
-private class Thread_Future[A](name: String, daemon: Boolean, body: => A) extends Future[A]
+private class Thread_Future[A](
+  name: String,
+  group: ThreadGroup,
+  pri: Int,
+  daemon: Boolean,
+  inherit_locals: Boolean,
+  uninterruptible: Boolean,
+  body: => A) extends Future[A]
 {
   private val result = Future.promise[A]
   private val thread =
-    Standard_Thread.fork(name, daemon) { result.fulfill_result(Exn.capture(body)) }
+    Isabelle_Thread.fork(name = name, group = group, pri = pri, daemon = daemon,
+      inherit_locals = inherit_locals, uninterruptible = uninterruptible)
+    { result.fulfill_result(Exn.capture(body)) }
 
   def peek: Option[Exn.Result[A]] = result.peek
   def join_result: Exn.Result[A] = result.join_result
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/isabelle_thread.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -0,0 +1,73 @@
+(*  Title:      Pure/Concurrent/isabelle_thread.ML
+    Author:     Makarius
+
+Isabelle-specific thread management.
+*)
+
+signature ISABELLE_THREAD =
+sig
+  val is_self: Thread.thread -> bool
+  val get_name: unit -> string
+  type params = {name: string, stack_limit: int option, interrupts: bool}
+  val attributes: params -> Thread.threadAttribute list
+  val fork: params -> (unit -> unit) -> Thread.thread
+  val join: Thread.thread -> unit
+  val interrupt_unsynchronized: Thread.thread -> unit
+end;
+
+structure Isabelle_Thread: ISABELLE_THREAD =
+struct
+
+(* self *)
+
+fun is_self thread = Thread.equal (Thread.self (), thread);
+
+
+(* unique name *)
+
+local
+  val name_var = Thread_Data.var () : string Thread_Data.var;
+  val count = Counter.make ();
+in
+
+fun get_name () =
+  (case Thread_Data.get name_var of
+    SOME name => name
+  | NONE => raise Fail "Isabelle-specific thread required");
+
+fun set_name base =
+  Thread_Data.put name_var (SOME ("Isabelle." ^ base ^ "-" ^ string_of_int (count ())));
+
+end;
+
+
+(* fork *)
+
+type params = {name: string, stack_limit: int option, interrupts: bool};
+
+fun attributes ({stack_limit, interrupts, ...}: params) =
+  Thread.MaximumMLStack stack_limit ::
+  Thread_Attributes.convert_attributes
+    (if interrupts then Thread_Attributes.public_interrupts else Thread_Attributes.no_interrupts);
+
+fun fork (params: params) body =
+  Thread.fork (fn () =>
+    Exn.trace General.exnMessage tracing (fn () =>
+      (set_name (#name params); body ())
+        handle exn => if Exn.is_interrupt exn then () (*sic!*) else Exn.reraise exn),
+    attributes params);
+
+
+(* join *)
+
+fun join thread =
+  while Thread.isActive thread
+  do OS.Process.sleep (seconds 0.1);
+
+
+(* interrupt *)
+
+fun interrupt_unsynchronized thread =
+  Thread.interrupt thread handle Thread _ => ();
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/isabelle_thread.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -0,0 +1,182 @@
+/*  Title:      Pure/Concurrent/isabelle_thread.scala
+    Author:     Makarius
+
+Isabelle-specific thread management.
+*/
+
+package isabelle
+
+
+import java.util.concurrent.{ThreadPoolExecutor, TimeUnit, LinkedBlockingQueue, ThreadFactory}
+
+
+object Isabelle_Thread
+{
+  /* self-thread */
+
+  def self: Isabelle_Thread =
+    Thread.currentThread match {
+      case thread: Isabelle_Thread => thread
+      case thread => error("Isabelle-specific thread required: " + thread)
+    }
+
+  def check_self: Boolean =
+    Thread.currentThread.isInstanceOf[Isabelle_Thread]
+
+
+  /* create threads */
+
+  private val counter = Counter.make()
+
+  def make_name(name: String = "", base: String = "thread"): String =
+    "Isabelle." + proper_string(name).getOrElse(base + counter())
+
+  def current_thread_group: ThreadGroup = Thread.currentThread.getThreadGroup
+
+  def create(
+    main: Runnable,
+    name: String = "",
+    group: ThreadGroup = current_thread_group,
+    pri: Int = Thread.NORM_PRIORITY,
+    daemon: Boolean = false,
+    inherit_locals: Boolean = false): Isabelle_Thread =
+  {
+    new Isabelle_Thread(main, name = make_name(name = name), group = group,
+      pri = pri, daemon = daemon, inherit_locals = inherit_locals)
+  }
+
+  def fork(
+    name: String = "",
+    group: ThreadGroup = current_thread_group,
+    pri: Int = Thread.NORM_PRIORITY,
+    daemon: Boolean = false,
+    inherit_locals: Boolean = false,
+    uninterruptible: Boolean = false)(body: => Unit): Isabelle_Thread =
+  {
+    val main: Runnable =
+      if (uninterruptible) { () => Isabelle_Thread.uninterruptible { body } }
+      else { () => body }
+    val thread =
+      create(main, name = name, group = group, pri = pri,
+        daemon = daemon, inherit_locals = inherit_locals)
+    thread.start
+    thread
+  }
+
+
+  /* thread pool */
+
+  lazy val pool: ThreadPoolExecutor =
+  {
+    val m = Value.Int.unapply(System.getProperty("isabelle.threads", "0")) getOrElse 0
+    val n = if (m > 0) m else (Runtime.getRuntime.availableProcessors max 1) min 8
+    val executor =
+      new ThreadPoolExecutor(n, n, 2500L, TimeUnit.MILLISECONDS, new LinkedBlockingQueue[Runnable])
+    executor.setThreadFactory(create(_, name = make_name(base = "worker"), group = current_thread_group))
+    executor
+  }
+
+
+  /* interrupt handlers */
+
+  object Interrupt_Handler
+  {
+    def apply(handle: Isabelle_Thread => Unit, name: String = "handler"): Interrupt_Handler =
+      new Interrupt_Handler(handle, name)
+
+    val interruptible: Interrupt_Handler =
+      Interrupt_Handler(_.raise_interrupt, name = "interruptible")
+
+    val uninterruptible: Interrupt_Handler =
+      Interrupt_Handler(_.postpone_interrupt, name = "uninterruptible")
+  }
+
+  class Interrupt_Handler private(handle: Isabelle_Thread => Unit, name: String)
+    extends Function[Isabelle_Thread, Unit]
+  {
+    def apply(thread: Isabelle_Thread) { handle(thread) }
+    override def toString: String = name
+  }
+
+  def interrupt_handler[A](handler: Interrupt_Handler)(body: => A): A =
+    if (handler == null) body
+    else self.interrupt_handler(handler)(body)
+
+  def interrupt_handler[A](handle: Isabelle_Thread => Unit)(body: => A): A =
+    self.interrupt_handler(Interrupt_Handler(handle))(body)
+
+  def interruptible[A](body: => A): A =
+    interrupt_handler(Interrupt_Handler.interruptible)(body)
+
+  def uninterruptible[A](body: => A): A =
+    interrupt_handler(Interrupt_Handler.uninterruptible)(body)
+
+  def try_uninterruptible[A](body: => A): A =
+    if (check_self) interrupt_handler(Interrupt_Handler.uninterruptible)(body)
+    else body
+}
+
+class Isabelle_Thread private(main: Runnable, name: String, group: ThreadGroup,
+    pri: Int, daemon: Boolean, inherit_locals: Boolean)
+  extends Thread(group, null, name, 0L, inherit_locals)
+{
+  thread =>
+
+  thread.setPriority(pri)
+  thread.setDaemon(daemon)
+
+  override def run { main.run() }
+
+  def is_self: Boolean = Thread.currentThread == thread
+
+
+  /* interrupt state */
+
+  // synchronized, with concurrent changes
+  private var interrupt_postponed: Boolean = false
+
+  def clear_interrupt: Boolean = synchronized
+  {
+    val was_interrupted = isInterrupted || interrupt_postponed
+    Exn.Interrupt.dispose()
+    interrupt_postponed = false
+    was_interrupted
+  }
+
+  def raise_interrupt: Unit = synchronized
+  {
+    interrupt_postponed = false
+    super.interrupt()
+  }
+
+  def postpone_interrupt: Unit = synchronized
+  {
+    interrupt_postponed = true
+    Exn.Interrupt.dispose()
+  }
+
+
+  /* interrupt handler */
+
+  // non-synchronized, only changed on self-thread
+  @volatile private var handler = Isabelle_Thread.Interrupt_Handler.interruptible
+
+  override def interrupt: Unit = handler(thread)
+
+  def interrupt_handler[A](new_handler: Isabelle_Thread.Interrupt_Handler)(body: => A): A =
+    if (new_handler == null) body
+    else {
+      require(is_self)
+
+      val old_handler = handler
+      handler = new_handler
+      try {
+        if (clear_interrupt) interrupt
+        body
+      }
+      finally {
+        handler = old_handler
+        if (clear_interrupt) interrupt
+      }
+    }
+}
--- a/src/Pure/Concurrent/standard_thread.ML	Wed Apr 15 12:33:42 2020 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,76 +0,0 @@
-(*  Title:      Pure/Concurrent/standard_thread.ML
-    Author:     Makarius
-
-Standard thread operations.
-*)
-
-signature STANDARD_THREAD =
-sig
-  val is_self: Thread.thread -> bool
-  val get_name: unit -> string option
-  val the_name: unit -> string
-  type params = {name: string, stack_limit: int option, interrupts: bool}
-  val attributes: params -> Thread.threadAttribute list
-  val fork: params -> (unit -> unit) -> Thread.thread
-  val join: Thread.thread -> unit
-  val interrupt_unsynchronized: Thread.thread -> unit
-end;
-
-structure Standard_Thread: STANDARD_THREAD =
-struct
-
-(* self *)
-
-fun is_self thread = Thread.equal (Thread.self (), thread);
-
-
-(* unique name *)
-
-local
-  val name_var = Thread_Data.var () : string Thread_Data.var;
-  val count = Counter.make ();
-in
-
-fun get_name () = Thread_Data.get name_var;
-
-fun the_name () =
-  (case get_name () of
-    NONE => raise Fail "Unknown thread name"
-  | SOME name => name);
-
-fun set_name base =
-  Thread_Data.put name_var (SOME (base ^ "/" ^ string_of_int (count ())));
-
-end;
-
-
-(* fork *)
-
-type params = {name: string, stack_limit: int option, interrupts: bool};
-
-fun attributes ({stack_limit, interrupts, ...}: params) =
-  Thread.MaximumMLStack stack_limit ::
-  Thread_Attributes.convert_attributes
-    (if interrupts then Thread_Attributes.public_interrupts else Thread_Attributes.no_interrupts);
-
-fun fork (params: params) body =
-  Thread.fork (fn () =>
-    Exn.trace General.exnMessage tracing (fn () =>
-      (set_name (#name params); body ())
-        handle exn => if Exn.is_interrupt exn then () (*sic!*) else Exn.reraise exn),
-    attributes params);
-
-
-(* join *)
-
-fun join thread =
-  while Thread.isActive thread
-  do OS.Process.sleep (seconds 0.1);
-
-
-(* interrupt *)
-
-fun interrupt_unsynchronized thread =
-  Thread.interrupt thread handle Thread _ => ();
-
-end;
--- a/src/Pure/Concurrent/standard_thread.scala	Wed Apr 15 12:33:42 2020 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,108 +0,0 @@
-/*  Title:      Pure/Concurrent/standard_thread.scala
-    Author:     Makarius
-
-Standard thread operations.
-*/
-
-package isabelle
-
-
-import java.lang.Thread
-import java.util.concurrent.{ThreadPoolExecutor, TimeUnit, LinkedBlockingQueue, ThreadFactory}
-
-
-object Standard_Thread
-{
-  /* fork */
-
-  def fork(name: String = "", daemon: Boolean = false)(body: => Unit): Thread =
-  {
-    val thread =
-      if (name == null || name == "") new Thread() { override def run = body }
-      else new Thread(name) { override def run = body }
-    thread.setDaemon(daemon)
-    thread.start
-    thread
-  }
-
-
-  /* pool */
-
-  lazy val pool: ThreadPoolExecutor =
-    {
-      val m = Value.Int.unapply(System.getProperty("isabelle.threads", "0")) getOrElse 0
-      val n = if (m > 0) m else (Runtime.getRuntime.availableProcessors max 1) min 8
-      val executor =
-        new ThreadPoolExecutor(n, n, 2500L, TimeUnit.MILLISECONDS, new LinkedBlockingQueue[Runnable])
-      val old_thread_factory = executor.getThreadFactory
-      executor.setThreadFactory(
-        new ThreadFactory {
-          def newThread(r: Runnable) =
-          {
-            val thread = old_thread_factory.newThread(r)
-            thread.setDaemon(true)
-            thread
-          }
-        })
-      executor
-    }
-
-
-  /* delayed events */
-
-  final class Delay private[Standard_Thread](
-    first: Boolean, delay: => Time, log: Logger, event: => Unit)
-  {
-    private var running: Option[Event_Timer.Request] = None
-
-    private def run: Unit =
-    {
-      val do_run = synchronized {
-        if (running.isDefined) { running = None; true } else false
-      }
-      if (do_run) {
-        try { event }
-        catch { case exn: Throwable if !Exn.is_interrupt(exn) => log(Exn.message(exn)); throw exn }
-      }
-    }
-
-    def invoke(): Unit = synchronized
-    {
-      val new_run =
-        running match {
-          case Some(request) => if (first) false else { request.cancel; true }
-          case None => true
-        }
-      if (new_run)
-        running = Some(Event_Timer.request(Time.now() + delay)(run))
-    }
-
-    def revoke(): Unit = synchronized
-    {
-      running match {
-        case Some(request) => request.cancel; running = None
-        case None =>
-      }
-    }
-
-    def postpone(alt_delay: Time): Unit = synchronized
-    {
-      running match {
-        case Some(request) =>
-          val alt_time = Time.now() + alt_delay
-          if (request.time < alt_time && request.cancel) {
-            running = Some(Event_Timer.request(alt_time)(run))
-          }
-        case None =>
-      }
-    }
-  }
-
-  // delayed event after first invocation
-  def delay_first(delay: => Time, log: Logger = No_Logger)(event: => Unit): Delay =
-    new Delay(true, delay, log, event)
-
-  // delayed event after last invocation
-  def delay_last(delay: => Time, log: Logger = No_Logger)(event: => Unit): Delay =
-    new Delay(false, delay, log, event)
-}
--- a/src/Pure/Concurrent/timeout.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Concurrent/timeout.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -24,7 +24,7 @@
 
       val request =
         Event_Timer.request {physical = false} (start + timeout)
-          (fn () => Standard_Thread.interrupt_unsynchronized self);
+          (fn () => Isabelle_Thread.interrupt_unsynchronized self);
       val result =
         Exn.capture (fn () => Thread_Attributes.with_attributes orig_atts (fn _ => f x)) ();
 
--- a/src/Pure/GUI/gui_thread.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/GUI/gui_thread.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -16,13 +16,13 @@
 
   def assert[A](body: => A): A =
   {
-    Predef.assert(SwingUtilities.isEventDispatchThread())
+    Predef.assert(SwingUtilities.isEventDispatchThread)
     body
   }
 
   def require[A](body: => A): A =
   {
-    Predef.require(SwingUtilities.isEventDispatchThread())
+    Predef.require(SwingUtilities.isEventDispatchThread)
     body
   }
 
@@ -31,36 +31,27 @@
 
   def now[A](body: => A): A =
   {
-    if (SwingUtilities.isEventDispatchThread()) body
+    if (SwingUtilities.isEventDispatchThread) body
     else {
-      lazy val result = { assert { Exn.capture(body) } }
-      SwingUtilities.invokeAndWait(new Runnable { def run = result })
+      lazy val result = assert { Exn.capture(body) }
+      SwingUtilities.invokeAndWait(() => result)
       Exn.release(result)
     }
   }
 
   def later(body: => Unit)
   {
-    if (SwingUtilities.isEventDispatchThread()) body
-    else SwingUtilities.invokeLater(new Runnable { def run = body })
+    if (SwingUtilities.isEventDispatchThread) body
+    else SwingUtilities.invokeLater(() => body)
   }
 
   def future[A](body: => A): Future[A] =
   {
-    if (SwingUtilities.isEventDispatchThread()) Future.value(body)
+    if (SwingUtilities.isEventDispatchThread) Future.value(body)
     else {
       val promise = Future.promise[A]
       later { promise.fulfill_result(Exn.capture(body)) }
       promise
     }
   }
-
-
-  /* delayed events */
-
-  def delay_first(delay: => Time)(event: => Unit): Standard_Thread.Delay =
-    Standard_Thread.delay_first(delay) { later { event } }
-
-  def delay_last(delay: => Time)(event: => Unit): Standard_Thread.Delay =
-    Standard_Thread.delay_last(delay) { later { event } }
 }
--- a/src/Pure/General/csv.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/csv.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -20,7 +20,7 @@
 
   sealed case class Record(fields: Any*)
   {
-    override def toString: String = fields.iterator.map(print_field(_)).mkString(",")
+    override def toString: String = fields.iterator.map(print_field).mkString(",")
   }
 
   sealed case class File(name: String, header: List[String], records: List[Record])
--- a/src/Pure/General/date.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/date.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -33,7 +33,7 @@
     }
 
     def apply(pats: String*): Format =
-      make(pats.toList.map(Date.Formatter.pattern(_)))
+      make(pats.toList.map(Date.Formatter.pattern))
 
     val default: Format = Format("dd-MMM-uuuu HH:mm:ss xx")
     val date: Format = Format("dd-MMM-uuuu")
@@ -56,7 +56,7 @@
     def variants(pats: List[String], locs: List[Locale] = Nil): List[DateTimeFormatter] =
       pats.flatMap(pat => {
         val fmt = pattern(pat)
-        if (locs.isEmpty) List(fmt) else locs.map(fmt.withLocale(_))
+        if (locs.isEmpty) List(fmt) else locs.map(fmt.withLocale)
       })
 
     @tailrec def try_variants(fmts: List[DateTimeFormatter], str: String,
--- a/src/Pure/General/exn.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/exn.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -70,7 +70,7 @@
   def release_first[A](results: List[Result[A]]): List[A] =
     results.find({ case Exn(exn) => !is_interrupt(exn) case _ => false }) match {
       case Some(Exn(exn)) => throw exn
-      case _ => results.map(release(_))
+      case _ => results.map(release)
     }
 
 
@@ -96,22 +96,10 @@
     def apply(): Throwable = new InterruptedException
     def unapply(exn: Throwable): Boolean = is_interrupt(exn)
 
+    def dispose() { Thread.interrupted }
     def expose() { if (Thread.interrupted) throw apply() }
     def impose() { Thread.currentThread.interrupt }
 
-    def postpone[A](body: => A): Option[A] =
-    {
-      val interrupted = Thread.interrupted
-      val result = capture { body }
-      if (interrupted) impose()
-      result match {
-        case Res(x) => Some(x)
-        case Exn(e) =>
-          if (is_interrupt(e)) { impose(); None }
-          else throw e
-      }
-    }
-
     val return_code = 130
   }
 
--- a/src/Pure/General/file.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/file.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -213,7 +213,7 @@
     val line =
       try { reader.readLine}
       catch { case _: IOException => null }
-    if (line == null) None else Some(line)
+    Option(line)
   }
 
   def read_lines(reader: BufferedReader, progress: String => Unit): List[String] =
--- a/src/Pure/General/file_watcher.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/file_watcher.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -84,13 +84,13 @@
 
     /* changed directory entries */
 
-    private val delay_changed = Standard_Thread.delay_last(delay)
+    private val delay_changed = Delay.last(delay)
     {
       val changed = state.change_result(st => (st.changed, st.copy(changed = Set.empty)))
       handle(changed)
     }
 
-    private val watcher_thread = Standard_Thread.fork("File_Watcher", daemon = true)
+    private val watcher_thread = Isabelle_Thread.fork(name = "file_watcher", daemon = true)
     {
       try {
         while (true) {
--- a/src/Pure/General/graph.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/graph.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -123,9 +123,9 @@
     (Map.empty[Key, Long] /: xs)(reach(0))
   }
   def node_height(count: Key => Long): Map[Key, Long] =
-    reachable_length(count, imm_preds(_), maximals)
+    reachable_length(count, imm_preds, maximals)
   def node_depth(count: Key => Long): Map[Key, Long] =
-    reachable_length(count, imm_succs(_), minimals)
+    reachable_length(count, imm_succs, minimals)
 
   /*reachable nodes with size limit*/
   def reachable_limit(
@@ -138,7 +138,7 @@
     {
       val (n, rs) = res
       if (rs.contains(x)) res
-      else ((n + count(x), rs + x) /: next(x))(reach _)
+      else ((n + count(x), rs + x) /: next(x))(reach)
     }
 
     @tailrec def reachs(size: Long, rs: Keys, rest: List[Key]): Keys =
--- a/src/Pure/General/json.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/json.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -9,6 +9,7 @@
 package isabelle
 
 
+import scala.util.matching.Regex
 import scala.util.parsing.combinator.Parsers
 import scala.util.parsing.combinator.lexical.Scanners
 import scala.util.parsing.input.CharArrayReader.EofCh
@@ -61,14 +62,14 @@
     def errorToken(msg: String): Token = Token(Kind.ERROR, msg)
 
     val white_space: String = " \t\n\r"
-    override val whiteSpace = ("[" + white_space + "]+").r
+    override val whiteSpace: Regex = ("[" + white_space + "]+").r
     def whitespace: Parser[Any] = many(character(white_space.contains(_)))
 
-    val letter: Parser[String] = one(character(Symbol.is_ascii_letter(_)))
-    val letters1: Parser[String] = many1(character(Symbol.is_ascii_letter(_)))
+    val letter: Parser[String] = one(character(Symbol.is_ascii_letter))
+    val letters1: Parser[String] = many1(character(Symbol.is_ascii_letter))
 
-    def digits: Parser[String] = many(character(Symbol.is_ascii_digit(_)))
-    def digits1: Parser[String] = many1(character(Symbol.is_ascii_digit(_)))
+    def digits: Parser[String] = many(character(Symbol.is_ascii_digit))
+    def digits1: Parser[String] = many1(character(Symbol.is_ascii_digit))
 
 
     /* keyword */
@@ -114,8 +115,8 @@
       one(character("eE".contains(_))) ~ maybe(character("-+".contains(_))) ~ digits1 ^^
         { case a ~ b ~ c => a + b + c }
 
-    def zero = one(character(c => c == '0'))
-    def nonzero = one(character(c => c != '0' && Symbol.is_ascii_digit(c)))
+    def zero: Parser[String] = one(character(c => c == '0'))
+    def nonzero: Parser[String] = one(character(c => c != '0' && Symbol.is_ascii_digit(c)))
     def positive: Parser[String] = nonzero ~ digits ^^ { case a ~ b => a + b }
 
 
@@ -326,7 +327,7 @@
     object Seconds
     {
       def unapply(json: T): Option[Time] =
-        Double.unapply(json).map(Time.seconds(_))
+        Double.unapply(json).map(Time.seconds)
     }
   }
 
@@ -402,9 +403,9 @@
     : Option[List[A]] = value_default(obj, name, Value.List.unapply(_, unapply), default)
 
   def strings(obj: T, name: String): Option[List[String]] =
-    list(obj, name, JSON.Value.String.unapply _)
+    list(obj, name, JSON.Value.String.unapply)
   def strings_default(obj: T, name: String, default: => List[String] = Nil): Option[List[String]] =
-    list_default(obj, name, JSON.Value.String.unapply _, default)
+    list_default(obj, name, JSON.Value.String.unapply, default)
 
   def seconds(obj: T, name: String): Option[Time] =
     value(obj, name, Value.Seconds.unapply)
--- a/src/Pure/General/linear_set.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/linear_set.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -17,7 +17,7 @@
 object Linear_Set extends SetFactory[Linear_Set]
 {
   private val empty_val: Linear_Set[Nothing] = new Linear_Set[Nothing](None, None, Map(), Map())
-  override def empty[A] = empty_val.asInstanceOf[Linear_Set[A]]
+  override def empty[A]: Linear_Set[A] = empty_val.asInstanceOf[Linear_Set[A]]
 
   implicit def canBuildFrom[A]: CanBuildFrom[Coll, A, Linear_Set[A]] = setCanBuildFrom[A]
   def newBuilder[A]: Builder[A, Linear_Set[A]] = new SetBuilder[A, Linear_Set[A]](empty[A])
--- a/src/Pure/General/mercurial.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/mercurial.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -44,7 +44,7 @@
 
   def find_repository(start: Path, ssh: SSH.System = SSH.Local): Option[Repository] =
   {
-    def find(root: Path): Option[Repository] =
+    @tailrec def find(root: Path): Option[Repository] =
       if (is_repository(root, ssh)) Some(repository(root, ssh = ssh))
       else if (root.is_root) None
       else find(root + Path.parent)
@@ -82,7 +82,7 @@
     val root: Path = ssh.expand_path(root_path)
     def root_url: String = ssh.hg_url + root.implode
 
-    override def toString: String = ssh.prefix + root.implode
+    override def toString: String = ssh.hg_url + root.implode
 
     def command(name: String, args: String = "", options: String = "",
       repository: Boolean = true): Process_Result =
@@ -95,7 +95,7 @@
     }
 
     def add(files: List[Path]): Unit =
-      hg.command("add", files.map(ssh.bash_path(_)).mkString(" "))
+      hg.command("add", files.map(ssh.bash_path).mkString(" "))
 
     def archive(target: String, rev: String = "", options: String = ""): Unit =
       hg.command("archive", opt_rev(rev) + " " + Bash.string(target), options).check
@@ -230,7 +230,7 @@
     remote_name: String = "",
     path_name: String = default_path_name,
     remote_exists: Boolean = false,
-    progress: Progress = No_Progress)
+    progress: Progress = new Progress)
   {
     /* local repository */
 
@@ -274,7 +274,7 @@
 
           while (repos.importing) {
             progress.echo("Awaiting remote repository ...")
-            Thread.sleep(500)
+            Time.seconds(0.5).sleep
             repos = phabricator.the_repository(repos.phid)
           }
 
--- a/src/Pure/General/multi_map.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/multi_map.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -14,7 +14,7 @@
 object Multi_Map extends ImmutableMapFactory[Multi_Map]
 {
   private val empty_val: Multi_Map[Any, Nothing] = new Multi_Map[Any, Nothing](Map.empty)
-  override def empty[A, B] = empty_val.asInstanceOf[Multi_Map[A, B]]
+  override def empty[A, B]: Multi_Map[A, B] = empty_val.asInstanceOf[Multi_Map[A, B]]
 
   implicit def canBuildFrom[A, B]: CanBuildFrom[Coll, (A, B), Multi_Map[A, B]] =
     new MapCanBuildFrom[A, B]
@@ -63,7 +63,7 @@
 
   override def stringPrefix = "Multi_Map"
 
-  override def empty = Multi_Map.empty
+  override def empty: Multi_Map[A, Nothing] = Multi_Map.empty
   override def isEmpty: Boolean = rep.isEmpty
 
   override def keySet: Set[A] = rep.keySet
--- a/src/Pure/General/name_space.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/name_space.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -575,7 +575,7 @@
 fun check context table (xname, pos) =
   let
     val ((name, reports), x) = check_reports context table (xname, [pos]);
-    val _ = Position.reports reports;
+    val _ = Context_Position.reports_generic context reports;
   in (name, x) end;
 
 fun defined (Table (_, tab)) name = Change_Table.defined tab name;
--- a/src/Pure/General/output.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/output.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -59,7 +59,6 @@
   val report_fn: (output list -> unit) Unsynchronized.ref
   val result_fn: (Properties.T -> output list -> unit) Unsynchronized.ref
   val protocol_message_fn: Output_Primitives.protocol_message_fn Unsynchronized.ref
-  val init_channels: unit -> unit
 end;
 
 structure Private_Output: PRIVATE_OUTPUT =
--- a/src/Pure/General/output.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/output.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -15,11 +15,11 @@
 
   def writeln_text(msg: String): String = clean_yxml(msg)
 
-  def warning_text(msg: String): String =
-    Library.prefix_lines("### ", clean_yxml(msg))
+  def warning_prefix(s: String): String = Library.prefix_lines("### ", s)
+  def warning_text(msg: String): String = warning_prefix(clean_yxml(msg))
 
-  def error_message_text(msg: String): String =
-    Library.prefix_lines("*** ", clean_yxml(msg))
+  def error_prefix(s: String): String = Library.prefix_lines("*** ", s)
+  def error_message_text(msg: String): String = error_prefix(clean_yxml(msg))
 
   def writeln(msg: String, stdout: Boolean = false, include_empty: Boolean = false)
   {
--- a/src/Pure/General/output_primitives.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/output_primitives.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -49,6 +49,7 @@
   | Text of string;
 
 type body = tree list;
+
 end;
 
 
--- a/src/Pure/General/path.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/path.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -73,7 +73,7 @@
   val current: Path = new Path(Nil)
   val root: Path = new Path(List(Root("")))
   def named_root(s: String): Path = new Path(List(root_elem(s)))
-  def make(elems: List[String]): Path = new Path(elems.reverse.map(basic_elem(_)))
+  def make(elems: List[String]): Path = new Path(elems.reverse.map(basic_elem))
   def basic(s: String): Path = new Path(List(basic_elem(s)))
   def variable(s: String): Path = new Path(List(variable_elem(s)))
   val parent: Path = new Path(List(Parent))
--- a/src/Pure/General/pretty.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/pretty.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -6,6 +6,8 @@
 
 package isabelle
 
+import scala.annotation.tailrec
+
 
 object Pretty
 {
@@ -69,7 +71,7 @@
   {
     val indent1 = force_nat(indent)
 
-    def body_length(prts: List[Tree], len: Double): Double =
+    @tailrec def body_length(prts: List[Tree], len: Double): Double =
     {
       val (line, rest) =
         Library.take_prefix[Tree]({ case Break(true, _, _) => false case _ => true }, prts)
@@ -104,7 +106,7 @@
 
   private def force_break(tree: Tree): Tree =
     tree match { case Break(false, wd, ind) => Break(true, wd, ind) case _ => tree }
-  private def force_all(trees: List[Tree]): List[Tree] = trees.map(force_break(_))
+  private def force_all(trees: List[Tree]): List[Tree] = trees.map(force_break)
 
   private def force_next(trees: List[Tree]): List[Tree] =
     trees match {
@@ -113,8 +115,8 @@
       case t :: ts => t :: force_next(ts)
     }
 
-  val default_margin = 76.0
-  val default_breakgain = default_margin / 20
+  val default_margin: Double = 76.0
+  val default_breakgain: Double = default_margin / 20
 
   def formatted(input: XML.Body,
     margin: Double = default_margin,
--- a/src/Pure/General/properties.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/properties.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -67,7 +67,7 @@
         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(_))
+        case Some(cache) => ps.map(cache.props)
       }
     }
   }
--- a/src/Pure/General/scan.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/scan.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -10,10 +10,10 @@
 import scala.annotation.tailrec
 import scala.collection.{IndexedSeq, Traversable, TraversableOnce}
 import scala.collection.immutable.PagedSeq
+import scala.util.matching.Regex
 import scala.util.parsing.input.{OffsetPosition, Position => InputPosition,
   Reader, CharSequenceReader}
 import scala.util.parsing.combinator.RegexParsers
-
 import java.io.{File => JFile, BufferedInputStream, FileInputStream, InputStream}
 import java.net.URL
 
@@ -39,7 +39,7 @@
 
   trait Parsers extends RegexParsers
   {
-    override val whiteSpace = "".r
+    override val whiteSpace: Regex = "".r
 
 
     /* optional termination */
@@ -237,7 +237,7 @@
     {
       require(depth >= 0)
 
-      val comment_text =
+      val comment_text: Parser[List[String]] =
         rep1(many1(sym => sym != "*" && sym != "(") | """\*(?!\))|\((?!\*)""".r)
 
       def apply(in: Input) =
@@ -399,7 +399,7 @@
       else this ++ other.raw_iterator
 
     def -- (remove: Traversable[String]): Lexicon =
-      if (remove.exists(contains(_)))
+      if (remove.exists(contains))
         Lexicon.empty ++ iterator.filterNot(a => remove.exists(b => a == b))
       else this
 
--- a/src/Pure/General/sql.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/sql.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -37,7 +37,7 @@
     }
 
   def string(s: String): Source =
-    s.iterator.map(escape_char(_)).mkString("'", "", "'")
+    s.iterator.map(escape_char).mkString("'", "", "'")
 
   def ident(s: String): Source =
     Long_Name.implode(Long_Name.explode(s).map(a => quote(a.replace("\"", "\"\""))))
@@ -288,7 +288,7 @@
     }
     def date(column: Column): Date = stmt.db.date(res, column)
 
-    def timing(c1: Column, c2: Column, c3: Column) =
+    def timing(c1: Column, c2: Column, c3: Column): Timing =
       Timing(Time.ms(long(c1)), Time.ms(long(c2)), Time.ms(long(c3)))
 
     def get[A](column: Column, f: Column => A): Option[A] =
@@ -296,13 +296,13 @@
       val x = f(column)
       if (rep.wasNull) None else Some(x)
     }
-    def get_bool(column: Column): Option[Boolean] = get(column, bool _)
-    def get_int(column: Column): Option[Int] = get(column, int _)
-    def get_long(column: Column): Option[Long] = get(column, long _)
-    def get_double(column: Column): Option[Double] = get(column, double _)
-    def get_string(column: Column): Option[String] = get(column, string _)
-    def get_bytes(column: Column): Option[Bytes] = get(column, bytes _)
-    def get_date(column: Column): Option[Date] = get(column, date _)
+    def get_bool(column: Column): Option[Boolean] = get(column, bool)
+    def get_int(column: Column): Option[Int] = get(column, int)
+    def get_long(column: Column): Option[Long] = get(column, long)
+    def get_double(column: Column): Option[Double] = get(column, double)
+    def get_string(column: Column): Option[String] = get(column, string)
+    def get_bytes(column: Column): Option[Bytes] = get(column, bytes)
+    def get_date(column: Column): Option[Date] = get(column, date)
   }
 
 
--- a/src/Pure/General/ssh.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/ssh.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -10,6 +10,7 @@
 import java.io.{InputStream, OutputStream, ByteArrayOutputStream}
 
 import scala.collection.mutable
+import scala.util.matching.Regex
 
 import com.jcraft.jsch.{JSch, Logger => JSch_Logger, Session => JSch_Session, SftpException,
   OpenSSHConfig, UserInfo, Channel => JSch_Channel, ChannelExec, ChannelSftp, SftpATTRS}
@@ -21,7 +22,7 @@
 
   object Target
   {
-    val User_Host = "^([^@]+)@(.+)$".r
+    val User_Host: Regex = "^([^@]+)@(.+)$".r
 
     def parse(s: String): (String, String) =
       s match {
@@ -76,17 +77,19 @@
     jsch.setKnownHosts(File.platform_path(known_hosts))
 
     val identity_files =
-      space_explode(':', options.string("ssh_identity_files")).map(Path.explode(_))
+      space_explode(':', options.string("ssh_identity_files")).map(Path.explode)
     for (identity_file <- identity_files if identity_file.is_file)
       jsch.addIdentity(File.platform_path(identity_file))
 
     new Context(options, jsch)
   }
 
-  def open_session(options: Options, host: String, user: String = "", port: Int = 0,
+  def open_session(options: Options,
+      host: String, user: String = "", port: Int = 0, actual_host: String = "",
       proxy_host: String = "", proxy_user: String = "", proxy_port: Int = 0,
       permissive: Boolean = false): Session =
-    init_context(options).open_session(host = host, user = user, port = port,
+    init_context(options).open_session(
+      host = host, user = user, port = port, actual_host = actual_host,
       proxy_host = proxy_host, proxy_user = proxy_user, proxy_port = proxy_port,
       permissive = permissive)
 
@@ -120,16 +123,18 @@
         proper_string(nominal_user) getOrElse user)
     }
 
-    def open_session(host: String, user: String = "", port: Int = 0,
+    def open_session(
+      host: String, user: String = "", port: Int = 0, actual_host: String = "",
       proxy_host: String = "", proxy_user: String = "", proxy_port: Int = 0,
       permissive: Boolean = false): Session =
     {
-      if (proxy_host == "") connect_session(host = host, user = user, port = port)
+      val connect_host = proper_string(actual_host) getOrElse host
+      if (proxy_host == "") connect_session(host = connect_host, user = user, port = port)
       else {
         val proxy = connect_session(host = proxy_host, port = proxy_port, user = proxy_user)
 
         val fw =
-          try { proxy.port_forwarding(remote_host = host, remote_port = make_port(port)) }
+          try { proxy.port_forwarding(remote_host = connect_host, remote_port = make_port(port)) }
           catch { case exn: Throwable => proxy.close; throw exn }
 
         try {
@@ -232,7 +237,7 @@
 
     val exit_status: Future[Int] =
       Future.thread("ssh_wait") {
-        while (!channel.isClosed) Thread.sleep(exec_wait_delay.ms)
+        while (!channel.isClosed) exec_wait_delay.sleep
         channel.getExitStatus
       }
 
@@ -271,7 +276,7 @@
             if (line_buffer.size > 0) line_flush()
             finished = true
           }
-          else Thread.sleep(exec_wait_delay.ms)
+          else exec_wait_delay.sleep
         }
 
         result.toList
@@ -317,9 +322,6 @@
     override def hg_url: String =
       "ssh://" + user_prefix(nominal_user) + nominal_host + "/"
 
-    override def prefix: String =
-      user_prefix(session.getUserName) + host + port_suffix(session.getPort) + ":"
-
     override def toString: String =
       user_prefix(session.getUserName) + host + port_suffix(session.getPort) +
       (if (session.isConnected) "" else " (disconnected)")
@@ -430,12 +432,12 @@
     def read_file(path: Path, local_path: Path): Unit =
       sftp.get(remote_path(path), File.platform_path(local_path))
     def read_bytes(path: Path): Bytes = using(open_input(path))(Bytes.read_stream(_))
-    def read(path: Path): String = using(open_input(path))(File.read_stream(_))
+    def read(path: Path): String = using(open_input(path))(File.read_stream)
 
     def write_file(path: Path, local_path: Path): Unit =
       sftp.put(File.platform_path(local_path), remote_path(path))
     def write_bytes(path: Path, bytes: Bytes): Unit =
-      using(open_output(path))(bytes.write_stream(_))
+      using(open_output(path))(bytes.write_stream)
     def write(path: Path, text: String): Unit =
       using(open_output(path))(stream => Bytes(text).write_stream(stream))
 
@@ -480,7 +482,6 @@
   trait System
   {
     def hg_url: String = ""
-    def prefix: String = ""
 
     def expand_path(path: Path): Path = path.expand
     def bash_path(path: Path): String = File.bash_path(path)
--- a/src/Pure/General/symbol.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/symbol.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -501,10 +501,9 @@
 (* length *)
 
 fun sym_len s =
-  if not (is_printable s) then (0: int)
-  else if String.isPrefix "\092<long" s then 2
-  else if String.isPrefix "\092<Long" s then 2
-  else 1;
+  if String.isPrefix "\092<long" s orelse String.isPrefix "\092<Long" s then 2
+  else if is_printable s then 1
+  else 0;
 
 fun sym_length ss = fold (fn s => fn n => sym_len s + n) ss 0;
 
--- a/src/Pure/General/symbol.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/symbol.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -23,6 +23,7 @@
 
   /* spaces */
 
+  val space_char = ' '
   val space = " "
 
   private val static_spaces = space * 4000
@@ -37,6 +38,8 @@
 
   /* ASCII characters */
 
+  def is_ascii_printable(c: Char): Boolean = space_char <= c && c <= '~'
+
   def is_ascii_letter(c: Char): Boolean = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z'
 
   def is_ascii_digit(c: Char): Boolean = '0' <= c && c <= '9'
@@ -113,8 +116,8 @@
     {
       private val matcher = new Matcher(text)
       private var i = 0
-      def hasNext = i < text.length
-      def next =
+      def hasNext: Boolean = i < text.length
+      def next: Symbol =
       {
         val n = matcher(i, text.length)
         val s =
@@ -135,10 +138,10 @@
   def length(text: CharSequence): Int = iterator(text).length
 
   def trim_blanks(text: CharSequence): String =
-    Library.trim(is_blank(_), explode(text)).mkString
+    Library.trim(is_blank, explode(text)).mkString
 
   def all_blank(str: String): Boolean =
-    iterator(str).forall(is_blank(_))
+    iterator(str).forall(is_blank)
 
   def trim_blank_lines(text: String): String =
     cat_lines(split_lines(text).dropWhile(all_blank).reverse.dropWhile(all_blank).reverse)
@@ -191,7 +194,7 @@
       if (i < 0) sym
       else index(i).chr + sym - index(i).sym
     }
-    def decode(symbol_range: Range): Text.Range = symbol_range.map(decode(_))
+    def decode(symbol_range: Range): Text.Range = symbol_range.map(decode)
 
     override def hashCode: Int = hash
     override def equals(that: Any): Boolean =
@@ -448,7 +451,7 @@
 
     /* classification */
 
-    val letters = recode_set(
+    val letters: Set[String] = recode_set(
       "A", "B", "C", "D", "E", "F", "G", "H", "I", "J", "K", "L", "M",
       "N", "O", "P", "Q", "R", "S", "T", "U", "V", "W", "X", "Y", "Z",
       "a", "b", "c", "d", "e", "f", "g", "h", "i", "j", "k", "l", "m",
@@ -481,12 +484,12 @@
       "\\<Xi>", "\\<Pi>", "\\<Sigma>", "\\<Upsilon>", "\\<Phi>",
       "\\<Psi>", "\\<Omega>")
 
-    val blanks = recode_set(space, "\t", "\n", "\u000B", "\f", "\r", "\r\n")
+    val blanks: Set[String] = recode_set(space, "\t", "\n", "\u000B", "\f", "\r", "\r\n")
 
     val sym_chars =
       Set("!", "#", "$", "%", "&", "*", "+", "-", "/", "<", "=", ">", "?", "@", "^", "_", "|", "~")
 
-    val symbolic = recode_set((for { (sym, _) <- symbols; if raw_symbolic(sym) } yield sym): _*)
+    val symbolic: Set[String] = recode_set((for {(sym, _) <- symbols; if raw_symbolic(sym)} yield sym): _*)
 
 
     /* misc symbols */
@@ -654,4 +657,23 @@
   def esub_decoded: Symbol = symbols.esub_decoded
   def bsup_decoded: Symbol = symbols.bsup_decoded
   def esup_decoded: Symbol = symbols.esup_decoded
+
+
+  /* metric */
+
+  def is_printable(sym: Symbol): Boolean =
+    if (is_ascii(sym)) is_ascii_printable(sym(0))
+    else !is_control(sym)
+
+  object Metric extends Pretty.Metric
+  {
+    val unit = 1.0
+    def apply(str: String): Double =
+      (for (s <- iterator(str)) yield {
+        val sym = encode(s)
+        if (sym.startsWith("\\<long") || sym.startsWith("\\<Long")) 2
+        else if (is_printable(sym)) 1
+        else 0
+      }).sum
+  }
 }
--- a/src/Pure/General/time.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/time.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -65,4 +65,6 @@
   }
 
   def instant: Instant = Instant.ofEpochMilli(ms)
+
+  def sleep { Thread.sleep(ms) }
 }
--- a/src/Pure/General/url.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/url.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -24,7 +24,7 @@
     }
     else c.toString
 
-  def escape_special(s: String): String = s.iterator.map(escape_special(_)).mkString
+  def escape_special(s: String): String = s.iterator.map(escape_special).mkString
 
   def escape_name(name: String): String =
     name.iterator.map({ case '\'' => "%27" case c => c.toString }).mkString
--- a/src/Pure/General/value.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/value.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -70,7 +70,7 @@
       val s = t.seconds
       if (s.toInt.toDouble == s) s.toInt.toString else t.toString
     }
-    def unapply(s: java.lang.String): Option[Time] = Double.unapply(s).map(Time.seconds(_))
+    def unapply(s: java.lang.String): Option[Time] = Double.unapply(s).map(Time.seconds)
     def parse(s: java.lang.String): Time =
       unapply(s) getOrElse error("Bad real (for seconds): " + quote(s))
   }
--- a/src/Pure/General/word.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/General/word.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -53,11 +53,11 @@
       }
     def unapply(str: String): Option[Case] =
       if (str.nonEmpty) {
-        if (Codepoint.iterator(str).forall(Character.isLowerCase(_))) Some(Lowercase)
-        else if (Codepoint.iterator(str).forall(Character.isUpperCase(_))) Some(Uppercase)
+        if (Codepoint.iterator(str).forall(Character.isLowerCase)) Some(Lowercase)
+        else if (Codepoint.iterator(str).forall(Character.isUpperCase)) Some(Uppercase)
         else {
           val it = Codepoint.iterator(str)
-          if (Character.isUpperCase(it.next) && it.forall(Character.isLowerCase(_)))
+          if (Character.isUpperCase(it.next) && it.forall(Character.isLowerCase))
             Some(Capitalized)
           else None
         }
--- a/src/Pure/Isar/keyword.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Isar/keyword.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -169,7 +169,7 @@
       val kinds1 = kinds + (name -> kind)
       val load_commands1 =
         if (kind == THY_LOAD) {
-          if (!Symbol.iterator(name).forall(Symbol.is_ascii(_)))
+          if (!Symbol.iterator(name).forall(Symbol.is_ascii))
             error("Bad theory load command " + quote(name))
           load_commands + (name -> exts)
         }
--- a/src/Pure/Isar/line_structure.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Isar/line_structure.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -9,7 +9,7 @@
 
 object Line_Structure
 {
-  val init = Line_Structure()
+  val init: Line_Structure = Line_Structure()
 }
 
 sealed case class Line_Structure(
--- a/src/Pure/Isar/method.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Isar/method.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -659,7 +659,9 @@
     maps (fn p => map (pair p) (Markup.keyword3 :: Completion.suppress_abbrevs ""))
       (keyword_positions text);
 
-val report = Position.reports o reports_of;
+fun report text_range =
+  if Context_Position.pide_reports ()
+  then Position.reports (reports_of text_range) else ();
 
 end;
 
--- a/src/Pure/Isar/outer_syntax.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Isar/outer_syntax.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -193,7 +193,7 @@
       if (tok.is_ignored) ignored += tok
       else if (keywords.is_before_command(tok) ||
         tok.is_command &&
-          (!content.exists(keywords.is_before_command(_)) || content.exists(_.is_command)))
+          (!content.exists(keywords.is_before_command) || content.exists(_.is_command)))
       { flush(); content += tok }
       else { content ++= ignored; ignored.clear; content += tok }
     }
--- a/src/Pure/Isar/proof_context.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Isar/proof_context.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -475,7 +475,7 @@
   let
     val source = Syntax.read_input text;
     val (c, reports) = check_class ctxt (Input.source_content source);
-    val _ = Position.reports reports;
+    val _ = Context_Position.reports ctxt reports;
   in c end;
 
 
@@ -536,11 +536,11 @@
         | Type.Nonterminal => if strict then err () else 0);
     in (Type (d, replicate args dummyT), reports) end;
 
-fun read_type_name ctxt flags text =
+fun read_type_name flags ctxt text =
   let
     val source = Syntax.read_input text;
-    val (T, reports) = check_type_name ctxt flags (Input.source_content source);
-    val _ = Position.reports reports;
+    val (T, reports) = check_type_name flags ctxt (Input.source_content source);
+    val _ = Context_Position.reports ctxt reports;
   in T end;
 
 
@@ -583,12 +583,12 @@
       | _ => ());
   in (t, reports) end;
 
-fun read_const ctxt flags text =
+fun read_const flags ctxt text =
   let
     val source = Syntax.read_input text;
     val (c, pos) = Input.source_content source;
-    val (t, reports) = check_const ctxt flags (c, [pos]);
-    val _ = Position.reports reports;
+    val (t, reports) = check_const flags ctxt (c, [pos]);
+    val _ = Context_Position.reports ctxt reports;
   in t end;
 
 
@@ -792,7 +792,7 @@
 fun check_tfree ctxt v =
   let
     val (sorting_report, [TFree a]) = prepare_sortsT ctxt [TFree v];
-    val _ = if Context_Position.is_visible ctxt then Output.report sorting_report else ();
+    val _ = if Context_Position.reports_enabled ctxt then Output.report sorting_report else ();
   in a end;
 
 end;
@@ -1103,7 +1103,7 @@
   let
     val ([x], ctxt') = Variable.add_fixes_binding [Binding.reset_pos b] ctxt;
     val mx' = Mixfix.reset_pos mx;
-    val _ = add_syntax [(x, T, if Context_Position.is_visible ctxt then mx else mx')] ctxt';
+    val _ = add_syntax [(x, T, if Context_Position.reports_enabled ctxt then mx else mx')] ctxt';
   in mx' end;
 
 fun prep_var prep_typ internal (b, raw_T, mx) ctxt =
--- a/src/Pure/Isar/specification.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Isar/specification.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -323,7 +323,7 @@
 fun gen_alias decl check (b, arg) lthy =
   let
     val (c, reports) = check {proper = true, strict = false} lthy arg;
-    val _ = Position.reports reports;
+    val _ = Context_Position.reports lthy reports;
   in decl b c lthy end;
 
 val alias =
--- a/src/Pure/Isar/token.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Isar/token.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -760,7 +760,7 @@
     val old_reports = maps reports_of_value src;
     val args1 = map init_assignable (args_of_src src);
     fun reported_text () =
-      if Context_Position.is_visible_generic context then
+      if Context_Position.reports_enabled_generic context then
         let val new_reports = maps (reports_of_value o closure) args1 in
           if old_reports <> new_reports then
             map (fn (p, m) => Position.reported_text p m "") new_reports
--- a/src/Pure/Isar/token.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Isar/token.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -249,9 +249,9 @@
 
   private class Token_Reader(tokens: List[Token], val pos: Pos) extends Reader
   {
-    def first = tokens.head
-    def rest = new Token_Reader(tokens.tail, pos.advance(first))
-    def atEnd = tokens.isEmpty
+    def first: Token = tokens.head
+    def rest: Token_Reader = new Token_Reader(tokens.tail, pos.advance(first))
+    def atEnd: Boolean = tokens.isEmpty
   }
 
   def reader(tokens: List[Token], start: Token.Pos): Reader =
@@ -321,8 +321,8 @@
     source.startsWith(Symbol.open) ||
     source.startsWith(Symbol.open_decoded))
 
-  def is_open_bracket: Boolean = is_keyword && Word.open_brackets.exists(is_keyword(_))
-  def is_close_bracket: Boolean = is_keyword && Word.close_brackets.exists(is_keyword(_))
+  def is_open_bracket: Boolean = is_keyword && Word.open_brackets.exists(is_keyword)
+  def is_close_bracket: Boolean = is_keyword && Word.close_brackets.exists(is_keyword)
 
   def is_begin: Boolean = is_keyword("begin")
   def is_end: Boolean = is_command("end")
@@ -341,7 +341,7 @@
   {
     val s = content
     is_name && Path.is_wellformed(s) &&
-      !s.exists(Symbol.is_ascii_blank(_)) &&
+      !s.exists(Symbol.is_ascii_blank) &&
       !Path.is_reserved(s)
   }
 }
--- a/src/Pure/ML/ml_compiler.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/ML/ml_compiler.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -50,11 +50,11 @@
 
 fun report_parse_tree redirect depth name_space parse_tree =
   let
-    val is_visible =
+    val reports_enabled =
       (case Context.get_generic_context () of
-        SOME context => Context_Position.is_visible_generic context
-      | NONE => true);
-    fun is_reported pos = is_visible andalso Position.is_reported pos;
+        SOME context => Context_Position.reports_enabled_generic context
+      | NONE => Context_Position.pide_reports ());
+    fun is_reported pos = reports_enabled andalso Position.is_reported pos;
 
 
     (* syntax reports *)
--- a/src/Pure/ML/ml_console.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/ML/ml_console.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -50,6 +50,9 @@
       if (more_args.nonEmpty) getopts.usage()
 
 
+      val sessions_structure = Sessions.load_structure(options, dirs = dirs)
+      val store = Sessions.store(options)
+
       // build logic
       if (!no_build && !raw_ml_system) {
         val progress = new Console_Progress()
@@ -62,23 +65,20 @@
 
       // process loop
       val process =
-        ML_Process(options, logic = logic, args = List("-i"), dirs = dirs, redirect = true,
+        ML_Process(options, sessions_structure, store,
+          logic = logic, args = List("-i"), redirect = true,
           modes = if (raw_ml_system) Nil else modes ::: List("ASCII"),
           raw_ml_system = raw_ml_system,
-          store = Some(Sessions.store(options)),
           session_base =
             if (raw_ml_system) None
             else Some(Sessions.base_info(
               options, logic, dirs = dirs, include_sessions = include_sessions).check_base))
 
-      val tty_loop = new TTY_Loop(process.stdin, process.stdout, Some(process.interrupt _))
-      val process_result = Future.thread[Int]("process_result") {
+      POSIX_Interrupt.handler { process.interrupt } {
+        new TTY_Loop(process.stdin, process.stdout).join
         val rc = process.join
-        tty_loop.cancel  // FIXME does not quite work, cannot interrupt blocking read on System.in
-        rc
+        if (rc != 0) sys.exit(rc)
       }
-      tty_loop.join
-      process_result.join
     }
   }
 }
--- a/src/Pure/ML/ml_process.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/ML/ml_process.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -13,31 +13,28 @@
 object ML_Process
 {
   def apply(options: Options,
+    sessions_structure: Sessions.Structure,
+    store: Sessions.Store,
     logic: String = "",
+    raw_ml_system: Boolean = false,
+    use_prelude: List[String] = Nil,
+    eval_main: String = "",
     args: List[String] = Nil,
-    dirs: List[Path] = Nil,
     modes: List[String] = Nil,
-    raw_ml_system: Boolean = false,
     cwd: JFile = null,
     env: Map[String, String] = Isabelle_System.settings(),
     redirect: Boolean = false,
     cleanup: () => Unit = () => (),
-    sessions_structure: Option[Sessions.Structure] = None,
-    session_base: Option[Sessions.Base] = None,
-    store: Option[Sessions.Store] = None): Bash.Process =
+    session_base: Option[Sessions.Base] = None): Bash.Process =
   {
     val logic_name = Isabelle_System.default_logic(logic)
-    val _store = store.getOrElse(Sessions.store(options))
-
-    val sessions_structure1 =
-      sessions_structure.getOrElse(Sessions.load_structure(options, dirs = dirs))
 
     val heaps: List[String] =
       if (raw_ml_system) Nil
       else {
-        sessions_structure1.selection(Sessions.Selection.session(logic_name)).
+        sessions_structure.selection(Sessions.Selection.session(logic_name)).
           build_requirements(List(logic_name)).
-          map(a => File.platform_path(_store.the_heap(a)))
+          map(a => File.platform_path(store.the_heap(a)))
       }
 
     val eval_init =
@@ -96,8 +93,8 @@
               ML_Syntax.print_pair(ML_Syntax.print_string_bytes, ML_Syntax.print_properties))(list)
 
           List("Resources.init_session_base" +
-            " {session_positions = " + print_sessions(sessions_structure1.session_positions) +
-            ", session_directories = " + print_table(sessions_structure1.dest_session_directories) +
+            " {session_positions = " + print_sessions(sessions_structure.session_positions) +
+            ", session_directories = " + print_table(sessions_structure.dest_session_directories) +
             ", docs = " + print_list(base.doc_names) +
             ", global_theories = " + print_table(base.global_theories.toList) +
             ", loaded_theories = " + print_list(base.loaded_theories.keys) + "}")
@@ -105,9 +102,11 @@
 
     // process
     val eval_process =
-      if (heaps.isEmpty)
-        List("PolyML.print_depth " + ML_Syntax.print_int(options.int("ML_print_depth")))
-      else List("Isabelle_Process.init ()")
+      proper_string(eval_main).getOrElse(
+        if (heaps.isEmpty) {
+          "PolyML.print_depth " + ML_Syntax.print_int(options.int("ML_print_depth"))
+        }
+        else "Isabelle_Process.init ()")
 
     // ISABELLE_TMP
     val isabelle_tmp = Isabelle_System.tmp_dir("process")
@@ -128,8 +127,8 @@
     // bash
     val bash_args =
       ml_runtime_options :::
-      (eval_init ::: eval_modes ::: eval_options ::: eval_session_base ::: eval_process)
-        .flatMap(eval => List("--eval", eval)) ::: args
+      (eval_init ::: eval_modes ::: eval_options ::: eval_session_base).flatMap(List("--eval", _)) :::
+      use_prelude.flatMap(List("--use", _)) ::: List("--eval", eval_process) ::: args
 
     Bash.process(
       "exec " + options.string("ML_process_policy") + """ "$ML_HOME/poly" -q """ +
@@ -182,8 +181,12 @@
     val more_args = getopts(args)
     if (args.isEmpty || more_args.nonEmpty) getopts.usage()
 
-    val rc = ML_Process(options, logic = logic, args = eval_args, dirs = dirs, modes = modes).
-      result().print_stdout.rc
+    val sessions_structure = Sessions.load_structure(options, dirs = dirs)
+    val store = Sessions.store(options)
+
+    val rc =
+      ML_Process(options, sessions_structure, store, logic = logic, args = eval_args, modes = modes)
+        .result().print_stdout.rc
     sys.exit(rc)
   })
 }
--- a/src/Pure/ML/ml_statistics.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/ML/ml_statistics.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -208,7 +208,7 @@
     chart(fields._1, fields._2)
 
   def show_frames(fields: List[ML_Statistics.Fields] = ML_Statistics.main_fields): Unit =
-    fields.map(chart(_)).foreach(c =>
+    fields.map(chart).foreach(c =>
       GUI_Thread.later {
         new Frame {
           iconImage = GUI.isabelle_image()
--- a/src/Pure/ML/ml_syntax.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/ML/ml_syntax.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -37,13 +37,13 @@
 
   private def print_symbol(s: Symbol.Symbol): String =
     if (s.startsWith("\\<")) s
-    else UTF8.bytes(s).iterator.map(print_byte(_)).mkString
+    else UTF8.bytes(s).iterator.map(print_byte).mkString
 
   def print_string_bytes(str: String): String =
-    quote(UTF8.bytes(str).iterator.map(print_byte(_)).mkString)
+    quote(UTF8.bytes(str).iterator.map(print_byte).mkString)
 
   def print_string_symbols(str: String): String =
-    quote(Symbol.iterator(str).map(print_symbol(_)).mkString)
+    quote(Symbol.iterator(str).map(print_symbol).mkString)
 
 
   /* pair */
@@ -61,5 +61,5 @@
   /* properties */
 
   def print_properties(props: Properties.T): String =
-    print_list(print_pair(print_string_bytes(_), print_string_bytes(_))(_))(props)
+    print_list(print_pair(print_string_bytes, print_string_bytes))(props)
 }
--- a/src/Pure/PIDE/command.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/command.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -58,7 +58,9 @@
 
 fun read_file_node file_node master_dir pos src_path =
   let
-    val _ = Position.report pos Markup.language_path;
+    val _ =
+      if Context_Position.pide_reports ()
+      then Position.report pos Markup.language_path else ();
     val _ =
       (case try Url.explode file_node of
         NONE => ()
--- a/src/Pure/PIDE/document.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/document.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -603,7 +603,8 @@
 
     val parents =
       if null parents_reports then [Theory.get_pure ()] else map #1 parents_reports;
-    val _ = Position.reports (map #2 parents_reports);
+    val _ = List.app (fn (thy, r) => Context_Position.reports_global thy [r]) parents_reports;
+
     val thy = Resources.begin_theory master_dir header parents;
     val _ = Output.status [Markup.markup_only Markup.initialized];
   in thy end;
--- a/src/Pure/PIDE/document.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/document.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -95,12 +95,12 @@
         copy(errors = errors.map(msg1 => Exn.cat_message(msg1, msg2)))
     }
 
-    val no_header = Header()
+    val no_header: Header = Header()
     def bad_header(msg: String): Header = Header(errors = List(msg))
 
     object Name
     {
-      val empty = Name("")
+      val empty: Name = Name("")
 
       object Ordering extends scala.math.Ordering[Name]
       {
@@ -375,7 +375,7 @@
     }
 
     def purge_suppressed: Option[Nodes] =
-      graph.keys_iterator.filter(is_suppressed(_)).toList match {
+      graph.keys_iterator.filter(is_suppressed).toList match {
         case Nil => None
         case del => Some(new Nodes((graph /: del)(_.del_node(_))))
       }
@@ -527,7 +527,7 @@
 
   object Snapshot
   {
-    val init = State.init.snapshot()
+    val init: Snapshot = State.init.snapshot()
   }
 
   abstract class Snapshot
@@ -737,13 +737,13 @@
     {
       execs.get(id) match {
         case Some(st) =>
-          val new_st = st.accumulate(self_id(st), other_id _, message, xml_cache)
+          val new_st = st.accumulate(self_id(st), other_id, message, xml_cache)
           val execs1 = execs + (id -> new_st)
           (new_st, copy(execs = execs1, commands_redirection = redirection(new_st)))
         case None =>
           commands.get(id) match {
             case Some(st) =>
-              val new_st = st.accumulate(self_id(st), other_id _, message, xml_cache)
+              val new_st = st.accumulate(self_id(st), other_id, message, xml_cache)
               val commands1 = commands + (id -> new_st)
               (new_st, copy(commands = commands1, commands_redirection = redirection(new_st)))
             case None => fail
@@ -1039,11 +1039,11 @@
 
         /* local node content */
 
-        def convert(offset: Text.Offset) = (offset /: edits)((i, edit) => edit.convert(i))
-        def revert(offset: Text.Offset) = (offset /: reverse_edits)((i, edit) => edit.revert(i))
+        def convert(offset: Text.Offset): Text.Offset = (offset /: edits)((i, edit) => edit.convert(i))
+        def revert(offset: Text.Offset): Text.Offset = (offset /: reverse_edits)((i, edit) => edit.revert(i))
 
-        def convert(range: Text.Range) = range.map(convert(_))
-        def revert(range: Text.Range) = range.map(revert(_))
+        def convert(range: Text.Range): Text.Range = range.map(convert)
+        def revert(range: Text.Range): Text.Range = range.map(revert)
 
         val node_name: Node.Name = name
         val node: Node = version.nodes(name)
@@ -1166,7 +1166,7 @@
                 case some => Some(some)
               }
           }
-          for (Text.Info(r, Some(x)) <- cumulate(range, None, elements, result1 _, status))
+          for (Text.Info(r, Some(x)) <- cumulate(range, None, elements, result1, status))
             yield Text.Info(r, x)
         }
 
--- a/src/Pure/PIDE/document_id.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/document_id.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -17,7 +17,7 @@
   type Exec = Generic
 
   val none: Generic = 0
-  val make = Counter.make()
+  val make: Counter = Counter.make()
 
   def apply(id: Generic): String = Value.Long.apply(id)
   def unapply(s: String): Option[Generic] = Value.Long.unapply(s)
--- a/src/Pure/PIDE/document_status.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/document_status.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -55,7 +55,7 @@
         runs = runs)
     }
 
-    val empty = make(Iterator.empty)
+    val empty: Command_Status = make(Iterator.empty)
 
     def merge(status_iterator: Iterator[Command_Status]): Command_Status =
       if (status_iterator.hasNext) {
--- a/src/Pure/PIDE/headless.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/headless.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -224,7 +224,7 @@
           }
           else result
 
-        val (load_theories, load_state1) = load_state.next(dep_graph, finished_theory(_))
+        val (load_theories, load_state1) = load_state.next(dep_graph, finished_theory)
 
         (load_theories,
           copy(already_committed = already_committed1, result = result1, load_state = load_state1))
@@ -244,7 +244,7 @@
       // commit: must not block, must not fail
       commit: Option[(Document.Snapshot, Document_Status.Node_Status) => Unit] = None,
       commit_cleanup_delay: Time = default_commit_cleanup_delay,
-      progress: Progress = No_Progress): Use_Theories_Result =
+      progress: Progress = new Progress): Use_Theories_Result =
     {
       val dependencies =
       {
@@ -304,12 +304,12 @@
       val consumer =
       {
         val delay_nodes_status =
-          Standard_Thread.delay_first(nodes_status_delay max Time.zero) {
+          Delay.first(nodes_status_delay max Time.zero) {
             progress.nodes_status(use_theories_state.value.nodes_status)
           }
 
         val delay_commit_clean =
-          Standard_Thread.delay_first(commit_cleanup_delay max Time.zero) {
+          Delay.first(commit_cleanup_delay max Time.zero) {
             val clean_theories = use_theories_state.change_result(_.clean_theories)
             if (clean_theories.nonEmpty) {
               progress.echo("Removing " + clean_theories.length + " theories ...")
@@ -350,7 +350,7 @@
                     (theory_progress, st.update(nodes_status1))
                   })
 
-              theory_progress.foreach(progress.theory(_))
+              theory_progress.foreach(progress.theory)
 
               check_state()
 
@@ -396,21 +396,21 @@
 
   object Resources
   {
-    def apply(base_info: Sessions.Base_Info, log: Logger = No_Logger): Resources =
-      new Resources(base_info, log = log)
+    def apply(options: Options, base_info: Sessions.Base_Info, log: Logger = No_Logger): Resources =
+      new Resources(options, base_info, log = log)
 
     def make(
       options: Options,
       session_name: String,
       session_dirs: List[Path] = Nil,
       include_sessions: List[String] = Nil,
-      progress: Progress = No_Progress,
+      progress: Progress = new Progress,
       log: Logger = No_Logger): Resources =
     {
       val base_info =
         Sessions.base_info(options, session_name, dirs = session_dirs,
           include_sessions = include_sessions, progress = progress)
-      apply(base_info, log = log)
+      apply(options, base_info, log = log)
     }
 
     final class Theory private[Headless](
@@ -496,7 +496,7 @@
       lazy val theory_graph: Document.Node.Name.Graph[Unit] =
         Document.Node.Name.make_graph(
           for ((name, theory) <- theories.toList)
-          yield ((name, ()), theory.node_header.imports.filter(theories.isDefinedAt(_))))
+          yield ((name, ()), theory.node_header.imports.filter(theories.isDefinedAt)))
 
       def is_required(name: Document.Node.Name): Boolean = required.isDefinedAt(name)
 
@@ -542,7 +542,7 @@
         : ((List[Document.Node.Name], List[Document.Node.Name], List[Document.Edit_Text]), State) =
       {
         val all_nodes = theory_graph.topological_order
-        val purge = nodes.getOrElse(all_nodes).filterNot(is_required(_)).toSet
+        val purge = nodes.getOrElse(all_nodes).filterNot(is_required).toSet
 
         val retain = theory_graph.all_preds(all_nodes.filterNot(purge)).toSet
         val (retained, purged) = all_nodes.partition(retain)
@@ -554,6 +554,7 @@
   }
 
   class Resources private[Headless](
+      val options: Options,
       val session_base_info: Sessions.Base_Info,
       log: Logger = No_Logger)
     extends isabelle.Resources(
@@ -561,37 +562,20 @@
   {
     resources =>
 
-    def options: Options = session_base_info.options
+    val store: Sessions.Store = Sessions.store(options)
 
 
     /* session */
 
-    def start_session(print_mode: List[String] = Nil, progress: Progress = No_Progress): Session =
+    def start_session(print_mode: List[String] = Nil, progress: Progress = new Progress): Session =
     {
       val session = new Session(session_base_info.session, options, resources)
 
-      val session_error = Future.promise[String]
-      var session_phase: Session.Consumer[Session.Phase] = null
-      session_phase =
-        Session.Consumer(getClass.getName) {
-          case Session.Ready =>
-            session.phase_changed -= session_phase
-            session_error.fulfill("")
-          case Session.Terminated(result) if !result.ok =>
-            session.phase_changed -= session_phase
-            session_error.fulfill("Session start failed: return code " + result.rc)
-          case _ =>
-        }
-      session.phase_changed += session_phase
+      progress.echo("Starting session " + session_base_info.session + " ...")
+      Isabelle_Process(session, options, session_base_info.sessions_structure, store,
+        logic = session_base_info.session, modes = print_mode).await_startup
 
-      progress.echo("Starting session " + session_base_info.session + " ...")
-      Isabelle_Process.start(session, options,
-        logic = session_base_info.session, dirs = session_base_info.dirs, modes = print_mode)
-
-      session_error.join match {
-        case "" => session
-        case msg => session.stop(); error(msg)
-      }
+      session
     }
 
 
--- a/src/Pure/PIDE/markup.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/markup.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -58,8 +58,8 @@
 
   /* basic markup */
 
-  val Empty = Markup("", Nil)
-  val Broken = Markup("broken", Nil)
+  val Empty: Markup = Markup("", Nil)
+  val Broken: Markup = Markup("broken", Nil)
 
   class Markup_String(val name: String, prop: String)
   {
@@ -549,133 +549,72 @@
   /* protocol message functions */
 
   val FUNCTION = "function"
-  val Function = new Properties.String(FUNCTION)
 
-  val COMMAND_TIMING: Properties.Entry = (FUNCTION, "command_timing")
-  val THEORY_TIMING: Properties.Entry = (FUNCTION, "theory_timing")
+  class Function(val name: String)
+  {
+    val PROPERTY: Properties.Entry = (FUNCTION, name)
+  }
 
-  val Commands_Accepted: Properties.T = List((FUNCTION, "commands_accepted"))
-  val Assign_Update: Properties.T = List((FUNCTION, "assign_update"))
-  val Removed_Versions: Properties.T = List((FUNCTION, "removed_versions"))
-
-  object Protocol_Handler
+  class Properties_Function(name: String) extends Function(name)
   {
-    def unapply(props: Properties.T): Option[(String)] =
+    def unapply(props: Properties.T): Option[Properties.T] =
       props match {
-        case List((FUNCTION, "protocol_handler"), (NAME, name)) => Some(name)
+        case PROPERTY :: args => Some(args)
         case _ => None
       }
   }
 
-  val INVOKE_SCALA = "invoke_scala"
-  object Invoke_Scala
+  class Name_Function(name: String) extends Function(name)
   {
-    def unapply(props: Properties.T): Option[(String, String)] =
+    def unapply(props: Properties.T): Option[(String)] =
       props match {
-        case List((FUNCTION, INVOKE_SCALA), (NAME, name), (ID, id)) => Some((name, id))
+        case List(PROPERTY, (NAME, a)) => Some(a)
         case _ => None
       }
   }
 
-  val CANCEL_SCALA = "cancel_scala"
-  object Cancel_Scala
+  object Command_Timing extends Properties_Function("command_timing")
+  object Theory_Timing extends Properties_Function("theory_timing")
+  object ML_Statistics extends Properties_Function("ML_statistics")
+  object Task_Statistics extends Properties_Function("task_statistics")
+
+  object Loading_Theory extends Name_Function("loading_theory")
+  object Build_Session_Finished extends Function("build_session_finished")
+
+  object Protocol_Handler extends Name_Function("protocol_handler")
+
+  object Commands_Accepted extends Function("commands_accepted")
+  object Assign_Update extends Function("assign_update")
+  object Removed_Versions extends Function("removed_versions")
+
+  object Invoke_Scala extends Function("invoke_scala")
   {
-    def unapply(props: Properties.T): Option[String] =
+    def unapply(props: Properties.T): Option[(String, String)] =
       props match {
-        case List((FUNCTION, CANCEL_SCALA), (ID, id)) => Some(id)
-        case _ => None
-      }
-  }
-
-  object ML_Statistics
-  {
-    def unapply(props: Properties.T): Option[Properties.T] =
-      props match {
-        case (FUNCTION, "ML_statistics") :: props => Some(props)
+        case List(PROPERTY, (NAME, name), (ID, id)) => Some((name, id))
         case _ => None
       }
   }
 
-  object Task_Statistics
+  object Cancel_Scala extends Function("cancel_scala")
   {
-    def unapply(props: Properties.T): Option[Properties.T] =
+    def unapply(props: Properties.T): Option[String] =
       props match {
-        case (FUNCTION, "task_statistics") :: props => Some(props)
+        case List(PROPERTY, (ID, id)) => Some(id)
         case _ => None
       }
   }
 
-  val LOADING_THEORY = "loading_theory"
-  object Loading_Theory
-  {
-    def unapply(props: Properties.T): Option[String] =
-      props match {
-        case List((FUNCTION, LOADING_THEORY), (NAME, name)) => Some(name)
-        case _ => None
-      }
-  }
-
-  val BUILD_SESSION_FINISHED = "build_session_finished"
-  val Build_Session_Finished: Properties.T = List((FUNCTION, BUILD_SESSION_FINISHED))
-
   val PRINT_OPERATIONS = "print_operations"
 
 
-
   /* export */
 
   val EXPORT = "export"
-
-  object Export
-  {
-    sealed case class Args(
-      id: Option[String],
-      serial: Long,
-      theory_name: String,
-      name: String,
-      executable: Boolean,
-      compress: Boolean,
-      strict: Boolean)
-    {
-      def compound_name: String = isabelle.Export.compound_name(theory_name, name)
-    }
-
-    val THEORY_NAME = "theory_name"
-    val EXECUTABLE = "executable"
-    val COMPRESS = "compress"
-    val STRICT = "strict"
-
-    def dest_inline(props: Properties.T): Option[(Args, Path)] =
-      props match {
-        case
-          List(
-            (SERIAL, Value.Long(serial)),
-            (THEORY_NAME, theory_name),
-            (NAME, name),
-            (EXECUTABLE, Value.Boolean(executable)),
-            (COMPRESS, Value.Boolean(compress)),
-            (STRICT, Value.Boolean(strict)),
-            (FILE, file)) if isabelle.Path.is_valid(file) =>
-          val args = Args(None, serial, theory_name, name, executable, compress, strict)
-          Some((args, isabelle.Path.explode(file)))
-        case _ => None
-      }
-
-    def unapply(props: Properties.T): Option[Args] =
-      props match {
-        case List(
-            (FUNCTION, EXPORT),
-            (ID, id),
-            (SERIAL, Value.Long(serial)),
-            (THEORY_NAME, theory_name),
-            (NAME, name),
-            (EXECUTABLE, Value.Boolean(executable)),
-            (COMPRESS, Value.Boolean(compress)),
-            (STRICT, Value.Boolean(strict))) =>
-          Some(Args(proper_string(id), serial, theory_name, name, executable, compress, strict))
-        case _ => None
-      }
-  }
+  val THEORY_NAME = "theory_name"
+  val EXECUTABLE = "executable"
+  val COMPRESS = "compress"
+  val STRICT = "strict"
 
 
   /* debugger output */
--- a/src/Pure/PIDE/protocol.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/protocol.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -9,6 +9,19 @@
 
 object Protocol
 {
+  /* markers for inlined messages */
+
+  val Loading_Theory_Marker = Protocol_Message.Marker("loading_theory")
+  val Export_Marker = Protocol_Message.Marker("export")
+  val Meta_Info_Marker = Protocol_Message.Marker("meta_info")
+  val Timing_Marker = Protocol_Message.Marker("Timing")
+  val Command_Timing_Marker = Protocol_Message.Marker("command_timing")
+  val Theory_Timing_Marker = Protocol_Message.Marker("theory_timing")
+  val ML_Statistics_Marker = Protocol_Message.Marker("ML_statistics")
+  val Task_Statistics_Marker = Protocol_Message.Marker("task_statistics")
+  val Error_Message_Marker = Protocol_Message.Marker("error_message")
+
+
   /* document editing */
 
   object Commands_Accepted
@@ -32,7 +45,7 @@
             case a :: bs => (a, bs)
             case _ => throw new XML.XML_Body(body)
           }
-        Some(triple(long, list(string), list(decode_upd _))(Symbol.decode_yxml(text)))
+        Some(triple(long, list(string), list(decode_upd))(Symbol.decode_yxml(text)))
       }
       catch {
         case ERROR(_) => None
@@ -59,11 +72,11 @@
 
   object Command_Timing
   {
-    def unapply(props: Properties.T): Option[(Document_ID.Generic, isabelle.Timing)] =
+    def unapply(props: Properties.T): Option[(Properties.T, Document_ID.Generic, isabelle.Timing)] =
       props match {
-        case Markup.COMMAND_TIMING :: args =>
+        case Markup.Command_Timing(args) =>
           (args, args) match {
-            case (Position.Id(id), Markup.Timing_Properties(timing)) => Some((id, timing))
+            case (Position.Id(id), Markup.Timing_Properties(timing)) => Some((args, id, timing))
             case _ => None
           }
         case _ => None
@@ -77,7 +90,7 @@
   {
     def unapply(props: Properties.T): Option[(String, isabelle.Timing)] =
       props match {
-        case Markup.THEORY_TIMING :: args =>
+        case Markup.Theory_Timing(args) =>
           (args, args) match {
             case (Markup.Name(name), Markup.Timing_Properties(timing)) => Some((name, timing))
             case _ => None
@@ -89,6 +102,12 @@
 
   /* result messages */
 
+  def is_message(pred: XML.Elem => Boolean, body: XML.Body): Boolean =
+    body match {
+      case List(elem: XML.Elem) => pred(elem)
+      case _ => false
+    }
+
   def is_result(msg: XML.Tree): Boolean =
     msg match {
       case XML.Elem(Markup(Markup.RESULT, _), _) => true
@@ -150,12 +169,76 @@
   def is_exported(msg: XML.Tree): Boolean =
     is_writeln(msg) || is_warning(msg) || is_legacy(msg) || is_error(msg)
 
-  def message_text(msg: XML.Tree): String =
+  def message_text(body: XML.Body,
+    margin: Double = Pretty.default_margin,
+    breakgain: Double = Pretty.default_breakgain,
+    metric: Pretty.Metric = Pretty.Default_Metric): String =
+  {
+    val text =
+      Pretty.string_of(Protocol_Message.expose_no_reports(body),
+        margin = margin, breakgain = breakgain, metric = metric)
+
+    if (is_message(is_warning, body) || is_message(is_legacy, body)) Output.warning_prefix(text)
+    else if (is_message(is_error, body)) Output.error_prefix(text)
+    else text
+  }
+
+
+  /* export */
+
+  object Export
   {
-    val text = Pretty.string_of(List(msg))
-    if (is_warning(msg) || is_legacy(msg)) Library.prefix_lines("### ", text)
-    else if (is_error(msg)) Library.prefix_lines("*** ", text)
-    else text
+    sealed case class Args(
+      id: Option[String],
+      serial: Long,
+      theory_name: String,
+      name: String,
+      executable: Boolean,
+      compress: Boolean,
+      strict: Boolean)
+    {
+      def compound_name: String = isabelle.Export.compound_name(theory_name, name)
+    }
+
+    object Marker
+    {
+      def unapply(line: String): Option[(Args, Path)] =
+        line match {
+          case Export_Marker(text) =>
+            val props = XML.Decode.properties(YXML.parse_body(text))
+            props match {
+              case
+                List(
+                  (Markup.SERIAL, Value.Long(serial)),
+                  (Markup.THEORY_NAME, theory_name),
+                  (Markup.NAME, name),
+                  (Markup.EXECUTABLE, Value.Boolean(executable)),
+                  (Markup.COMPRESS, Value.Boolean(compress)),
+                  (Markup.STRICT, Value.Boolean(strict)),
+                  (Markup.FILE, file)) if isabelle.Path.is_valid(file) =>
+                val args = Args(None, serial, theory_name, name, executable, compress, strict)
+                Some((args, isabelle.Path.explode(file)))
+              case _ => None
+            }
+          case _ => None
+        }
+    }
+
+    def unapply(props: Properties.T): Option[Args] =
+      props match {
+        case
+          List(
+            (Markup.FUNCTION, Markup.EXPORT),
+            (Markup.ID, id),
+            (Markup.SERIAL, Value.Long(serial)),
+            (Markup.THEORY_NAME, theory_name),
+            (Markup.NAME, name),
+            (Markup.EXECUTABLE, Value.Boolean(executable)),
+            (Markup.COMPRESS, Value.Boolean(compress)),
+            (Markup.STRICT, Value.Boolean(strict))) =>
+          Some(Args(proper_string(id), serial, theory_name, name, executable, compress, strict))
+        case _ => None
+      }
   }
 
 
@@ -310,7 +393,7 @@
   def define_commands_bulk(commands: List[Command])
   {
     val (irregular, regular) = commands.partition(command => YXML.detect(command.source))
-    irregular.foreach(define_command(_))
+    irregular.foreach(define_command)
     regular match {
       case Nil =>
       case List(command) => define_command(command)
--- a/src/Pure/PIDE/protocol_handlers.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/protocol_handlers.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -51,14 +51,17 @@
     def init(name: String): State =
     {
       val new_handler =
-        try { Some(Class.forName(name).newInstance.asInstanceOf[Session.Protocol_Handler]) }
+        try {
+          Some(Class.forName(name).getDeclaredConstructor().newInstance()
+            .asInstanceOf[Session.Protocol_Handler])
+        }
         catch { case exn: Throwable => bad_handler(exn, name); None }
       new_handler match { case Some(handler) => init(handler) case None => this }
     }
 
     def invoke(msg: Prover.Protocol_Output): Boolean =
       msg.properties match {
-        case Markup.Function(a) if functions.isDefinedAt(a) =>
+        case (Markup.FUNCTION, a) :: _ if functions.isDefinedAt(a) =>
           try { functions(a)(msg) }
           catch {
             case exn: Throwable =>
--- a/src/Pure/PIDE/protocol_message.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/protocol_message.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -6,7 +6,8 @@
 
 signature PROTOCOL_MESSAGE =
 sig
-  val inline: string -> Properties.T -> unit
+  val marker_text: string -> string -> unit
+  val marker: string -> Properties.T -> unit
   val command_positions: string -> XML.body -> XML.body
   val command_positions_yxml: string -> string -> string
 end;
@@ -14,8 +15,11 @@
 structure Protocol_Message: PROTOCOL_MESSAGE =
 struct
 
-fun inline a args =
-  writeln ("\f" ^ a ^ " = " ^ YXML.string_of_body (XML.Encode.properties args));
+fun marker_text a text =
+  Output.physical_writeln ("\f" ^ a ^ " = " ^ encode_lines text);
+
+fun marker a props =
+  marker_text a (YXML.string_of_body (XML.Encode.properties props));
 
 
 fun command_positions id =
--- a/src/Pure/PIDE/protocol_message.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/protocol_message.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -9,6 +9,34 @@
 
 object Protocol_Message
 {
+  /* message markers */
+
+  object Marker
+  {
+    def apply(a: String): Marker =
+      new Marker { override def name: String = a }
+
+    def test(line: String): Boolean = line.startsWith("\f")
+  }
+
+  abstract class Marker private
+  {
+    def name: String
+    val prefix: String = "\f" + name + " = "
+
+    def apply(text: String): String = prefix + Library.encode_lines(text)
+    def apply(props: Properties.T): String = apply(YXML.string_of_body(XML.Encode.properties(props)))
+
+    def test(line: String): Boolean = line.startsWith(prefix)
+    def test_yxml(line: String): Boolean = test(line) && YXML.detect(line)
+
+    def unapply(line: String): Option[String] =
+      Library.try_unprefix(prefix, line).map(Library.decode_lines)
+
+    override def toString: String = "Marker(" + quote(name) + ")"
+  }
+
+
   /* inlined reports */
 
   private val report_elements =
@@ -25,6 +53,14 @@
       case t => t
     }
 
+  def expose_no_reports(body: XML.Body): XML.Body =
+    body flatMap {
+      case XML.Wrapped_Elem(markup, body, ts) => List(XML.Wrapped_Elem(markup, body, expose_no_reports(ts)))
+      case XML.Elem(Markup(Markup.NO_REPORT, _), ts) => ts
+      case XML.Elem(markup, ts) => List(XML.Elem(markup, expose_no_reports(ts)))
+      case t => List(t)
+    }
+
   def reports(props: Properties.T, body: XML.Body): List[XML.Elem] =
     body flatMap {
       case XML.Wrapped_Elem(Markup(Markup.REPORT, ps), body, ts) =>
--- a/src/Pure/PIDE/prover.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/prover.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -32,20 +32,20 @@
     def properties: Properties.T = message.markup.properties
     def body: XML.Body = message.body
 
-    def is_init = kind == Markup.INIT
-    def is_exit = kind == Markup.EXIT
-    def is_stdout = kind == Markup.STDOUT
-    def is_stderr = kind == Markup.STDERR
-    def is_system = kind == Markup.SYSTEM
-    def is_status = kind == Markup.STATUS
-    def is_report = kind == Markup.REPORT
-    def is_syslog = is_init || is_exit || is_system || is_stderr
+    def is_init: Boolean = kind == Markup.INIT
+    def is_exit: Boolean = kind == Markup.EXIT
+    def is_stdout: Boolean = kind == Markup.STDOUT
+    def is_stderr: Boolean = kind == Markup.STDERR
+    def is_system: Boolean = kind == Markup.SYSTEM
+    def is_status: Boolean = kind == Markup.STATUS
+    def is_report: Boolean = kind == Markup.REPORT
+    def is_syslog: Boolean = is_init || is_exit || is_system || is_stderr
 
     override def toString: String =
     {
       val res =
         if (is_status || is_report) message.body.map(_.toString).mkString
-        else Pretty.string_of(message.body)
+        else Pretty.string_of(message.body, metric = Symbol.Metric)
       if (properties.isEmpty)
         kind.toString + " [[" + res + "]]"
       else
@@ -90,7 +90,7 @@
   private def exit_message(result: Process_Result)
   {
     output(Markup.EXIT, Markup.Process_Result(result),
-      List(XML.Text("Return code: " + result.rc.toString)))
+      List(XML.Text(result.print_return_code)))
   }
 
 
@@ -112,8 +112,10 @@
     }
   }
 
-  private val process_manager = Standard_Thread.fork("process_manager")
+  private val process_manager = Isabelle_Thread.fork(name = "process_manager")
   {
+    val stdout = physical_output(false)
+
     val (startup_failed, startup_errors) =
     {
       var finished: Option[Boolean] = None
@@ -127,7 +129,7 @@
           }
           catch { case _: IOException => finished = Some(false) }
         }
-        Thread.sleep(10)
+        Time.seconds(0.05).sleep
       }
       (finished.isEmpty || !finished.get, result.toString.trim)
     }
@@ -136,13 +138,13 @@
     if (startup_failed) {
       terminate_process()
       process_result.join
+      stdout.join
       exit_message(Process_Result(127))
     }
     else {
       val (command_stream, message_stream) = channel.rendezvous()
 
       command_input_init(command_stream)
-      val stdout = physical_output(false)
       val stderr = physical_output(true)
       val message = message_output(message_stream)
 
@@ -168,7 +170,7 @@
 
     var count = 10
     while (!process_result.is_finished && count > 0) {
-      Thread.sleep(100)
+      Time.seconds(0.1).sleep
       count -= 1
     }
     if (!process_result.is_finished) terminate_process()
@@ -216,7 +218,7 @@
       if (err) ("standard_error", process.stderr, Markup.STDERR)
       else ("standard_output", process.stdout, Markup.STDOUT)
 
-    Standard_Thread.fork(name) {
+    Isabelle_Thread.fork(name = name) {
       try {
         var result = new StringBuilder(100)
         var finished = false
@@ -254,7 +256,7 @@
     class Protocol_Error(msg: String) extends Exception(msg)
 
     val name = "message_output"
-    Standard_Thread.fork(name) {
+    Isabelle_Thread.fork(name = name) {
       val default_buffer = new Array[Byte](65536)
       var c = -1
 
--- a/src/Pure/PIDE/rendering.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/rendering.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -21,20 +21,20 @@
     // background
     val unprocessed1, running1, canceled, bad, intensify, entity, active, active_result,
       markdown_bullet1, markdown_bullet2, markdown_bullet3, markdown_bullet4 = Value
-    val background_colors = values
+    val background_colors: ValueSet = values
 
     // foreground
     val quoted, antiquoted = Value
-    val foreground_colors = values -- background_colors
+    val foreground_colors: ValueSet = values -- background_colors
 
     // message underline
     val writeln, information, warning, legacy, error = Value
-    val message_underline_colors = values -- background_colors -- foreground_colors
+    val message_underline_colors: ValueSet = values -- background_colors -- foreground_colors
 
     // message background
     val writeln_message, information_message, tracing_message,
       warning_message, legacy_message, error_message = Value
-    val message_background_colors =
+    val message_background_colors: ValueSet =
       values -- background_colors -- foreground_colors -- message_underline_colors
 
     // text
@@ -42,7 +42,7 @@
       tfree, tvar, free, skolem, bound, `var`, inner_numeral, inner_quoted,
       inner_cartouche, comment1, comment2, comment3, dynamic, class_parameter,
       antiquote, raw_text, plain_text = Value
-    val text_colors =
+    val text_colors: ValueSet =
       values -- background_colors -- foreground_colors -- message_underline_colors --
       message_background_colors
 
--- a/src/Pure/PIDE/resources.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/resources.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -23,19 +23,17 @@
 
   /* file formats */
 
-  val file_formats: File_Format.Environment = File_Format.environment()
-
   def make_theory_name(name: Document.Node.Name): Option[Document.Node.Name] =
-    file_formats.get(name).flatMap(_.make_theory_name(resources, name))
+    File_Format.registry.get(name).flatMap(_.make_theory_name(resources, name))
 
   def make_theory_content(thy_name: Document.Node.Name): Option[String] =
-    file_formats.get_theory(thy_name).flatMap(_.make_theory_content(resources, thy_name))
+    File_Format.registry.get_theory(thy_name).flatMap(_.make_theory_content(resources, thy_name))
 
   def make_preview(snapshot: Document.Snapshot): Option[Present.Preview] =
-    file_formats.get(snapshot.node_name).flatMap(_.make_preview(snapshot))
+    File_Format.registry.get(snapshot.node_name).flatMap(_.make_preview(snapshot))
 
   def is_hidden(name: Document.Node.Name): Boolean =
-    !name.is_theory || name.theory == Sessions.root_name || file_formats.is_theory(name)
+    !name.is_theory || name.theory == Sessions.root_name || File_Format.registry.is_theory(name)
 
   def thy_path(path: Path): Path = path.ext("thy")
 
@@ -265,10 +263,10 @@
 
   def dependencies(
       thys: List[(Document.Node.Name, Position.T)],
-      progress: Progress = No_Progress): Dependencies[Unit] =
+      progress: Progress = new Progress): Dependencies[Unit] =
     Dependencies.empty[Unit].require_thys((), thys, progress = progress)
 
-  def session_dependencies(info: Sessions.Info, progress: Progress = No_Progress)
+  def session_dependencies(info: Sessions.Info, progress: Progress = new Progress)
     : Dependencies[Options] =
   {
     (Dependencies.empty[Options] /: info.theories)({ case (dependencies, (options, thys)) =>
@@ -303,7 +301,7 @@
     def require_thy(adjunct: A,
       thy: (Document.Node.Name, Position.T),
       initiators: List[Document.Node.Name] = Nil,
-      progress: Progress = No_Progress): Dependencies[A] =
+      progress: Progress = new Progress): Dependencies[A] =
     {
       val (name, pos) = thy
 
@@ -337,7 +335,7 @@
 
     def require_thys(adjunct: A,
         thys: List[(Document.Node.Name, Position.T)],
-        progress: Progress = No_Progress,
+        progress: Progress = new Progress,
         initiators: List[Document.Node.Name] = Nil): Dependencies[A] =
       (this /: thys)(_.require_thy(adjunct, _, progress = progress, initiators = initiators))
 
--- a/src/Pure/PIDE/session.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/session.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -9,7 +9,7 @@
   val get_name: unit -> string
   val welcome: unit -> string
   val get_keywords: unit -> Keyword.keywords
-  val init: HTML.symbols -> bool -> bool -> Path.T -> string -> string -> (string * string) list ->
+  val init: HTML.symbols -> bool -> Path.T -> string -> string -> (string * string) list ->
     (Path.T * Path.T) list -> Path.T -> string -> string * string -> bool -> unit
   val shutdown: unit -> unit
   val finish: unit -> unit
@@ -48,13 +48,13 @@
 
 (* init *)
 
-fun init symbols build info info_path doc doc_output doc_variants doc_files graph_file
+fun init symbols info info_path doc doc_output doc_variants doc_files graph_file
     parent (chapter, name) verbose =
   (Synchronized.change session (fn ({name = parent_name, ...}, parent_finished) =>
     if parent_name <> parent orelse not parent_finished then
       error ("Unfinished parent session " ^ quote parent ^ " for " ^ quote name)
     else ({chapter = chapter, name = name}, false));
-    Present.init symbols build info info_path (if doc = "false" then "" else doc)
+    Present.init symbols info info_path (if doc = "false" then "" else doc)
       doc_output doc_variants doc_files graph_file (chapter, name) verbose);
 
 
--- a/src/Pure/PIDE/session.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/session.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -61,7 +61,10 @@
   /* events */
 
   //{{{
-  case class Statistics(props: Properties.T)
+  case class Command_Timing(props: Properties.T)
+  case class Theory_Timing(props: Properties.T)
+  case class Runtime_Statistics(props: Properties.T)
+  case class Task_Statistics(props: Properties.T)
   case class Global_Options(options: Options)
   case object Caret_Focus
   case class Raw_Edits(doc_blobs: Document.Blobs, edits: List[Document.Edit_Text])
@@ -175,7 +178,10 @@
 
   /* outlets */
 
-  val statistics = new Session.Outlet[Session.Statistics](dispatcher)
+  val command_timings = new Session.Outlet[Session.Command_Timing](dispatcher)
+  val theory_timings = new Session.Outlet[Session.Theory_Timing](dispatcher)
+  val runtime_statistics = new Session.Outlet[Session.Runtime_Statistics](dispatcher)
+  val task_statistics = new Session.Outlet[Session.Task_Statistics](dispatcher)
   val global_options = new Session.Outlet[Session.Global_Options](dispatcher)
   val caret_focus = new Session.Outlet[Session.Caret_Focus.type](dispatcher)
   val raw_edits = new Session.Outlet[Session.Raw_Edits](dispatcher)
@@ -213,7 +219,7 @@
   private val _phase = Synchronized[Session.Phase](Session.Inactive)
   private def phase_=(new_phase: Session.Phase): Unit = _phase.change(_ => post_phase(new_phase))
 
-  def phase = _phase.value
+  def phase: Session.Phase = _phase.value
   def is_ready: Boolean = phase == Session.Ready
 
 
@@ -262,7 +268,7 @@
       nodes = Set.empty
       commands = Set.empty
     }
-    private val delay_flush = Standard_Thread.delay_first(output_delay) { flush() }
+    private val delay_flush = Delay.first(output_delay) { flush() }
 
     def invoke(assign: Boolean, edited_nodes: List[Document.Node.Name], cmds: List[Command]): Unit =
       synchronized {
@@ -307,7 +313,7 @@
   private object consolidation
   {
     private val delay =
-      Standard_Thread.delay_first(consolidate_delay) { manager.send(Consolidate_Execution) }
+      Delay.first(consolidate_delay) { manager.send(Consolidate_Execution) }
 
     private val init_state: Option[Set[Document.Node.Name]] = Some(Set.empty)
     private val state = Synchronized(init_state)
@@ -348,7 +354,7 @@
   /* file formats */
 
   lazy val file_formats: File_Format.Session =
-    resources.file_formats.start_session(session)
+    File_Format.registry.start_session(session)
 
 
   /* protocol handlers */
@@ -376,7 +382,7 @@
   /* manager thread */
 
   private val delay_prune =
-    Standard_Thread.delay_first(prune_delay) { manager.send(Prune_History) }
+    Delay.first(prune_delay) { manager.send(Prune_History) }
 
   private val manager: Consumer_Thread[Any] =
   {
@@ -479,20 +485,27 @@
               case Markup.Protocol_Handler(name) if prover.defined =>
                 init_protocol_handler(name)
 
-              case Protocol.Command_Timing(state_id, timing) if prover.defined =>
+              case Protocol.Command_Timing(props, state_id, timing) if prover.defined =>
+                command_timings.post(Session.Command_Timing(props))
                 val message = XML.elem(Markup.STATUS, List(XML.Elem(Markup.Timing(timing), Nil)))
                 change_command(_.accumulate(state_id, xml_cache.elem(message), xml_cache))
 
-              case Protocol.Theory_Timing(_, _) =>
-                // FIXME
+              case Markup.Theory_Timing(props) =>
+                theory_timings.post(Session.Theory_Timing(props))
 
-              case Markup.Export(args)
+              case Markup.ML_Statistics(props) =>
+                runtime_statistics.post(Session.Runtime_Statistics(props))
+
+              case Markup.Task_Statistics(props) =>
+                task_statistics.post(Session.Task_Statistics(props))
+
+              case Protocol.Export(args)
               if args.id.isDefined && Value.Long.unapply(args.id.get).isDefined =>
                 val id = Value.Long.unapply(args.id.get).get
                 val export = Export.make_entry("", args, msg.bytes, cache = xz_cache)
                 change_command(_.add_export(id, (args.serial, export)))
 
-              case Markup.Commands_Accepted =>
+              case List(Markup.Commands_Accepted.PROPERTY) =>
                 msg.text match {
                   case Protocol.Commands_Accepted(ids) =>
                     ids.foreach(id =>
@@ -500,7 +513,7 @@
                   case _ => bad_output()
                 }
 
-              case Markup.Assign_Update =>
+              case List(Markup.Assign_Update.PROPERTY) =>
                 msg.text match {
                   case Protocol.Assign_Update(id, edited, update) =>
                     try {
@@ -514,7 +527,7 @@
                 }
                 delay_prune.invoke()
 
-              case Markup.Removed_Versions =>
+              case List(Markup.Removed_Versions.PROPERTY) =>
                 msg.text match {
                   case Protocol.Removed(removed) =>
                     try {
@@ -525,12 +538,6 @@
                   case _ => bad_output()
                 }
 
-              case Markup.ML_Statistics(props) =>
-                statistics.post(Session.Statistics(props))
-
-              case Markup.Task_Statistics(props) =>
-                // FIXME
-
               case _ => bad_output()
             }
           }
@@ -566,15 +573,15 @@
         //{{{
         arg match {
           case output: Prover.Output =>
-            if (output.is_stdout || output.is_stderr)
-              raw_output_messages.post(output)
-            else handle_output(output)
-
             if (output.is_syslog) {
               syslog += output.message
               syslog_messages.post(output)
             }
 
+            if (output.is_stdout || output.is_stderr)
+              raw_output_messages.post(output)
+            else handle_output(output)
+
             all_messages.post(output)
 
           case input: Prover.Input =>
@@ -645,7 +652,7 @@
           case Session.Change_Flush if prover.defined =>
             val state = global_state.value
             if (!state.removing_versions)
-              postponed_changes.flush(state).foreach(handle_change(_))
+              postponed_changes.flush(state).foreach(handle_change)
 
           case bad =>
             if (verbose) Output.warning("Ignoring bad message: " + bad.toString)
@@ -680,7 +687,7 @@
   {
     val snapshot = this.snapshot()
     if (snapshot.is_outdated) {
-      Thread.sleep(output_delay.ms)
+      output_delay.sleep
       await_stable_snapshot()
     }
     else snapshot
@@ -698,22 +705,17 @@
       })
   }
 
-  def send_stop()
+  def stop(): Process_Result =
   {
     val was_ready =
-      _phase.guarded_access(phase =>
-        phase match {
+      _phase.guarded_access(
+        {
           case Session.Startup | Session.Shutdown => None
           case Session.Terminated(_) => Some((false, phase))
           case Session.Inactive => Some((false, post_phase(Session.Terminated(Process_Result(0)))))
           case Session.Ready => Some((true, post_phase(Session.Shutdown)))
         })
     if (was_ready) manager.send(Stop)
-  }
-
-  def stop(): Process_Result =
-  {
-    send_stop()
     prover.await_reset()
 
     change_parser.shutdown()
--- a/src/Pure/PIDE/xml.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/xml.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -217,7 +217,7 @@
       else
         lookup(x) match {
           case Some(y) => y
-          case None => x.map(cache_tree(_))
+          case None => x.map(cache_tree)
         }
     }
 
--- a/src/Pure/PIDE/yxml.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/PIDE/yxml.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -21,10 +21,10 @@
   val is_X = (c: Char) => c == X
   val is_Y = (c: Char) => c == Y
 
-  val X_string = X.toString
-  val Y_string = Y.toString
-  val XY_string = X_string + Y_string
-  val XYX_string = XY_string + X_string
+  val X_string: String = X.toString
+  val Y_string: String = Y.toString
+  val XY_string: String = X_string + Y_string
+  val XYX_string: String = XY_string + X_string
 
   def detect(s: String): Boolean = s.exists(c => c == X || c == Y)
   def detect_elem(s: String): Boolean = s.startsWith(XY_string)
--- a/src/Pure/ROOT.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/ROOT.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -96,15 +96,16 @@
 ML_file "General/change_table.ML";
 ML_file "General/graph.ML";
 
+ML_file "System/options.ML";
+
 
 subsection "Fundamental structures";
 
 ML_file "name.ML";
 ML_file "term.ML";
 ML_file "context.ML";
+ML_file "config.ML";
 ML_file "context_position.ML";
-ML_file "System/options.ML";
-ML_file "config.ML";
 ML_file "soft_type_system.ML";
 
 
@@ -116,7 +117,7 @@
 ML_file "ML/ml_statistics.ML";
 
 ML_file "Concurrent/thread_data_virtual.ML";
-ML_file "Concurrent/standard_thread.ML";
+ML_file "Concurrent/isabelle_thread.ML";
 ML_file "Concurrent/single_assignment.ML";
 
 ML_file "Concurrent/par_exn.ML";
--- a/src/Pure/Syntax/syntax_phases.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Syntax/syntax_phases.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -950,7 +950,7 @@
 fun check_typs ctxt raw_tys =
   let
     val (sorting_report, tys) = Proof_Context.prepare_sortsT ctxt raw_tys;
-    val _ = if Context_Position.is_visible ctxt then Output.report sorting_report else ();
+    val _ = if Context_Position.reports_enabled ctxt then Output.report sorting_report else ();
   in
     tys
     |> apply_typ_check ctxt
@@ -974,8 +974,8 @@
         else I) ps tys' [];
 
     val _ =
-      if Context_Position.is_visible ctxt then Output.report (sorting_report @ typing_report)
-      else ();
+      if Context_Position.reports_enabled ctxt
+      then Output.report (sorting_report @ typing_report) else ();
   in Term_Sharing.terms (Proof_Context.theory_of ctxt) ts' end;
 
 fun check_props ctxt = map (Type.constraint propT) #> check_terms ctxt;
--- a/src/Pure/System/bash.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/bash.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -38,7 +38,7 @@
     val _ = cleanup_files ();
 
     val system_thread =
-      Standard_Thread.fork {name = "bash", stack_limit = NONE, interrupts = false} (fn () =>
+      Isabelle_Thread.fork {name = "bash", stack_limit = NONE, interrupts = false} (fn () =>
         Thread_Attributes.with_attributes Thread_Attributes.private_interrupts (fn _ =>
           let
             val _ = File.write script_path script;
@@ -86,7 +86,7 @@
           in () end;
 
     fun cleanup () =
-     (Standard_Thread.interrupt_unsynchronized system_thread;
+     (Isabelle_Thread.interrupt_unsynchronized system_thread;
       cleanup_files ());
   in
     let
@@ -132,7 +132,7 @@
     val _ = cleanup_files ();
 
     val system_thread =
-      Standard_Thread.fork {name = "bash", stack_limit = NONE, interrupts = false} (fn () =>
+      Isabelle_Thread.fork {name = "bash", stack_limit = NONE, interrupts = false} (fn () =>
         Thread_Attributes.with_attributes Thread_Attributes.private_interrupts (fn _ =>
           let
             val _ = File.write script_path script;
@@ -182,7 +182,7 @@
           in () end;
 
     fun cleanup () =
-     (Standard_Thread.interrupt_unsynchronized system_thread;
+     (Isabelle_Thread.interrupt_unsynchronized system_thread;
       cleanup_files ());
   in
     let
--- a/src/Pure/System/bash.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/bash.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -10,6 +10,8 @@
 import java.io.{File => JFile, BufferedReader, InputStreamReader,
   BufferedWriter, OutputStreamWriter}
 
+import scala.annotation.tailrec
+
 
 object Bash
 {
@@ -35,10 +37,10 @@
 
   def string(s: String): String =
     if (s == "") "\"\""
-    else UTF8.bytes(s).iterator.map(bash_chr(_)).mkString
+    else UTF8.bytes(s).iterator.map(bash_chr).mkString
 
   def strings(ss: List[String]): String =
-    ss.iterator.map(Bash.string(_)).mkString(" ")
+    ss.iterator.map(Bash.string).mkString(" ")
 
 
   /* process and result */
@@ -100,41 +102,36 @@
 
     private val pid = stdout.readLine
 
-    def interrupt()
-    { Exn.Interrupt.postpone { Isabelle_System.kill("INT", pid) } }
-
-    private def kill(signal: String): Boolean =
-      Exn.Interrupt.postpone {
-        Isabelle_System.kill(signal, pid)
-        Isabelle_System.kill("0", pid)._2 == 0 } getOrElse true
-
-    private def multi_kill(signal: String): Boolean =
+    @tailrec private def kill(signal: String, count: Int = 1): Boolean =
     {
-      var running = true
-      var count = 10
-      while (running && count > 0) {
-        if (kill(signal)) {
-          Exn.Interrupt.postpone {
-            Thread.sleep(100)
-            count -= 1
-          }
+      count <= 0 ||
+      {
+        Isabelle_System.kill(signal, pid)
+        val running = Isabelle_System.kill("0", pid)._2 == 0
+        if (running) {
+          Time.seconds(0.1).sleep
+          kill(signal, count - 1)
         }
-        else running = false
+        else false
       }
-      running
     }
 
-    def terminate()
+    def terminate(): Unit = Isabelle_Thread.try_uninterruptible
     {
-      multi_kill("INT") && multi_kill("TERM") && kill("KILL")
+      kill("INT", count = 7) && kill("TERM", count = 3) && kill("KILL")
       proc.destroy
       do_cleanup()
     }
 
+    def interrupt(): Unit = Isabelle_Thread.try_uninterruptible
+    {
+      Isabelle_System.kill("INT", pid)
+    }
+
 
     // JVM shutdown hook
 
-    private val shutdown_hook = new Thread { override def run = terminate() }
+    private val shutdown_hook = Isabelle_Thread.create(() => terminate())
 
     try { Runtime.getRuntime.addShutdownHook(shutdown_hook) }
     catch { case _: IllegalStateException => }
--- a/src/Pure/System/command_line.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/command_line.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -6,8 +6,7 @@
 
 signature COMMAND_LINE =
 sig
-  val tool: (unit -> int) -> unit
-  val tool0: (unit -> unit) -> unit
+  val tool: (unit -> unit) -> unit
 end;
 
 structure Command_Line: COMMAND_LINE =
@@ -17,11 +16,8 @@
   Thread_Attributes.uninterruptible (fn restore_attributes => fn () =>
     let
       val rc =
-        restore_attributes body () handle exn =>
-          Exn.capture_exit 2 (fn () => (Runtime.exn_error_message exn; raise exn)) ();
+        (restore_attributes body (); 0)
+          handle exn => Exn.capture_exit 2 (fn () => (Runtime.exn_error_message exn; raise exn)) ();
     in exit rc end) ();
 
-fun tool0 body = tool (fn () => (body (); 0));
-
 end;
-
--- a/src/Pure/System/command_line.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/command_line.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -23,18 +23,22 @@
 
   var debug = false
 
-  def tool(body: => Int): Nothing =
+  def tool(body: => Unit)
   {
-    val rc =
-      try { body }
-      catch {
-        case exn: Throwable =>
-          Output.error_message(Exn.message(exn) + (if (debug) "\n" + Exn.trace(exn) else ""))
-          Exn.return_code(exn, 2)
+    val thread =
+      Isabelle_Thread.fork(name = "command_line", inherit_locals = true) {
+        val rc =
+          try { body; 0 }
+          catch {
+            case exn: Throwable =>
+              Output.error_message(Exn.message(exn) + (if (debug) "\n" + Exn.trace(exn) else ""))
+              Exn.return_code(exn, 2)
+          }
+        sys.exit(rc)
       }
-    sys.exit(rc)
+    thread.join
   }
 
-  def tool0(body: => Unit): Nothing = tool { body; 0 }
+  def ML_tool(body: List[String]): String =
+    "Command_Line.tool (fn () => (" + body.mkString("; ") + "));"
 }
-
--- a/src/Pure/System/getopts.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/getopts.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -7,6 +7,9 @@
 package isabelle
 
 
+import scala.annotation.tailrec
+
+
 object Getopts
 {
   def apply(usage_text: String, option_specs: (String, String => Unit)*): Getopts =
@@ -44,7 +47,7 @@
         cat_error(msg, "The error(s) above occurred in command-line option " + print_option(opt))
     }
 
-  private def getopts(args: List[List[Char]]): List[List[Char]] =
+  @tailrec private def getopts(args: List[List[Char]]): List[List[Char]] =
     args match {
       case List('-', '-') :: rest_args => rest_args
       case ('-' :: opt :: rest_opts) :: rest_args if is_option0(opt) && rest_opts.nonEmpty =>
--- a/src/Pure/System/invoke_scala.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/invoke_scala.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -123,6 +123,6 @@
 
   val functions =
     List(
-      Markup.INVOKE_SCALA -> invoke_scala _,
-      Markup.CANCEL_SCALA -> cancel_scala _)
+      Markup.Invoke_Scala.name -> invoke_scala,
+      Markup.Cancel_Scala.name -> cancel_scala)
 }
--- a/src/Pure/System/isabelle_process.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/isabelle_process.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -7,12 +7,15 @@
 signature ISABELLE_PROCESS =
 sig
   val is_active: unit -> bool
+  exception STOP
+  exception EXIT of int
   val protocol_command: string -> (string list -> unit) -> unit
   val reset_tracing: Document_ID.exec -> unit
   val crashes: exn list Synchronized.var
   val init_options: unit -> unit
   val init_options_interactive: unit -> unit
   val init: unit -> unit
+  val init_build: unit -> unit
 end;
 
 structure Isabelle_Process: ISABELLE_PROCESS =
@@ -27,9 +30,17 @@
 val _ = Output.add_mode isabelle_processN Output.default_output Output.default_escape;
 val _ = Markup.add_mode isabelle_processN YXML.output_markup;
 
+val protocol_modes1 = [Syntax_Trans.no_bracketsN, Syntax_Trans.no_type_bracketsN];
+val protocol_modes2 = [isabelle_processN, Pretty.symbolicN];
+
 
 (* protocol commands *)
 
+exception STOP;
+exception EXIT of int;
+
+val is_protocol_exn = fn STOP => true | EXIT _ => true | _ => false;
+
 local
 
 val commands =
@@ -49,7 +60,9 @@
     NONE => error ("Undefined Isabelle protocol command " ^ quote name)
   | SOME cmd =>
       (Runtime.exn_trace_system (fn () => cmd args)
-        handle _ (*sic!*) => error ("Isabelle protocol command failure: " ^ quote name)));
+        handle exn =>
+          if is_protocol_exn exn then Exn.reraise exn
+          else error ("Isabelle protocol command failure: " ^ quote name)));
 
 end;
 
@@ -91,16 +104,39 @@
       end);
 
 
-(* output channels *)
+(* init protocol -- uninterruptible *)
+
+val crashes = Synchronized.var "Isabelle_Process.crashes" ([]: exn list);
+
+local
+
+fun recover crash =
+  (Synchronized.change crashes (cons crash);
+    Output.physical_stderr
+      "Recovered from Isabelle process crash -- see also Isabelle_Process.crashes\n");
+
+in
 
-val serial_props = Markup.serial_properties o serial;
+fun init_protocol modes = Thread_Attributes.uninterruptible (fn _ => fn (address, password) =>
+  let
+    val _ = SHA1.test_samples ()
+      handle exn as Fail msg => (Output.physical_stderr (msg ^ "\n"); Exn.reraise exn);
+
+    val _ = Output.physical_stderr Symbol.STX;
 
-fun init_channels out_stream =
-  let
+
+    (* streams *)
+
+    val (in_stream, out_stream) = Socket_IO.open_streams address;
+    val _ = Byte_Message.write_line out_stream password;
+
     val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream TextIO.stdOut, IO.LINE_BUF);
     val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream TextIO.stdErr, IO.LINE_BUF);
     val _ = BinIO.StreamIO.setBufferMode (BinIO.getOutstream out_stream, IO.BLOCK_BUF);
 
+
+    (* messages *)
+
     val msg_channel = Message_Channel.make out_stream;
 
     fun message name props body =
@@ -115,88 +151,80 @@
               (false, SOME id') => props @ [(Markup.idN, id')]
             | _ => props);
         in message name props' (XML.blob ss) end;
-  in
-    Private_Output.status_fn := standard_message [] Markup.statusN;
-    Private_Output.report_fn := standard_message [] Markup.reportN;
-    Private_Output.result_fn :=
-      (fn props => fn s => standard_message (props @ serial_props ()) Markup.resultN s);
-    Private_Output.writeln_fn := (fn s => standard_message (serial_props ()) Markup.writelnN s);
-    Private_Output.state_fn := (fn s => standard_message (serial_props ()) Markup.stateN s);
-    Private_Output.information_fn :=
-      (fn s => standard_message (serial_props ()) Markup.informationN s);
-    Private_Output.tracing_fn :=
-      (fn s => (update_tracing (); standard_message (serial_props ()) Markup.tracingN s));
-    Private_Output.warning_fn := (fn s => standard_message (serial_props ()) Markup.warningN s);
-    Private_Output.legacy_fn := (fn s => standard_message (serial_props ()) Markup.legacyN s);
-    Private_Output.error_message_fn :=
-      (fn (i, s) => standard_message (Markup.serial_properties i) Markup.errorN s);
-    Private_Output.system_message_fn :=
-      (fn ss => message Markup.systemN [] (XML.blob ss));
-    Private_Output.protocol_message_fn :=
-      (fn props => fn body => message Markup.protocolN props body);
+
+    fun report_message ss =
+      if Context_Position.pide_reports ()
+      then standard_message [] Markup.reportN ss else ();
+
+    val serial_props = Markup.serial_properties o serial;
 
-    Session.init_protocol_handlers ();
-    message Markup.initN [] [XML.Text (Session.welcome ())];
-    msg_channel
-  end;
+    val message_context =
+      Unsynchronized.setmp Private_Output.status_fn (standard_message [] Markup.statusN) #>
+      Unsynchronized.setmp Private_Output.report_fn report_message #>
+      Unsynchronized.setmp Private_Output.result_fn
+        (fn props => fn s => standard_message (props @ serial_props ()) Markup.resultN s) #>
+      Unsynchronized.setmp Private_Output.writeln_fn
+        (fn s => standard_message (serial_props ()) Markup.writelnN s) #>
+      Unsynchronized.setmp Private_Output.state_fn
+        (fn s => standard_message (serial_props ()) Markup.stateN s) #>
+      Unsynchronized.setmp Private_Output.information_fn
+        (fn s => standard_message (serial_props ()) Markup.informationN s) #>
+      Unsynchronized.setmp Private_Output.tracing_fn
+        (fn s => (update_tracing (); standard_message (serial_props ()) Markup.tracingN s)) #>
+      Unsynchronized.setmp Private_Output.warning_fn
+        (fn s => standard_message (serial_props ()) Markup.warningN s) #>
+      Unsynchronized.setmp Private_Output.legacy_fn
+        (fn s => standard_message (serial_props ()) Markup.legacyN s) #>
+      Unsynchronized.setmp Private_Output.error_message_fn
+        (fn (i, s) => standard_message (Markup.serial_properties i) Markup.errorN s) #>
+      Unsynchronized.setmp Private_Output.system_message_fn
+        (fn ss => message Markup.systemN [] (XML.blob ss)) #>
+      Unsynchronized.setmp Private_Output.protocol_message_fn
+        (fn props => fn body => message Markup.protocolN props body) #>
+      Unsynchronized.setmp print_mode
+        ((! print_mode @ #1 modes) |> fold (update op =) (#2 modes));
 
 
-(* protocol loop -- uninterruptible *)
-
-val crashes = Synchronized.var "Isabelle_Process.crashes" ([]: exn list);
-
-local
+    (* protocol *)
 
-fun recover crash =
-  (Synchronized.change crashes (cons crash);
-    Output.physical_stderr
-      "Recovered from Isabelle process crash -- see also Isabelle_Process.crashes\n");
-
-in
+    fun protocol_loop () =
+      let
+        val _ =
+          (case Byte_Message.read_message in_stream of
+            NONE => raise STOP
+          | SOME [] => Output.system_message "Isabelle process: no input"
+          | SOME (name :: args) => run_command name args)
+          handle exn =>
+            if is_protocol_exn exn then Exn.reraise exn
+            else (Runtime.exn_system_message exn handle crash => recover crash);
+      in protocol_loop () end;
 
-fun loop stream =
-  let
-    val continue =
-      (case Byte_Message.read_message stream of
-        NONE => false
-      | SOME [] => (Output.system_message "Isabelle process: no input"; true)
-      | SOME (name :: args) => (run_command name args; true))
-      handle exn => (Runtime.exn_system_message exn handle crash => recover crash; true);
+    fun protocol () =
+     (Session.init_protocol_handlers ();
+      message Markup.initN [] [XML.Text (Session.welcome ())];
+      protocol_loop ());
+
+    val result = Exn.capture (message_context protocol) ();
+
+
+    (* shutdown *)
+
+    val _ = Future.shutdown ();
+    val _ = Execution.reset ();
+    val _ = Message_Channel.shutdown msg_channel;
+    val _ = BinIO.closeIn in_stream;
+    val _ = BinIO.closeOut out_stream;
+
   in
-    if continue then loop stream
-    else (Future.shutdown (); Execution.reset (); ())
-  end;
+    (case result of
+      Exn.Exn STOP => ()
+    | Exn.Exn (EXIT rc) => exit rc
+    | _ => Exn.release result)
+  end);
 
 end;
 
 
-(* init protocol *)
-
-val default_modes1 = [Syntax_Trans.no_bracketsN, Syntax_Trans.no_type_bracketsN];
-val default_modes2 = [isabelle_processN, Pretty.symbolicN];
-
-val init_protocol = Thread_Attributes.uninterruptible (fn _ => fn (address, password) =>
-  let
-    val _ = SHA1.test_samples ()
-      handle exn as Fail msg => (Output.physical_stderr (msg ^ "\n"); Exn.reraise exn);
-    val _ = Output.physical_stderr Symbol.STX;
-
-    val _ = Context.put_generic_context NONE;
-    val _ =
-      Unsynchronized.change print_mode
-        (fn mode => (mode @ default_modes1) |> fold (update op =) default_modes2);
-
-    val (in_stream, out_stream) = Socket_IO.open_streams address;
-    val _ = Byte_Message.write_line out_stream password;
-    val msg_channel = init_channels out_stream;
-    val _ = loop in_stream;
-    val _ = Message_Channel.shutdown msg_channel;
-    val _ = Private_Output.init_channels ();
-
-    val _ = print_mode := [];
-  in () end);
-
-
 (* init options *)
 
 fun init_options () =
@@ -218,14 +246,17 @@
 
 (* generic init *)
 
-fun init () =
+fun init_modes modes =
   let
     val address = Options.default_string \<^system_option>\<open>system_channel_address\<close>;
     val password = Options.default_string \<^system_option>\<open>system_channel_password\<close>;
   in
     if address <> "" andalso password <> ""
-    then init_protocol (address, password)
+    then init_protocol modes (address, password)
     else init_options ()
   end;
 
+fun init () = init_modes (protocol_modes1, protocol_modes2);
+fun init_build () = init_modes ([], protocol_modes2);
+
 end;
--- a/src/Pure/System/isabelle_process.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/isabelle_process.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -12,39 +12,18 @@
 
 object Isabelle_Process
 {
-  def start(session: Session,
+  def apply(
+    session: Session,
     options: Options,
+    sessions_structure: Sessions.Structure,
+    store: Sessions.Store,
     logic: String = "",
-    args: List[String] = Nil,
-    dirs: List[Path] = Nil,
+    raw_ml_system: Boolean = false,
+    use_prelude: List[String] = Nil,
+    eval_main: String = "",
     modes: List[String] = Nil,
     cwd: JFile = null,
-    env: Map[String, String] = Isabelle_System.settings(),
-    sessions_structure: Option[Sessions.Structure] = None,
-    store: Option[Sessions.Store] = None,
-    phase_changed: Session.Phase => Unit = null)
-  {
-    if (phase_changed != null)
-      session.phase_changed += Session.Consumer("Isabelle_Process")(phase_changed)
-
-    session.start(receiver =>
-      Isabelle_Process(options, logic = logic, args = args, dirs = dirs, modes = modes,
-        cwd = cwd, env = env, receiver = receiver, xml_cache = session.xml_cache,
-        sessions_structure = sessions_structure, store = store))
-  }
-
-  def apply(
-    options: Options,
-    logic: String = "",
-    args: List[String] = Nil,
-    dirs: List[Path] = Nil,
-    modes: List[String] = Nil,
-    cwd: JFile = null,
-    env: Map[String, String] = Isabelle_System.settings(),
-    receiver: Prover.Receiver = (msg: Prover.Message) => Output.writeln(msg.toString, stdout = true),
-    xml_cache: XML.Cache = XML.make_cache(),
-    sessions_structure: Option[Sessions.Structure] = None,
-    store: Option[Sessions.Store] = None): Prover =
+    env: Map[String, String] = Isabelle_System.settings()): Isabelle_Process =
   {
     val channel = System_Channel()
     val process =
@@ -52,12 +31,51 @@
         val channel_options =
           options.string.update("system_channel_address", channel.address).
             string.update("system_channel_password", channel.password)
-        ML_Process(channel_options, logic = logic, args = args, dirs = dirs, modes = modes,
-            cwd = cwd, env = env, sessions_structure = sessions_structure, store = store)
+        ML_Process(channel_options, sessions_structure, store,
+          logic = logic, raw_ml_system = raw_ml_system,
+          use_prelude = use_prelude, eval_main = eval_main,
+          modes = modes, cwd = cwd, env = env)
       }
       catch { case exn @ ERROR(_) => channel.shutdown(); throw exn }
     process.stdin.close
 
-    new Prover(receiver, xml_cache, channel, process)
+    new Isabelle_Process(session, channel, process)
   }
 }
+
+class Isabelle_Process private(session: Session, channel: System_Channel, process: Bash.Process)
+{
+  private val startup = Future.promise[String]
+  private val terminated = Future.promise[Process_Result]
+
+  session.phase_changed +=
+    Session.Consumer(getClass.getName) {
+      case Session.Ready =>
+        startup.fulfill("")
+      case Session.Terminated(result) =>
+        if (!result.ok && !startup.is_finished) {
+          val syslog = session.syslog_content()
+          val err = "Session startup failed" + (if (syslog.isEmpty) "" else ":\n" + syslog)
+          startup.fulfill(err)
+        }
+        terminated.fulfill(result)
+      case _ =>
+    }
+
+  session.start(receiver => new Prover(receiver, session.xml_cache, channel, process))
+
+  def await_startup(): Isabelle_Process =
+    startup.join match {
+      case "" => this
+      case err => session.stop(); error(err)
+    }
+
+  def await_shutdown(): Process_Result =
+  {
+    val result = terminated.join
+    session.stop()
+    result
+  }
+
+  def terminate { process.terminate }
+}
--- a/src/Pure/System/isabelle_system.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/isabelle_system.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -46,16 +46,28 @@
 
   /** implicit settings environment **/
 
+  abstract class Service
+
   @volatile private var _settings: Option[Map[String, String]] = None
+  @volatile private var _services: Option[List[Service]] = None
+
+  private def uninitialized: Boolean = _services.isEmpty   // unsynchronized check
 
   def settings(): Map[String, String] =
   {
-    if (_settings.isEmpty) init()  // unsynchronized check
+    if (uninitialized) init()
     _settings.get
   }
 
-  def init(isabelle_root: String = "", cygwin_root: String = ""): Unit = synchronized {
-    if (_settings.isEmpty) {
+  def services(): List[Service] =
+  {
+    if (uninitialized) init()
+    _services.get
+  }
+
+  def init(isabelle_root: String = "", cygwin_root: String = ""): Unit = synchronized
+  {
+    if (uninitialized) {
       val isabelle_root1 =
         bootstrap_directory(isabelle_root, "ISABELLE_ROOT", "isabelle.root", "Isabelle root")
 
@@ -124,12 +136,32 @@
       }
       _settings = Some(settings)
       set_cygwin_root()
+
+      val variable = "ISABELLE_SCALA_SERVICES"
+      val services =
+        for (name <- space_explode(':', settings.getOrElse(variable, getenv_error(variable))))
+        yield {
+          def err(msg: String): Nothing =
+            error("Bad entry " + quote(name) + " in " + variable + "\n" + msg)
+          try {
+            Class.forName(name).asInstanceOf[Class[Service]]
+              .getDeclaredConstructor().newInstance()
+          }
+          catch {
+            case _: ClassNotFoundException => err("Class not found")
+            case exn: Throwable => err(Exn.message(exn))
+          }
+        }
+      _services = Some(services)
     }
   }
 
 
   /* getenv */
 
+  private def getenv_error(name: String): Nothing =
+    error("Undefined Isabelle environment variable: " + quote(name))
+
   def getenv(name: String, env: Map[String, String] = settings()): String =
     env.getOrElse(name, "")
 
@@ -287,7 +319,7 @@
         proc.getInputStream.close
         proc.getErrorStream.close
         proc.destroy
-        Thread.interrupted
+        Exn.Interrupt.dispose()
       }
     (output, rc)
   }
@@ -365,24 +397,6 @@
     Path.split(getenv_strict("ISABELLE_COMPONENTS"))
 
 
-  /* classes */
-
-  def init_classes[A](variable: String): List[A] =
-  {
-    for (name <- space_explode(':', Isabelle_System.getenv_strict(variable)))
-    yield {
-      def err(msg: String): Nothing =
-        error("Bad entry " + quote(name) + " in " + variable + "\n" + msg)
-
-      try { Class.forName(name).asInstanceOf[Class[A]].newInstance() }
-      catch {
-        case _: ClassNotFoundException => err("Class not found")
-        case exn: Throwable => err(Exn.message(exn))
-      }
-    }
-  }
-
-
   /* default logic */
 
   def default_logic(args: String*): String =
--- a/src/Pure/System/isabelle_tool.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/isabelle_tool.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -98,8 +98,7 @@
   /* internal tools */
 
   private lazy val internal_tools: List[Isabelle_Tool] =
-    Isabelle_System.init_classes[Isabelle_Scala_Tools]("ISABELLE_SCALA_TOOLS")
-      .flatMap(_.tools.toList)
+    Isabelle_System.services.collect { case c: Isabelle_Scala_Tools => c.tools.toList }.flatten
 
   private def list_internal(): List[(String, String)] =
     for (tool <- internal_tools.toList)
@@ -108,7 +107,7 @@
   private def find_internal(name: String): Option[List[String] => Unit] =
     internal_tools.collectFirst({
       case tool if tool.name == name =>
-        args => Command_Line.tool0 { tool.body(args) }
+        args => Command_Line.tool { tool.body(args) }
       })
 
 
@@ -116,7 +115,7 @@
 
   def main(args: Array[String])
   {
-    Command_Line.tool0 {
+    Command_Line.tool {
       args.toList match {
         case Nil | List("-?") =>
           val tool_descriptions =
@@ -140,7 +139,7 @@
 
 sealed case class Isabelle_Tool(name: String, description: String, body: List[String] => Unit)
 
-class Isabelle_Scala_Tools(val tools: Isabelle_Tool*)
+class Isabelle_Scala_Tools(val tools: Isabelle_Tool*) extends Isabelle_System.Service
 
 class Tools extends Isabelle_Scala_Tools(
   Build.isabelle_tool,
--- a/src/Pure/System/linux.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/linux.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -62,12 +62,12 @@
   def check_reboot_required(): Unit =
     if (reboot_required()) error("Reboot required")
 
-  def package_update(progress: Progress = No_Progress): Unit =
+  def package_update(progress: Progress = new Progress): Unit =
     progress.bash(
       """apt-get update -y && apt-get upgrade -y && apt autoremove -y""",
       echo = true).check
 
-  def package_install(packages: List[String], progress: Progress = No_Progress): Unit =
+  def package_install(packages: List[String], progress: Progress = new Progress): Unit =
     progress.bash("apt-get install -y -- " + Bash.strings(packages), echo = true).check
 
   def package_installed(name: String): Boolean =
--- a/src/Pure/System/message_channel.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/message_channel.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -59,11 +59,11 @@
   let
     val mbox = Mailbox.create ();
     val thread =
-      Standard_Thread.fork {name = "channel", stack_limit = NONE, interrupts = false}
+      Isabelle_Thread.fork {name = "channel", stack_limit = NONE, interrupts = false}
         (message_output mbox stream);
     fun send msg = Mailbox.send mbox (SOME msg);
     fun shutdown () =
-      (Mailbox.send mbox NONE; Mailbox.await_empty mbox; Standard_Thread.join thread);
+      (Mailbox.send mbox NONE; Mailbox.await_empty mbox; Isabelle_Thread.join thread);
   in Message_Channel {send = send, shutdown = shutdown} end;
 
 end;
--- a/src/Pure/System/numa.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/numa.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -26,7 +26,7 @@
       }
 
     if (numa_nodes_linux.is_file) {
-      space_explode(',', File.read(numa_nodes_linux).trim).flatMap(read(_))
+      space_explode(',', File.read(numa_nodes_linux).trim).flatMap(read)
     }
     else Nil
   }
--- a/src/Pure/System/options.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/options.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -55,7 +55,7 @@
           case word :: rest if word == strip => rest
           case _ => words
         }
-      Word.implode(words1.map(Word.perhaps_capitalize(_)))
+      Word.implode(words1.map(Word.perhaps_capitalize))
     }
 
     def unknown: Boolean = typ == Unknown
@@ -70,19 +70,19 @@
   private val OPTIONS = Path.explode("etc/options")
   private val PREFS = Path.explode("$ISABELLE_HOME_USER/etc/preferences")
 
-  val options_syntax =
+  val options_syntax: Outer_Syntax =
     Outer_Syntax.empty + ":" + "=" + "--" + Symbol.comment + Symbol.comment_decoded +
       (SECTION, Keyword.DOCUMENT_HEADING) +
       (PUBLIC, Keyword.BEFORE_COMMAND) +
       (OPTION, Keyword.THY_DECL)
 
-  val prefs_syntax = Outer_Syntax.empty + "="
+  val prefs_syntax: Outer_Syntax = Outer_Syntax.empty + "="
 
   trait Parser extends Parse.Parser
   {
-    val option_name = atom("option name", _.is_name)
-    val option_type = atom("option type", _.is_name)
-    val option_value =
+    val option_name: Parser[String] = atom("option name", _.is_name)
+    val option_type: Parser[String] = atom("option type", _.is_name)
+    val option_value: Parser[String] =
       opt(token("-", tok => tok.is_sym_ident && tok.content == "-")) ~ atom("nat", _.is_nat) ^^
         { case s ~ n => if (s.isDefined) "-" + n else n } |
       atom("option value", tok => tok.is_name || tok.is_float)
--- a/src/Pure/System/process_result.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/process_result.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -6,6 +6,28 @@
 
 package isabelle
 
+object Process_Result
+{
+  def print_return_code(rc: Int): String = "Return code: " + rc + rc_text(rc)
+  def print_rc(rc: Int): String = "return code " + rc + rc_text(rc)
+
+  def rc_text(rc: Int): String =
+    return_code_text.get(rc) match {
+      case None => ""
+      case Some(text) => " (" + text + ")"
+    }
+  private val return_code_text =
+    Map(0 -> "OK",
+      1 -> "ERROR",
+      2 -> "FAILURE",
+      130 -> "INTERRUPT",
+      131 -> "QUIT SIGNAL",
+      137 -> "KILL SIGNAL",
+      138 -> "BUS ERROR",
+      139 -> "SEGMENTATION VIOLATION",
+      143 -> "TERMINATION SIGNAL")
+}
+
 final case class Process_Result(
   rc: Int,
   out_lines: List[String] = Nil,
@@ -15,7 +37,11 @@
 {
   def out: String = cat_lines(out_lines)
   def err: String = cat_lines(err_lines)
-  def errors(errs: List[String]): Process_Result = copy(err_lines = err_lines ::: errs)
+
+  def output(outs: List[String]): Process_Result =
+    copy(out_lines = out_lines ::: outs.flatMap(split_lines))
+  def errors(errs: List[String]): Process_Result =
+    copy(err_lines = err_lines ::: errs.flatMap(split_lines))
   def error(err: String): Process_Result = errors(List(err))
 
   def was_timeout: Process_Result = copy(rc = 1, timeout = true)
@@ -32,6 +58,9 @@
 
   def check: Process_Result = check_rc(_ == 0)
 
+  def print_return_code: String = Process_Result.print_return_code(rc)
+  def print_rc: String = Process_Result.print_rc(rc)
+
   def print: Process_Result =
   {
     Output.warning(err)
--- a/src/Pure/System/progress.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/progress.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -34,10 +34,17 @@
   def echo_error_message(msg: String) { echo(Output.error_message_text(msg)) }
 
   def timeit[A](message: String = "", enabled: Boolean = true)(e: => A): A =
-    Timing.timeit(message, enabled, echo(_))(e)
+    Timing.timeit(message, enabled, echo)(e)
 
-  def stopped: Boolean = false
-  def interrupt_handler[A](e: => A): A = e
+  @volatile protected var is_stopped = false
+  def stop { is_stopped = true }
+  def stopped: Boolean =
+  {
+    if (Thread.interrupted) is_stopped = true
+    is_stopped
+  }
+
+  def interrupt_handler[A](e: => A): A = POSIX_Interrupt.handler { stop } { e }
   def expose_interrupt() { if (stopped) throw Exn.Interrupt() }
   override def toString: String = if (stopped) "Progress(stopped)" else "Progress"
 
@@ -55,8 +62,6 @@
   }
 }
 
-object No_Progress extends Progress
-
 class Console_Progress(verbose: Boolean = false, stderr: Boolean = false) extends Progress
 {
   override def echo(msg: String): Unit =
@@ -64,15 +69,6 @@
 
   override def theory(theory: Progress.Theory): Unit =
     if (verbose) echo(theory.message)
-
-  @volatile private var is_stopped = false
-  override def interrupt_handler[A](e: => A): A =
-    POSIX_Interrupt.handler { is_stopped = true } { e }
-  override def stopped: Boolean =
-  {
-    if (Thread.interrupted) is_stopped = true
-    is_stopped
-  }
 }
 
 class File_Progress(path: Path, verbose: Boolean = false) extends Progress
--- a/src/Pure/System/tty_loop.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/System/tty_loop.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -10,13 +10,14 @@
 import java.io.{IOException, Writer, Reader, InputStreamReader, BufferedReader}
 
 
-class TTY_Loop(writer: Writer, reader: Reader,
-  writer_lock: AnyRef = new Object,
-  interrupt: Option[() => Unit] = None)
+class TTY_Loop(
+  writer: Writer,
+  reader: Reader,
+  writer_lock: AnyRef = new Object)
 {
-  private val console_output = Future.thread[Unit]("console_output") {
+  private val console_output = Future.thread[Unit]("console_output", uninterruptible = true) {
     try {
-      var result = new StringBuilder(100)
+      val result = new StringBuilder(100)
       var finished = false
       while (!finished) {
         var c = -1
@@ -40,32 +41,25 @@
     catch { case e: IOException => case Exn.Interrupt() => }
   }
 
-  private val console_input = Future.thread[Unit]("console_input") {
+  private val console_input = Future.thread[Unit]("console_input", uninterruptible = true) {
     val console_reader = new BufferedReader(new InputStreamReader(System.in))
-    def body
-    {
-      try {
-        var finished = false
-        while (!finished) {
-          console_reader.readLine() match {
-            case null =>
-              writer.close()
-              finished = true
-            case line =>
-              writer_lock.synchronized {
-                writer.write(line)
-                writer.write("\n")
-                writer.flush()
-              }
-          }
+    try {
+      var finished = false
+      while (!finished) {
+        console_reader.readLine() match {
+          case null =>
+            writer.close()
+            finished = true
+          case line =>
+            writer_lock.synchronized {
+              writer.write(line)
+              writer.write("\n")
+              writer.flush()
+            }
         }
       }
-      catch { case e: IOException => case Exn.Interrupt() => }
     }
-    interrupt match {
-      case None => body
-      case Some(int) => POSIX_Interrupt.handler { int() } { body }
-    }
+    catch { case e: IOException => case Exn.Interrupt() => }
   }
 
   def join { console_output.join; console_input.join }
--- a/src/Pure/Thy/export.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/export.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -83,7 +83,7 @@
         let
           val path = Path.expand (Path.explode ("$ISABELLE_EXPORT_TMP/export" ^ serial_string ()));
           val _ = YXML.write_body path body;
-        in Protocol_Message.inline (#2 function) (tl args @ [(Markup.fileN, Path.implode path)]) end
+        in Protocol_Message.marker (#2 function) (tl args @ [(Markup.fileN, Path.implode path)]) end
       else raise Output.Protocol_Message props
   | [] => raise Output.Protocol_Message props);
 
--- a/src/Pure/Thy/export.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/export.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -143,7 +143,7 @@
       regex.pattern.matcher(compound_name(theory_name, name)).matches
   }
 
-  def make_entry(session_name: String, args: Markup.Export.Args, body: Bytes,
+  def make_entry(session_name: String, args: Protocol.Export.Args, body: Bytes,
     cache: XZ.Cache = XZ.cache()): Entry =
   {
     Entry(session_name, args.theory_name, args.name, args.executable,
@@ -213,7 +213,7 @@
             (results, true)
           })
 
-    def apply(session_name: String, args: Markup.Export.Args, body: Bytes): Unit =
+    def apply(session_name: String, args: Protocol.Export.Args, body: Bytes): Unit =
       consumer.send(make_entry(session_name, args, body, cache = cache) -> args.strict)
 
     def shutdown(close: Boolean = false): List[String] =
@@ -299,7 +299,7 @@
     store: Sessions.Store,
     session_name: String,
     export_dir: Path,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     export_prune: Int = 0,
     export_list: Boolean = false,
     export_patterns: List[String] = Nil)
@@ -312,7 +312,7 @@
         // list
         if (export_list) {
           (for ((theory_name, name) <- export_names) yield compound_name(theory_name, name)).
-            sorted.foreach(progress.echo(_))
+            sorted.foreach(progress.echo)
         }
 
         // export
@@ -348,7 +348,7 @@
 
   /* Isabelle tool wrapper */
 
-  val default_export_dir = Path.explode("export")
+  val default_export_dir: Path = Path.explode("export")
 
   val isabelle_tool = Isabelle_Tool("export", "retrieve theory exports", args =>
   {
--- a/src/Pure/Thy/export_theory.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/export_theory.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -23,14 +23,14 @@
       else None
 
     def theories: List[Theory] =
-      theory_graph.topological_order.flatMap(theory(_))
+      theory_graph.topological_order.flatMap(theory)
   }
 
   def read_session(
     store: Sessions.Store,
     sessions_structure: Sessions.Structure,
     session_name: String,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     cache: Term.Cache = Term.make_cache()): Session =
   {
     val thys =
--- a/src/Pure/Thy/file_format.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/file_format.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -12,12 +12,12 @@
   sealed case class Theory_Context(name: Document.Node.Name, content: String)
 
 
-  /* environment */
+  /* registry */
 
-  def environment(): Environment =
-    new Environment(Isabelle_System.init_classes[File_Format]("ISABELLE_FILE_FORMATS"))
+  lazy val registry: Registry =
+    new Registry(Isabelle_System.services.collect { case c: File_Format => c })
 
-  final class Environment private [File_Format](file_formats: List[File_Format])
+  final class Registry private [File_Format](file_formats: List[File_Format])
   {
     override def toString: String = file_formats.mkString("File_Format.Environment(", ",", ")")
 
@@ -53,10 +53,10 @@
   object No_Agent extends Agent
 }
 
-trait File_Format
+trait File_Format extends Isabelle_System.Service
 {
   def format_name: String
-  override def toString = format_name
+  override def toString: String = "File_Format(" + format_name + ")"
 
   def file_ext: String
   def detect(name: String): Boolean = name.endsWith("." + file_ext)
--- a/src/Pure/Thy/html.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/html.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -368,7 +368,7 @@
         (if (entry.is_italic) List("  font-style: italic;") else Nil) :::
         List("}"))
 
-    ("/* Isabelle fonts */" :: Isabelle_Fonts.fonts(hidden = true).map(font_face(_)))
+    ("/* Isabelle fonts */" :: Isabelle_Fonts.fonts(hidden = true).map(font_face))
       .mkString("", "\n\n", "\n")
   }
 
--- a/src/Pure/Thy/latex.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/latex.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -10,6 +10,7 @@
 import java.io.{File => JFile}
 
 import scala.annotation.tailrec
+import scala.util.matching.Regex
 
 
 object Latex
@@ -102,7 +103,7 @@
 
     object File_Line_Error
     {
-      val Pattern = """^(.*?\.\w\w\w):(\d+): (.*)$""".r
+      val Pattern: Regex = """^(.*?\.\w\w\w):(\d+): (.*)$""".r
       def unapply(line: String): Option[(Path, Int, String)] =
         line match {
           case Pattern(file, Value.Int(line), message) =>
--- a/src/Pure/Thy/present.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/present.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -10,7 +10,7 @@
   val theory_qualifier: theory -> string
   val document_option: Options.T -> {enabled: bool, disabled: bool}
   val document_variants: Options.T -> (string * string) list
-  val init: HTML.symbols -> bool -> bool -> Path.T -> string -> string -> (string * string) list ->
+  val init: HTML.symbols -> bool -> Path.T -> string -> string -> (string * string) list ->
     (Path.T * Path.T) list -> Path.T -> string * string -> bool -> unit
   val finish: unit -> unit
   val theory_output: theory -> string list -> unit
@@ -165,28 +165,25 @@
 
 (* init session *)
 
-fun init symbols build info info_path doc document_output doc_variants doc_files graph_file
+fun init symbols info info_path doc document_output doc_variants doc_files graph_file
     (chapter, name) verbose =
-  if not build andalso not info andalso doc = "" then
-    (Synchronized.change browser_info (K empty_browser_info); session_info := NONE)
-  else
-    let
-      val doc_output =
-        if document_output = "" then NONE else SOME (Path.explode document_output);
+  let
+    val doc_output =
+      if document_output = "" then NONE else SOME (Path.explode document_output);
+
+    val documents = if doc = "" orelse null doc_files then [] else doc_variants;
+    val readme = if File.exists readme_html_path then SOME readme_html_path else NONE;
 
-      val documents = if doc = "" orelse null doc_files then [] else doc_variants;
-      val readme = if File.exists readme_html_path then SOME readme_html_path else NONE;
-
-      val docs =
-        (case readme of NONE => [] | SOME p => [(Url.File p, "README")]) @
-          map (fn (name, _) => (Url.File (Path.ext doc (Path.basic name)), name)) documents;
-    in
-      session_info :=
-        SOME (make_session_info (symbols, name, chapter, info_path, info, doc,
-          doc_output, doc_files, graph_file, documents, verbose, readme));
-      Synchronized.change browser_info (K empty_browser_info);
-      add_html_index (0, HTML.begin_session_index symbols name (Url.File session_graph_path) docs)
-    end;
+    val docs =
+      (case readme of NONE => [] | SOME p => [(Url.File p, "README")]) @
+        map (fn (name, _) => (Url.File (Path.ext doc (Path.basic name)), name)) documents;
+  in
+    session_info :=
+      SOME (make_session_info (symbols, name, chapter, info_path, info, doc,
+        doc_output, doc_files, graph_file, documents, verbose, readme));
+    Synchronized.change browser_info (K empty_browser_info);
+    add_html_index (0, HTML.begin_session_index symbols name (Url.File session_graph_path) docs)
+  end;
 
 
 (* isabelle tool wrappers *)
--- a/src/Pure/Thy/sessions.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/sessions.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -116,7 +116,7 @@
   }
 
   def deps(sessions_structure: Structure,
-      progress: Progress = No_Progress,
+      progress: Progress = new Progress,
       inlined_files: Boolean = false,
       verbose: Boolean = false,
       list_files: Boolean = false,
@@ -143,21 +143,19 @@
 
     val doc_names = Doc.doc_names()
 
+    val bootstrap =
+      Sessions.Base.bootstrap(
+        sessions_structure.session_directories,
+        sessions_structure.global_theories)
+
     val session_bases =
-      (Map.empty[String, Base] /: sessions_structure.imports_topological_order)({
+      (Map("" -> bootstrap) /: sessions_structure.imports_topological_order)({
         case (session_bases, session_name) =>
           progress.expose_interrupt()
 
           val info = sessions_structure(session_name)
           try {
-            val parent_base: Sessions.Base =
-              info.parent match {
-                case None =>
-                  Base.bootstrap(
-                    sessions_structure.session_directories,
-                    sessions_structure.global_theories)
-                case Some(parent) => session_bases(parent)
-              }
+            val parent_base = session_bases(info.parent.getOrElse(""))
 
             val imports_base: Sessions.Base =
             {
@@ -235,7 +233,7 @@
                 (Graph_Display.empty_graph /: required_subgraph.topological_order)(
                   { case (g, session) =>
                       val a = session_node(session)
-                      val bs = required_subgraph.imm_preds(session).toList.map(session_node(_))
+                      val bs = required_subgraph.imm_preds(session).toList.map(session_node)
                       ((g /: (a :: bs))(_.default_node(_, Nil)) /: bs)(_.add_edge(_, a)) })
 
               (graph0 /: dependencies.entries)(
@@ -271,7 +269,7 @@
                   if !ok(path.canonical_file)
                   path1 = File.relative_path(info.dir.canonical, path).getOrElse(path)
                 } yield (path1, name)).toList
-              val bad_dirs = (for { (path1, _) <- bad } yield path1.toString).toSet.toList.sorted
+              val bad_dirs = (for { (path1, _) <- bad } yield path1.toString).distinct.sorted
 
               val errs1 =
                 for { (path1, name) <- bad }
@@ -336,8 +334,6 @@
   /* base info */
 
   sealed case class Base_Info(
-    options: Options,
-    dirs: List[Path],
     session: String,
     sessions_structure: Structure,
     errors: List[String],
@@ -349,7 +345,7 @@
 
   def base_info(options: Options,
     session: String,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     dirs: List[Path] = Nil,
     include_sessions: List[String] = Nil,
     session_ancestor: Option[String] = None,
@@ -371,7 +367,7 @@
           deps.get(ancestor.get) match {
             case Some(ancestor_base)
             if !selected_sessions.imports_requirements(List(ancestor.get)).contains(session) =>
-              ancestor_base.loaded_theories.defined(_)
+              ancestor_base.loaded_theories.defined _
             case _ =>
               error("Bad ancestor " + quote(ancestor.get) + " for session " + quote(session))
           }
@@ -420,7 +416,7 @@
 
     val deps1 = Sessions.deps(selected_sessions1, progress = progress)
 
-    Base_Info(options, dirs, session1, full_sessions1, deps1.errors, deps1(session1), infos1)
+    Base_Info(session1, full_sessions1, deps1.errors, deps1(session1), infos1)
   }
 
 
@@ -654,7 +650,7 @@
 
     def check_sessions(names: List[String])
     {
-      val bad_sessions = SortedSet(names.filterNot(defined(_)): _*).toList
+      val bad_sessions = SortedSet(names.filterNot(defined): _*).toList
       if (bad_sessions.nonEmpty)
         error("Undefined session(s): " + commas_quote(bad_sessions))
     }
@@ -712,7 +708,7 @@
     def selection_deps(
       options: Options,
       selection: Selection,
-      progress: Progress = No_Progress,
+      progress: Progress = new Progress,
       loading_sessions: Boolean = false,
       inlined_files: Boolean = false,
       verbose: Boolean = false): Deps =
@@ -746,8 +742,8 @@
 
   /* parser */
 
-  val ROOT = Path.explode("ROOT")
-  val ROOTS = Path.explode("ROOTS")
+  val ROOT: Path = Path.explode("ROOT")
+  val ROOTS: Path = Path.explode("ROOTS")
 
   private val CHAPTER = "chapter"
   private val SESSION = "session"
@@ -761,7 +757,7 @@
   private val DOCUMENT_FILES = "document_files"
   private val EXPORT_FILES = "export_files"
 
-  val root_syntax =
+  val root_syntax: Outer_Syntax =
     Outer_Syntax.empty + "(" + ")" + "+" + "," + "=" + "[" + "]" +
       GLOBAL + IN +
       (CHAPTER, Keyword.THY_DECL) +
@@ -897,7 +893,7 @@
 
   def directories(dirs: List[Path], select_dirs: List[Path]): List[(Boolean, Path)] =
   {
-    val default_dirs = Isabelle_System.components().filter(is_session_dir(_))
+    val default_dirs = Isabelle_System.components().filter(is_session_dir)
     for { (select, dir) <- (default_dirs ::: dirs).map((false, _)) ::: select_dirs.map((true, _)) }
     yield (select, dir.canonical)
   }
@@ -1075,7 +1071,7 @@
       input_dirs.map(_ + heap(name)).find(_.is_file)
 
     def find_heap_digest(name: String): Option[String] =
-      find_heap(name).flatMap(read_heap_digest(_))
+      find_heap(name).flatMap(read_heap_digest)
 
     def the_heap(name: String): Path =
       find_heap(name) getOrElse
@@ -1107,7 +1103,7 @@
         if (output || has_session_info(db, name)) Some(db) else { db.close; None }
       }
       else if (output) Some(SQLite.open_database(output_database(name)))
-      else input_dirs.map(_ + database(name)).find(_.is_file).map(SQLite.open_database(_))
+      else input_dirs.map(_ + database(name)).find(_.is_file).map(SQLite.open_database)
     }
 
     def open_database(name: String, output: Boolean = false): SQL.Database =
--- a/src/Pure/Thy/thy_info.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/thy_info.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -352,7 +352,9 @@
     val dir = Path.dir thy_path;
     val header = Thy_Header.make (name, pos) imports keywords;
 
-    val _ = Position.reports (map #2 imports ~~ map Theory.get_markup parents);
+    val _ =
+      (imports ~~ parents) |> List.app (fn ((_, pos), thy) =>
+        Context_Position.reports_global thy [(pos, Theory.get_markup thy)]);
 
     val exec_id = Document_ID.make ();
     val _ =
--- a/src/Pure/Thy/thy_syntax.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Thy/thy_syntax.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -136,7 +136,7 @@
             val new_commands = insert_text(Some(cmd), text) - cmd
             edit_text(rest.toList ::: es, new_commands)
 
-          case Some((cmd, cmd_start)) =>
+          case Some((cmd, _)) =>
             edit_text(es, insert_text(Some(cmd), e.text))
 
           case None =>
--- a/src/Pure/Tools/build.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/build.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -46,7 +46,7 @@
 
 (* session timing *)
 
-fun session_timing name verbose f x =
+fun session_timing session_name verbose f x =
   let
     val start = Timing.start ();
     val y = f x;
@@ -58,10 +58,10 @@
 
     val timing_props =
       [("threads", threads)] @ Markup.timing_properties timing @ [("factor", factor)];
-    val _ = writeln ("\fTiming = " ^ YXML.string_of_body (XML.Encode.properties timing_props));
+    val _ = Protocol_Message.marker "Timing" timing_props;
     val _ =
       if verbose then
-        Output.physical_stderr ("Timing " ^ name ^ " (" ^
+        Output.physical_stderr ("Timing " ^ session_name ^ " (" ^
           threads ^ " threads, " ^ Timing.message timing ^ ", factor " ^ factor ^ ")\n")
       else ();
   in y end;
@@ -73,7 +73,7 @@
   (case props of
     function :: args =>
       if function = Markup.ML_statistics orelse function = Markup.task_statistics then
-        Protocol_Message.inline (#2 function) args
+        Protocol_Message.marker (#2 function) args
       else if function = Markup.command_timing then
         let
           val name = the_default "" (Properties.get args Markup.nameN);
@@ -86,15 +86,16 @@
           if is_significant then
             (case approximative_id name pos of
               SOME id =>
-                Protocol_Message.inline (#2 function) (Markup.command_timing_properties id elapsed)
+                Protocol_Message.marker (#2 function)
+                  (Markup.command_timing_properties id elapsed)
             | NONE => ())
           else ()
         end
       else if function = Markup.theory_timing then
-        Protocol_Message.inline (#2 function) args
+        Protocol_Message.marker (#2 function) args
       else
         (case Markup.dest_loading_theory props of
-          SOME name => writeln ("\floading_theory = " ^ name)
+          SOME name => Protocol_Message.marker_text "loading_theory" name
         | NONE => Export.protocol_message props output)
   | [] => raise Output.Protocol_Message props);
 
@@ -133,14 +134,13 @@
 datatype args = Args of
  {symbol_codes: (string * int) list,
   command_timings: Properties.T list,
-  do_output: bool,
   verbose: bool,
   browser_info: Path.T,
   document_files: (Path.T * Path.T) list,
   graph_file: Path.T,
   parent_name: string,
   chapter: string,
-  name: string,
+  session_name: string,
   master_dir: Path.T,
   theories: (Options.T * (string * Position.T) list) list,
   session_positions: (string * Properties.T) list,
@@ -154,31 +154,31 @@
   let
     open XML.Decode;
     val position = Position.of_properties o properties;
-    val (symbol_codes, (command_timings, (do_output, (verbose, (browser_info,
-      (document_files, (graph_file, (parent_name, (chapter, (name, (master_dir,
+    val (symbol_codes, (command_timings, (verbose, (browser_info,
+      (document_files, (graph_file, (parent_name, (chapter, (session_name, (master_dir,
       (theories, (session_positions, (session_directories, (doc_names, (global_theories,
-      (loaded_theories, bibtex_entries))))))))))))))))) =
-      pair (list (pair string int)) (pair (list properties) (pair bool (pair bool (pair string
+      (loaded_theories, bibtex_entries)))))))))))))))) =
+      pair (list (pair string int)) (pair (list properties) (pair bool (pair string
         (pair (list (pair string string)) (pair string (pair string (pair string (pair string
           (pair string
             (pair (((list (pair Options.decode (list (pair string position))))))
               (pair (list (pair string properties))
                 (pair (list (pair string string)) (pair (list string)
-                  (pair (list (pair string string)) (pair (list string) (list string)))))))))))))))))
+                  (pair (list (pair string string)) (pair (list string) (list string))))))))))))))))
       (YXML.parse_body yxml);
   in
-    Args {symbol_codes = symbol_codes, command_timings = command_timings, do_output = do_output,
+    Args {symbol_codes = symbol_codes, command_timings = command_timings,
       verbose = verbose, browser_info = Path.explode browser_info,
       document_files = map (apply2 Path.explode) document_files,
       graph_file = Path.explode graph_file, parent_name = parent_name, chapter = chapter,
-      name = name, master_dir = Path.explode master_dir, theories = theories,
+      session_name = session_name, master_dir = Path.explode master_dir, theories = theories,
       session_positions = session_positions, session_directories = session_directories,
       doc_names = doc_names, global_theories = global_theories, loaded_theories = loaded_theories,
       bibtex_entries = bibtex_entries}
   end;
 
-fun build_session (Args {symbol_codes, command_timings, do_output, verbose, browser_info,
-    document_files, graph_file, parent_name, chapter, name, master_dir, theories, session_positions,
+fun build_session (Args {symbol_codes, command_timings, verbose, browser_info, document_files,
+    graph_file, parent_name, chapter, session_name, master_dir, theories, session_positions,
     session_directories, doc_names, global_theories, loaded_theories, bibtex_entries}) =
   let
     val symbols = HTML.make_symbols symbol_codes;
@@ -194,7 +194,6 @@
     val _ =
       Session.init
         symbols
-        do_output
         (Options.default_bool "browser_info")
         browser_info
         (Options.default_string "document")
@@ -203,48 +202,69 @@
         document_files
         graph_file
         parent_name
-        (chapter, name)
+        (chapter, session_name)
         verbose;
 
     val last_timing = get_timings (fold update_timings command_timings empty_timings);
 
     val res1 =
       theories |>
-        (List.app (build_theories symbols bibtex_entries last_timing name master_dir)
-          |> session_timing name verbose
+        (List.app (build_theories symbols bibtex_entries last_timing session_name master_dir)
+          |> session_timing session_name verbose
           |> Exn.capture);
     val res2 = Exn.capture Session.finish ();
 
     val _ = Resources.finish_session_base ();
     val _ = Par_Exn.release_all [res1, res2];
+    val _ =
+      if session_name = Context.PureN
+      then Theory.install_pure (Thy_Info.get_theory Context.PureN) else ();
   in () end;
 
-(*command-line tool*)
+
+(* command-line tool *)
+
+fun inline_errors exn =
+  Runtime.exn_message_list exn
+  |> List.app (fn msg => Protocol_Message.marker_text "error_message" (YXML.content_of msg));
+
 fun build args_file =
   let
     val _ = SHA1.test_samples ();
     val _ = Options.load_default ();
     val _ = Isabelle_Process.init_options ();
     val args = decode_args (File.read (Path.explode args_file));
-    fun error_message msg = writeln ("\ferror_message = " ^ encode_lines (YXML.content_of msg));
     val _ =
       Unsynchronized.setmp Private_Output.protocol_message_fn protocol_message
         build_session args
-      handle exn => (List.app error_message (Runtime.exn_message_list exn); Exn.reraise exn);
+      handle exn => (inline_errors exn; Exn.reraise exn);
     val _ = Private_Output.protocol_message_fn := Output.protocol_message_undefined;
     val _ = Options.reset_default ();
   in () end;
 
-(*PIDE version*)
+
+(* PIDE version *)
+
+fun build_session_finished errs =
+  XML.Encode.list XML.Encode.string errs
+  |> Output.protocol_message Markup.build_session_finished;
+
 val _ =
   Isabelle_Process.protocol_command "build_session"
     (fn [args_yxml] =>
         let
           val args = decode_args args_yxml;
-          val result = (build_session args; "") handle exn =>
-            (Runtime.exn_message exn handle _ (*sic!*) =>
-              "Exception raised, but failed to print details!");
-        in Output.protocol_message Markup.build_session_finished [XML.Text result] end
+          val errs =
+            Future.interruptible_task (fn () => (build_session args; [])) () handle exn =>
+              (Runtime.exn_message_list exn handle _ (*sic!*) =>
+                (build_session_finished ["CRASHED"];
+                  raise Isabelle_Process.EXIT 2));
+          val _ = build_session_finished errs;
+        in
+          if null errs
+          then raise Isabelle_Process.STOP
+          else raise Isabelle_Process.EXIT 1
+        end
       | _ => raise Match);
 
 end;
--- a/src/Pure/Tools/build.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/build.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -8,7 +8,7 @@
 package isabelle
 
 
-import scala.collection.SortedSet
+import scala.collection.{SortedSet, mutable}
 import scala.annotation.tailrec
 
 
@@ -34,11 +34,11 @@
   {
     type Timings = (List[Properties.T], Double)
 
-    def load_timings(progress: Progress, store: Sessions.Store, name: String): Timings =
+    def load_timings(progress: Progress, store: Sessions.Store, session_name: String): Timings =
     {
       val no_timings: Timings = (Nil, 0.0)
 
-      store.access_database(name) match {
+      store.access_database(session_name) match {
         case None => no_timings
         case Some(db) =>
           def ignore_error(msg: String) =
@@ -47,9 +47,9 @@
             no_timings
           }
           try {
-            val command_timings = store.read_command_timings(db, name)
+            val command_timings = store.read_command_timings(db, session_name)
             val session_timing =
-              store.read_session_timing(db, name) match {
+              store.read_session_timing(db, session_name) match {
                 case Markup.Elapsed(t) => t
                 case _ => 0.0
               }
@@ -68,11 +68,11 @@
       : Map[String, Double] =
     {
       val maximals = sessions_structure.build_graph.maximals.toSet
-      def desc_timing(name: String): Double =
+      def desc_timing(session_name: String): Double =
       {
-        if (maximals.contains(name)) timing(name)
+        if (maximals.contains(session_name)) timing(session_name)
         else {
-          val descendants = sessions_structure.build_descendants(List(name)).toSet
+          val descendants = sessions_structure.build_descendants(List(session_name)).toSet
           val g = sessions_structure.build_graph.restrict(descendants)
           (0.0 :: g.maximals.flatMap(desc => {
             val ps = g.all_preds(List(desc))
@@ -150,86 +150,53 @@
 
   /* PIDE protocol handler */
 
-  class Handler(progress: Progress, session: Session, session_name: String)
-    extends Session.Protocol_Handler
-  {
-    val result_error: Promise[String] = Future.promise
-
-    override def exit() { result_error.cancel }
-
-    private def build_session_finished(msg: Prover.Protocol_Output): Boolean =
-    {
-      val error_message =
-        try { Pretty.string_of(Symbol.decode_yxml(msg.text)) }
-        catch { case ERROR(msg) => msg }
-      result_error.fulfill(error_message)
-      session.send_stop()
-      true
-    }
-
-    private def loading_theory(msg: Prover.Protocol_Output): Boolean =
-      msg.properties match {
-        case Markup.Loading_Theory(name) =>
-          progress.theory(Progress.Theory(name, session = session_name))
-          true
-        case _ => false
-      }
-
-    val functions =
-      List(
-        Markup.BUILD_SESSION_FINISHED -> build_session_finished _,
-        Markup.LOADING_THEORY -> loading_theory _)
-  }
-
-
   /* job: running prover process */
 
   private class Job(progress: Progress,
-    name: String,
+    session_name: String,
     val info: Sessions.Info,
     deps: Sessions.Deps,
     store: Sessions.Store,
-    do_output: Boolean,
+    do_store: Boolean,
     verbose: Boolean,
-    pide: Boolean,
     val numa_node: Option[Int],
     command_timings: List[Properties.T])
   {
     val options = NUMA.policy_options(info.options, numa_node)
 
-    val sessions_structure = deps.sessions_structure
+    private val sessions_structure = deps.sessions_structure
 
     private val graph_file = Isabelle_System.tmp_file("session_graph", "pdf")
-    isabelle.graphview.Graph_File.write(options, graph_file, deps(name).session_graph_display)
+    graphview.Graph_File.write(options, graph_file, deps(session_name).session_graph_display)
 
     private val export_tmp_dir = Isabelle_System.tmp_dir("export")
     private val export_consumer =
-      Export.consumer(store.open_database(name, output = true), cache = store.xz_cache)
+      Export.consumer(store.open_database(session_name, output = true), cache = store.xz_cache)
 
     private val future_result: Future[Process_Result] =
-      Future.thread("build") {
+      Future.thread("build", uninterruptible = true) {
         val parent = info.parent.getOrElse("")
-        val base = deps(name)
+        val base = deps(parent)
         val args_yxml =
           YXML.string_of_body(
             {
               import XML.Encode._
-              pair(list(pair(string, int)), pair(list(properties), pair(bool, pair(bool,
+              pair(list(pair(string, int)), pair(list(properties), pair(bool,
                 pair(Path.encode, pair(list(pair(Path.encode, Path.encode)), pair(string,
                 pair(string, pair(string, pair(string, pair(Path.encode,
                 pair(list(pair(Options.encode, list(pair(string, properties)))),
                 pair(list(pair(string, properties)),
                 pair(list(pair(string, string)),
                 pair(list(string), pair(list(pair(string, string)),
-                pair(list(string), list(string))))))))))))))))))(
-              (Symbol.codes, (command_timings, (do_output, (verbose,
+                pair(list(string), list(string)))))))))))))))))(
+              (Symbol.codes, (command_timings, (verbose,
                 (store.browser_info, (info.document_files, (File.standard_path(graph_file),
-                (parent, (info.chapter, (name, (Path.current,
+                (parent, (info.chapter, (session_name, (Path.current,
                 (info.theories,
                 (sessions_structure.session_positions,
                 (sessions_structure.dest_session_directories,
                 (base.doc_names, (base.global_theories.toList,
-                (base.loaded_theories.keys, info.bibtex_entries.map(_.info)))))))))))))))))))
+                (base.loaded_theories.keys, info.bibtex_entries.map(_.info))))))))))))))))))
             })
 
         val env =
@@ -237,73 +204,175 @@
             ("ISABELLE_EXPORT_TMP" -> File.standard_path(export_tmp_dir)) +
             ("ISABELLE_ML_DEBUGGER" -> options.bool("ML_debugger").toString)
 
-        def save_heap: String =
-          (if (info.theories.isEmpty) "" else "ML_Heap.share_common_data (); ") +
-            "ML_Heap.save_child " +
-            ML_Syntax.print_string_bytes(File.platform_path(store.output_heap(name)))
+        val is_pure = Sessions.is_pure(session_name)
+
+        val use_prelude = if (is_pure) Thy_Header.ml_roots.map(_._1) else Nil
 
-        if (pide && !Sessions.is_pure(name)) {
+        val eval_store =
+          if (do_store) {
+            (if (info.theories.nonEmpty) List("ML_Heap.share_common_data ()") else Nil) :::
+            List("ML_Heap.save_child " +
+              ML_Syntax.print_string_bytes(File.platform_path(store.output_heap(session_name))))
+          }
+          else Nil
+
+        if (options.bool("pide_build")) {
           val resources = new Resources(sessions_structure, deps(parent))
           val session = new Session(options, resources)
-          val handler = new Handler(progress, session, name)
-          session.init_protocol_handler(handler)
+
+          val build_session_errors: Promise[List[String]] = Future.promise
+          val stdout = new StringBuilder(1000)
+          val messages = new mutable.ListBuffer[String]
+          val command_timings = new mutable.ListBuffer[Properties.T]
+          val theory_timings = new mutable.ListBuffer[Properties.T]
+          val runtime_statistics = new mutable.ListBuffer[Properties.T]
+          val task_statistics = new mutable.ListBuffer[Properties.T]
+
+          session.init_protocol_handler(new Session.Protocol_Handler
+            {
+              override def exit() { build_session_errors.cancel }
 
-          val session_result = Future.promise[Process_Result]
+              private def build_session_finished(msg: Prover.Protocol_Output): Boolean =
+              {
+                val errors =
+                  try {
+                    for (err <- XML.Decode.list(x => x)(Symbol.decode_yxml(msg.text)))
+                    yield {
+                      val prt = Protocol_Message.expose_no_reports(err)
+                      Pretty.string_of(prt, metric = Symbol.Metric)
+                    }
+                  }
+                  catch { case ERROR(err) => List(err) }
+                build_session_errors.fulfill(errors)
+                true
+              }
+
+              private def loading_theory(msg: Prover.Protocol_Output): Boolean =
+                msg.properties match {
+                  case Markup.Loading_Theory(name) =>
+                    progress.theory(Progress.Theory(name, session = session_name))
+                    true
+                  case _ => false
+                }
 
-          Isabelle_Process.start(session, options, logic = parent, cwd = info.dir.file, env = env,
-            sessions_structure = Some(sessions_structure), store = Some(store),
-            phase_changed =
-            {
-              case Session.Ready => session.protocol_command("build_session", args_yxml)
-              case Session.Terminated(result) => session_result.fulfill(result)
-              case _ =>
+              private def export(msg: Prover.Protocol_Output): Boolean =
+                msg.properties match {
+                  case Protocol.Export(args) =>
+                    export_consumer(session_name, args, msg.bytes)
+                    true
+                  case _ => false
+                }
+
+              private def command_timing(msg: Prover.Protocol_Output): Boolean =
+                msg.properties match {
+                  case Markup.Command_Timing(props) => command_timings += props; true
+                  case _ => false
+                }
+
+              private def theory_timing(msg: Prover.Protocol_Output): Boolean =
+                msg.properties match {
+                  case Markup.Theory_Timing(props) => theory_timings += props; true
+                  case _ => false
+                }
+
+              private def ml_stats(msg: Prover.Protocol_Output): Boolean =
+                msg.properties match {
+                  case Markup.ML_Statistics(props) => runtime_statistics += props; true
+                  case _ => false
+                }
+
+              private def task_stats(msg: Prover.Protocol_Output): Boolean =
+                msg.properties match {
+                  case Markup.Task_Statistics(props) => task_statistics += props; true
+                  case _ => false
+                }
+
+              val functions =
+                List(
+                  Markup.Build_Session_Finished.name -> build_session_finished,
+                  Markup.Loading_Theory.name -> loading_theory,
+                  Markup.EXPORT -> export,
+                  Markup.Command_Timing.name -> command_timing,
+                  Markup.Theory_Timing.name -> theory_timing,
+                  Markup.ML_Statistics.name -> ml_stats,
+                  Markup.Task_Statistics.name -> task_stats)
             })
 
-          val result = session_result.join
-          handler.result_error.join match {
-            case "" => result
-            case msg =>
-              result.copy(
-                rc = result.rc max 1,
-                out_lines = result.out_lines ::: split_lines(Output.error_message_text(msg)))
+          session.all_messages += Session.Consumer[Any]("build_session_output")
+            {
+              case msg: Prover.Output =>
+                val message = msg.message
+                if (msg.is_stdout) {
+                  stdout ++= Symbol.encode(XML.content(message))
+                }
+                else if (Protocol.is_exported(message)) {
+                  messages +=
+                    Symbol.encode(Protocol.message_text(List(message), metric = Symbol.Metric))
+                }
+              case _ =>
+            }
+
+          val eval_main = Command_Line.ML_tool("Isabelle_Process.init_build ()" :: eval_store)
+
+          val process =
+            Isabelle_Process(session, options, sessions_structure, store,
+              logic = parent, raw_ml_system = is_pure,
+              use_prelude = use_prelude, eval_main = eval_main,
+              cwd = info.dir.file, env = env)
+
+          val errors =
+            Isabelle_Thread.interrupt_handler(_ => process.terminate) {
+              Exn.capture { process.await_startup } match {
+                case Exn.Res(_) =>
+                  session.protocol_command("build_session", args_yxml)
+                  build_session_errors.join_result
+                case Exn.Exn(exn) => Exn.Res(List(Exn.message(exn)))
+              }
+            }
+
+          val process_result =
+            Isabelle_Thread.interrupt_handler(_ => process.terminate) { process.await_shutdown }
+          val process_output =
+            stdout.toString :: messages.toList :::
+            command_timings.toList.map(Protocol.Command_Timing_Marker.apply) :::
+            theory_timings.toList.map(Protocol.Theory_Timing_Marker.apply) :::
+            runtime_statistics.toList.map(Protocol.ML_Statistics_Marker.apply) :::
+            task_statistics.toList.map(Protocol.Task_Statistics_Marker.apply)
+
+          val result = process_result.output(process_output)
+
+          errors match {
+            case Exn.Res(Nil) => result
+            case Exn.Res(errs) =>
+              result.error_rc.output(
+                errs.flatMap(s => split_lines(Output.error_message_text(s))) :::
+                  errs.map(Protocol.Error_Message_Marker.apply))
+            case Exn.Exn(Exn.Interrupt()) =>
+              if (result.ok) result.copy(rc = Exn.Interrupt.return_code) else result
+            case Exn.Exn(exn) => throw exn
           }
         }
         else {
           val args_file = Isabelle_System.tmp_file("build")
           File.write(args_file, args_yxml)
 
-          val eval =
-            "Command_Line.tool0 (fn () => (" +
-            "Build.build " + ML_Syntax.print_string_bytes(File.standard_path(args_file)) +
-            (if (Sessions.is_pure(name)) "; Theory.install_pure (Thy_Info.get_theory Context.PureN)"
-             else "") + (if (do_output) "; " + save_heap else "") + "));"
+          val eval_build =
+            "Build.build " + ML_Syntax.print_string_bytes(File.standard_path(args_file))
+          val eval_main = Command_Line.ML_tool(eval_build :: eval_store)
 
           val process =
-            if (Sessions.is_pure(name)) {
-              ML_Process(options, raw_ml_system = true, cwd = info.dir.file,
-                args =
-                  (for ((root, _) <- Thy_Header.ml_roots) yield List("--use", root)).flatten :::
-                  List("--eval", eval),
-                env = env, sessions_structure = Some(deps.sessions_structure), store = Some(store),
-                cleanup = () => args_file.delete)
-            }
-            else {
-              ML_Process(options, parent, List("--eval", eval), cwd = info.dir.file,
-                env = env, sessions_structure = Some(deps.sessions_structure), store = Some(store),
-                cleanup = () => args_file.delete)
-            }
+            ML_Process(options, deps.sessions_structure, store,
+              logic = parent, raw_ml_system = is_pure,
+              use_prelude = use_prelude, eval_main = eval_main,
+              cwd = info.dir.file, env = env, cleanup = () => args_file.delete)
 
-          process.result(
-            progress_stdout = (line: String) =>
-              Library.try_unprefix("\floading_theory = ", line) match {
-                case Some(theory) =>
-                  progress.theory(Progress.Theory(theory, session = name))
-                case None =>
-                  for {
-                    text <- Library.try_unprefix("\fexport = ", line)
-                    (args, path) <-
-                      Markup.Export.dest_inline(XML.Decode.properties(YXML.parse_body(text)))
-                  } {
+          Isabelle_Thread.interrupt_handler(_ => process.terminate) {
+            process.result(
+              progress_stdout =
+                {
+                  case Protocol.Loading_Theory_Marker(theory) =>
+                    progress.theory(Progress.Theory(theory, session = session_name))
+                  case Protocol.Export.Marker((args, path)) =>
                     val body =
                       try { Bytes.read(path) }
                       catch {
@@ -311,15 +380,16 @@
                           error("Failed to read export " + quote(args.compound_name) + "\n" + msg)
                       }
                     path.file.delete
-                    export_consumer(name, args, body)
-                  }
-              },
-            progress_limit =
-              options.int("process_output_limit") match {
-                case 0 => None
-                case m => Some(m * 1000000L)
-              },
-            strict = false)
+                    export_consumer(session_name, args, body)
+                  case _ =>
+                },
+              progress_limit =
+                options.int("process_output_limit") match {
+                  case 0 => None
+                  case m => Some(m * 1000000L)
+                },
+              strict = false)
+          }
         }
       }
 
@@ -337,7 +407,7 @@
     {
       val result0 = future_result.join
       val result1 =
-        export_consumer.shutdown(close = true).map(Output.error_message_text(_)) match {
+        export_consumer.shutdown(close = true).map(Output.error_message_text) match {
           case Nil => result0
           case errs => result0.errors(errs).error_rc
         }
@@ -345,7 +415,7 @@
       Isabelle_System.rm_tree(export_tmp_dir)
 
       if (result1.ok)
-        Present.finish(progress, store.browser_info, graph_file, info, name)
+        Present.finish(progress, store.browser_info, graph_file, info, session_name)
 
       graph_file.delete
 
@@ -363,8 +433,8 @@
         else result1
 
       val heap_digest =
-        if (result2.ok && do_output && store.output_heap(name).is_file)
-          Some(Sessions.write_heap_digest(store.output_heap(name)))
+        if (result2.ok && do_store && store.output_heap(session_name).is_file)
+          Some(Sessions.write_heap_digest(store.output_heap(session_name)))
         else None
 
       (result2, heap_digest)
@@ -381,7 +451,7 @@
     def cancelled(name: String): Boolean = results(name)._1.isEmpty
     def apply(name: String): Process_Result = results(name)._1.getOrElse(Process_Result(1))
     def info(name: String): Sessions.Info = results(name)._2
-    val rc =
+    val rc: Int =
       (0 /: results.iterator.map(
         { case (_, (Some(r), _)) => r.rc case (_, (None, _)) => 1 }))(_ max _)
     def ok: Boolean = rc == 0
@@ -391,7 +461,7 @@
 
   def build(
     options: Options,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     check_unknown_files: Boolean = false,
     build_heap: Boolean = false,
     clean_build: Boolean = false,
@@ -407,7 +477,6 @@
     soft_build: Boolean = false,
     verbose: Boolean = false,
     export_files: Boolean = false,
-    pide: Boolean = false,
     requirements: Boolean = false,
     all_sessions: Boolean = false,
     base_sessions: List[String] = Nil,
@@ -417,7 +486,12 @@
     sessions: List[String] = Nil,
     selection: Sessions.Selection = Sessions.Selection.empty): Results =
   {
-    val build_options = options.int.update("completion_limit", 0).bool.update("ML_statistics", true)
+    val build_options =
+      options +
+        "ML_statistics" +
+        "completion_limit=0" +
+        "editor_tracing_messages=0" +
+        "pide_reports=false"
 
     val store = Sessions.store(build_options)
 
@@ -429,10 +503,12 @@
     val full_sessions =
       Sessions.load_structure(build_options, dirs = dirs, select_dirs = select_dirs, infos = infos)
 
-    def sources_stamp(deps: Sessions.Deps, name: String): String =
+    def sources_stamp(deps: Sessions.Deps, session_name: String): String =
     {
       val digests =
-        full_sessions(name).meta_digest :: deps.sources(name) ::: deps.imported_sources(name)
+        full_sessions(session_name).meta_digest ::
+        deps.sources(session_name) :::
+        deps.imported_sources(session_name)
       SHA1.digest(cat_lines(digests.map(_.toString).sorted)).toString
     }
 
@@ -518,8 +594,7 @@
 
     def sleep()
     {
-      try { Thread.sleep(500) }
-      catch { case Exn.Interrupt() => Exn.Interrupt.impose() }
+      Isabelle_Thread.interrupt_handler(_ => progress.stop) { Time.seconds(0.5).sleep }
     }
 
     val numa_nodes = new NUMA.Nodes(numa_shuffling)
@@ -535,89 +610,93 @@
 
       if (pending.is_empty) results
       else {
-        if (progress.stopped)
+        if (progress.stopped) {
           for ((_, (_, job)) <- running) job.terminate
+        }
 
         running.find({ case (_, (_, job)) => job.is_finished }) match {
-          case Some((name, (input_heaps, job))) =>
+          case Some((session_name, (input_heaps, job))) =>
             //{{{ finish job
 
             val (process_result, heap_digest) = job.join
 
-            val log_lines = process_result.out_lines.filterNot(_.startsWith("\f"))
+            val log_lines = process_result.out_lines.filterNot(Protocol_Message.Marker.test)
             val process_result_tail =
             {
               val tail = job.info.options.int("process_output_tail")
               process_result.copy(
                 out_lines =
-                  "(see also " + store.output_log(name).file.toString + ")" ::
+                  "(see also " + store.output_log(session_name).file.toString + ")" ::
                   (if (tail == 0) log_lines else log_lines.drop(log_lines.length - tail max 0)))
             }
 
 
             // write log file
             if (process_result.ok) {
-              File.write_gzip(store.output_log_gz(name), terminate_lines(log_lines))
+              File.write_gzip(store.output_log_gz(session_name), terminate_lines(log_lines))
             }
-            else File.write(store.output_log(name), terminate_lines(log_lines))
+            else File.write(store.output_log(session_name), terminate_lines(log_lines))
 
             // write database
             {
               val build_log =
-                Build_Log.Log_File(name, process_result.out_lines).
+                Build_Log.Log_File(session_name, process_result.out_lines).
                   parse_session_info(
                     command_timings = true,
                     theory_timings = true,
                     ml_statistics = true,
                     task_statistics = true)
 
-              using(store.open_database(name, output = true))(db =>
-                store.write_session_info(db, name,
+              using(store.open_database(session_name, output = true))(db =>
+                store.write_session_info(db, session_name,
                   build_log =
                     if (process_result.timeout) build_log.error("Timeout") else build_log,
                   build =
-                    Session_Info(sources_stamp(deps, name), input_heaps, heap_digest,
+                    Session_Info(sources_stamp(deps, session_name), input_heaps, heap_digest,
                       process_result.rc)))
             }
 
             // messages
-            process_result.err_lines.foreach(progress.echo(_))
+            process_result.err_lines.foreach(progress.echo)
 
             if (process_result.ok)
-              progress.echo("Finished " + name + " (" + process_result.timing.message_resources + ")")
+              progress.echo(
+                "Finished " + session_name + " (" + process_result.timing.message_resources + ")")
             else {
-              progress.echo(name + " FAILED")
+              progress.echo(session_name + " FAILED")
               if (!process_result.interrupted) progress.echo(process_result_tail.out)
             }
 
-            loop(pending - name, running - name,
-              results + (name -> Result(false, heap_digest, Some(process_result_tail), job.info)))
+            loop(pending - session_name, running - session_name,
+              results +
+                (session_name -> Result(false, heap_digest, Some(process_result_tail), job.info)))
             //}}}
           case None if running.size < (max_jobs max 1) =>
             //{{{ check/start next job
-            pending.dequeue(running.isDefinedAt(_)) match {
-              case Some((name, info)) =>
+            pending.dequeue(running.isDefinedAt) match {
+              case Some((session_name, info)) =>
                 val ancestor_results =
-                  deps.sessions_structure.build_requirements(List(name)).filterNot(_ == name).
-                    map(results(_))
+                  deps.sessions_structure.build_requirements(List(session_name)).
+                    filterNot(_ == session_name).map(results(_))
                 val ancestor_heaps = ancestor_results.flatMap(_.heap_digest)
 
-                val do_output = build_heap || Sessions.is_pure(name) || queue.is_inner(name)
+                val do_store =
+                  build_heap || Sessions.is_pure(session_name) || queue.is_inner(session_name)
 
                 val (current, heap_digest) =
                 {
-                  store.access_database(name) match {
+                  store.access_database(session_name) match {
                     case Some(db) =>
-                      using(db)(store.read_build(_, name)) match {
+                      using(db)(store.read_build(_, session_name)) match {
                         case Some(build) =>
-                          val heap_digest = store.find_heap_digest(name)
+                          val heap_digest = store.find_heap_digest(session_name)
                           val current =
                             !fresh_build &&
                             build.ok &&
-                            build.sources == sources_stamp(deps, name) &&
+                            build.sources == sources_stamp(deps, session_name) &&
                             build.input_heaps == ancestor_heaps &&
                             build.output_heap == heap_digest &&
-                            !(do_output && heap_digest.isEmpty)
+                            !(do_store && heap_digest.isEmpty)
                           (current, heap_digest)
                         case None => (false, None)
                       }
@@ -627,29 +706,32 @@
                 val all_current = current && ancestor_results.forall(_.current)
 
                 if (all_current)
-                  loop(pending - name, running,
-                    results + (name -> Result(true, heap_digest, Some(Process_Result(0)), info)))
+                  loop(pending - session_name, running,
+                    results +
+                      (session_name -> Result(true, heap_digest, Some(Process_Result(0)), info)))
                 else if (no_build) {
-                  if (verbose) progress.echo("Skipping " + name + " ...")
-                  loop(pending - name, running,
-                    results + (name -> Result(false, heap_digest, Some(Process_Result(1)), info)))
+                  if (verbose) progress.echo("Skipping " + session_name + " ...")
+                  loop(pending - session_name, running,
+                    results +
+                      (session_name -> Result(false, heap_digest, Some(Process_Result(1)), info)))
                 }
                 else if (ancestor_results.forall(_.ok) && !progress.stopped) {
-                  progress.echo((if (do_output) "Building " else "Running ") + name + " ...")
+                  progress.echo((if (do_store) "Building " else "Running ") + session_name + " ...")
 
-                  store.clean_output(name)
-                  using(store.open_database(name, output = true))(store.init_session_info(_, name))
+                  store.clean_output(session_name)
+                  using(store.open_database(session_name, output = true))(
+                    store.init_session_info(_, session_name))
 
-                  val numa_node = numa_nodes.next(used_node(_))
+                  val numa_node = numa_nodes.next(used_node)
                   val job =
-                    new Job(progress, name, info, deps, store, do_output, verbose, pide = pide,
-                      numa_node, queue.command_timings(name))
-                  loop(pending, running + (name -> (ancestor_heaps, job)), results)
+                    new Job(progress, session_name, info, deps, store, do_store, verbose,
+                      numa_node, queue.command_timings(session_name))
+                  loop(pending, running + (session_name -> (ancestor_heaps, job)), results)
                 }
                 else {
-                  progress.echo(name + " CANCELLED")
-                  loop(pending - name, running,
-                    results + (name -> Result(false, heap_digest, None, info)))
+                  progress.echo(session_name + " CANCELLED")
+                  loop(pending - session_name, running,
+                    results + (session_name -> Result(false, heap_digest, None, info)))
                 }
               case None => sleep(); loop(pending, running, results)
             }
@@ -667,7 +749,7 @@
         progress.echo_warning("Nothing to build")
         Map.empty[String, Result]
       }
-      else loop(queue, Map.empty, Map.empty)
+      else Isabelle_Thread.uninterruptible { loop(queue, Map.empty, Map.empty) }
 
     val results =
       new Results(
@@ -681,7 +763,7 @@
           progress.echo("Exporting " + info.name + " ...")
           for ((dir, prune, pats) <- info.export_files) {
             Export.export_files(store, name, info.dir + dir,
-              progress = if (verbose) progress else No_Progress,
+              progress = if (verbose) progress else new Progress,
               export_prune = prune,
               export_patterns = pats)
           }
@@ -730,7 +812,6 @@
     var base_sessions: List[String] = Nil
     var select_dirs: List[Path] = Nil
     var numa_shuffling = false
-    var pide = false
     var requirements = false
     var soft_build = false
     var exclude_session_groups: List[String] = Nil
@@ -756,7 +837,6 @@
     -B NAME      include session NAME and all descendants
     -D DIR       include session directory and select its sessions
     -N           cyclic shuffling of NUMA CPU nodes (performance tuning)
-    -P           build via PIDE protocol
     -R           operate on requirements of selected sessions
     -S           soft build: only observe changes of sources, not heap images
     -X NAME      exclude sessions from group NAME and all descendants
@@ -781,7 +861,6 @@
       "B:" -> (arg => base_sessions = base_sessions ::: List(arg)),
       "D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))),
       "N" -> (_ => numa_shuffling = true),
-      "P" -> (_ => pide = true),
       "R" -> (_ => requirements = true),
       "S" -> (_ => soft_build = true),
       "X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)),
@@ -831,7 +910,6 @@
           soft_build = soft_build,
           verbose = verbose,
           export_files = export_files,
-          pide = pide,
           requirements = requirements,
           all_sessions = all_sessions,
           base_sessions = base_sessions,
@@ -859,7 +937,7 @@
   /* build logic image */
 
   def build_logic(options: Options, logic: String,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     build_heap: Boolean = false,
     dirs: List[Path] = Nil,
     fresh: Boolean = false,
--- a/src/Pure/Tools/debugger.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/debugger.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -22,7 +22,7 @@
   if msg = "" then ()
   else
     Output.protocol_message
-      (Markup.debugger_output (Standard_Thread.the_name ()))
+      (Markup.debugger_output (Isabelle_Thread.get_name ()))
       [XML.Text (Markup.markup (kind, Markup.serial_properties (serial ())) msg)];
 
 val writeln_message = output_message Markup.writelnN;
@@ -250,7 +250,7 @@
         (SOME (fn (_, break) =>
           if not (is_debugging ()) andalso (! break orelse is_break () orelse is_stepping ())
           then
-            (case Standard_Thread.get_name () of
+            (case try Isabelle_Thread.get_name () of
               SOME thread_name => debugger_loop thread_name
             | NONE => ())
           else ()))));
--- a/src/Pure/Tools/debugger.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/debugger.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -23,7 +23,7 @@
   {
     def size: Int = debug_states.length + 1
     def reset: Context = copy(index = 0)
-    def select(i: Int) = copy(index = i + 1)
+    def select(i: Int): Context = copy(index = i + 1)
 
     def thread_state: Option[Debug_State] = debug_states.headOption
 
@@ -143,8 +143,8 @@
 
     val functions =
       List(
-        Markup.DEBUGGER_STATE -> debugger_state _,
-        Markup.DEBUGGER_OUTPUT -> debugger_output _)
+        Markup.DEBUGGER_STATE -> debugger_state,
+        Markup.DEBUGGER_OUTPUT -> debugger_output)
   }
 }
 
@@ -155,7 +155,7 @@
   private val state = Synchronized(Debugger.State())
 
   private val delay_update =
-    Standard_Thread.delay_first(session.output_delay) {
+    Delay.first(session.output_delay) {
       session.debugger_updates.post(Debugger.Update)
     }
 
--- a/src/Pure/Tools/doc.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/doc.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -46,7 +46,7 @@
 
   private def release_notes(): List[Entry] =
     Section("Release Notes", true) ::
-      Path.split(Isabelle_System.getenv_strict("ISABELLE_DOCS_RELEASE_NOTES")).flatMap(text_file(_))
+      Path.split(Isabelle_System.getenv_strict("ISABELLE_DOCS_RELEASE_NOTES")).flatMap(text_file)
 
   private def examples(): List[Entry] =
     Section("Examples", true) ::
--- a/src/Pure/Tools/dump.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/dump.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -94,7 +94,7 @@
     def apply(
       options: Options,
       aspects: List[Aspect] = Nil,
-      progress: Progress = No_Progress,
+      progress: Progress = new Progress,
       dirs: List[Path] = Nil,
       select_dirs: List[Path] = Nil,
       selection: Sessions.Selection = Sessions.Selection.empty,
@@ -106,9 +106,9 @@
         val options0 = if (NUMA.enabled) NUMA.policy_options(options) else options
         val options1 =
           options0 +
-            "completion_limit=0" +
             "ML_statistics=false" +
             "parallel_proofs=0" +
+            "completion_limit=0" +
             "editor_tracing_messages=0" +
             "editor_presentation"
         (options1 /: aspects)({ case (opts, aspect) => (opts /: aspect.options)(_ + _) })
@@ -363,7 +363,7 @@
           session.use_theories(used_theories.map(_.theory),
             unicode_symbols = unicode_symbols,
             progress = progress,
-            commit = Some(Consumer.apply _))
+            commit = Some(Consumer.apply))
 
         val bad_theories = Consumer.shutdown()
         val bad_msgs =
@@ -395,7 +395,7 @@
     options: Options,
     logic: String,
     aspects: List[Aspect] = Nil,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     log: Logger = No_Logger,
     dirs: List[Path] = Nil,
     select_dirs: List[Path] = Nil,
@@ -463,7 +463,7 @@
   Dump cumulative PIDE session database, with the following aspects:
 
 """ + Library.prefix_lines("    ", show_aspects) + "\n",
-      "A:" -> (arg => aspects = Library.distinct(space_explode(',', arg)).map(the_aspect(_))),
+      "A:" -> (arg => aspects = Library.distinct(space_explode(',', arg)).map(the_aspect)),
       "B:" -> (arg => base_sessions = base_sessions ::: List(arg)),
       "D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))),
       "O:" -> (arg => output_dir = Path.explode(arg)),
--- a/src/Pure/Tools/fontforge.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/fontforge.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -38,7 +38,7 @@
     }
 
     if (s.nonEmpty && s(0) == '\\') err('\\')
-    s.iterator.map(escape(_)).mkString(q.toString, "", q.toString)
+    s.iterator.map(escape).mkString(q.toString, "", q.toString)
   }
 
   def file_name(path: Path): Script = string(File.standard_path(path))
@@ -121,7 +121,7 @@
     }
 
     def set: Script =
-      List(fontname, familyname, fullname, weight, copyright, fontversion).map(string(_)).
+      List(fontname, familyname, fullname, weight, copyright, fontversion).map(string).
         mkString("SetFontNames(", ",", ")")
   }
 
--- a/src/Pure/Tools/mkroot.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/mkroot.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -25,7 +25,7 @@
     init_repos: Boolean = false,
     title: String = "",
     author: String = "",
-    progress: Progress = No_Progress)
+    progress: Progress = new Progress)
   {
     Isabelle_System.mkdirs(session_dir)
 
--- a/src/Pure/Tools/phabricator.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/phabricator.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -39,7 +39,7 @@
 
   val daemon_user = "phabricator"
 
-  val sshd_config = Path.explode("/etc/ssh/sshd_config")
+  val sshd_config: Path = Path.explode("/etc/ssh/sshd_config")
 
 
   /* installation parameters */
@@ -59,7 +59,7 @@
 
   val default_mailers: Path = Path.explode("mailers.json")
 
-  val default_system_port = SSH.default_port
+  val default_system_port: Int = SSH.default_port
   val alternative_system_port = 222
   val default_server_port = 2222
 
@@ -69,7 +69,7 @@
 
   /** global configuration **/
 
-  val global_config = Path.explode("/etc/" + isabelle_phabricator_name(ext = "conf"))
+  val global_config: Path = Path.explode("/etc/" + isabelle_phabricator_name(ext = "conf"))
 
   def global_config_script(
     init: String = "",
@@ -162,7 +162,7 @@
       else {
         val config = get_config(name)
         val result = progress.bash(Bash.strings(more_args), cwd = config.home.file, echo = true)
-        if (!result.ok) error("Return code: " + result.rc.toString)
+        if (!result.ok) error(result.print_return_code)
       }
     })
 
@@ -193,7 +193,7 @@
     command
   }
 
-  def mercurial_setup(mercurial_source: String, progress: Progress = No_Progress)
+  def mercurial_setup(mercurial_source: String, progress: Progress = new Progress)
   {
     progress.echo("\nMercurial installation from source " + quote(mercurial_source) + " ...")
     Isabelle_System.with_tmp_dir("mercurial")(tmp_dir =>
@@ -226,7 +226,7 @@
     repo: String = "",
     package_update: Boolean = false,
     mercurial_source: String = "",
-    progress: Progress = No_Progress)
+    progress: Progress = new Progress)
   {
     /* system environment */
 
@@ -599,7 +599,7 @@
     name: String = default_name,
     config_file: Option[Path] = None,
     test_user: String = "",
-    progress: Progress = No_Progress)
+    progress: Progress = new Progress)
   {
     Linux.check_system_root()
 
@@ -717,7 +717,7 @@
   def phabricator_setup_ssh(
     server_port: Int = default_server_port,
     system_port: Int = default_system_port,
-    progress: Progress = No_Progress)
+    progress: Progress = new Progress)
   {
     Linux.check_system_root()
 
--- a/src/Pure/Tools/print_operation.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/print_operation.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -38,6 +38,6 @@
       true
     }
 
-    val functions = List(Markup.PRINT_OPERATIONS -> put _)
+    val functions = List(Markup.PRINT_OPERATIONS -> put)
   }
 }
--- a/src/Pure/Tools/profiling_report.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/profiling_report.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -32,24 +32,22 @@
   val isabelle_tool =
     Isabelle_Tool("profiling_report", "report Poly/ML profiling information from log files", args =>
     {
-      Command_Line.tool0 {
-        val getopts =
-          Getopts("""
+      val getopts =
+        Getopts("""
 Usage: isabelle profiling_report [LOGS ...]
 
   Report Poly/ML profiling output from log files (potentially compressed).
 """)
-        val log_names = getopts(args)
-        for (name <- log_names) {
-          val log_file = Build_Log.Log_File(Path.explode(name))
-          val results =
-            for ((count, fun) <- profiling_report(log_file))
-              yield
-                String.format(Locale.ROOT, "%14d %s",
-                  count.asInstanceOf[AnyRef], fun.asInstanceOf[AnyRef])
-          if (results.nonEmpty)
-            Output.writeln(cat_lines((log_file.name + ":") :: results))
-        }
+      val log_names = getopts(args)
+      for (name <- log_names) {
+        val log_file = Build_Log.Log_File(Path.explode(name))
+        val results =
+          for ((count, fun) <- profiling_report(log_file))
+            yield
+              String.format(Locale.ROOT, "%14d %s",
+                count.asInstanceOf[AnyRef], fun.asInstanceOf[AnyRef])
+        if (results.nonEmpty)
+          Output.writeln(cat_lines((log_file.name + ":") :: results))
       }
     })
 }
--- a/src/Pure/Tools/rail.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/rail.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -191,7 +191,7 @@
   ANTIQUOTE of (bool * Antiquote.antiq) * Position.range;
 
 fun reports_of_tree ctxt =
-  if Context_Position.is_visible ctxt then
+  if Context_Position.reports_enabled ctxt then
     let
       fun reports r =
         if r = Position.no_range then []
@@ -309,7 +309,7 @@
     val toks = tokenize (Input.source_explode source);
     val _ = Context_Position.reports ctxt (maps reports_of_token toks);
     val rules = parse_rules (filter is_proper toks);
-    val _ = Position.reports (maps (reports_of_tree ctxt o #2) rules);
+    val _ = Context_Position.reports ctxt (maps (reports_of_tree ctxt o #2) rules);
   in map (apsnd prepare_tree) rules end;
 
 
--- a/src/Pure/Tools/server.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/server.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -35,8 +35,8 @@
 
     def split(msg: String): (String, String) =
     {
-      val name = msg.takeWhile(is_name_char(_))
-      val argument = msg.substring(name.length).dropWhile(Symbol.is_ascii_blank(_))
+      val name = msg.takeWhile(is_name_char)
+      val argument = msg.substring(name.length).dropWhile(Symbol.is_ascii_blank)
       (name, argument)
     }
 
@@ -173,12 +173,11 @@
     private val out = new BufferedOutputStream(socket.getOutputStream)
     private val out_lock: AnyRef = new Object
 
-    def tty_loop(interrupt: Option[() => Unit] = None): TTY_Loop =
+    def tty_loop(): TTY_Loop =
       new TTY_Loop(
         new OutputStreamWriter(out),
         new InputStreamReader(in),
-        writer_lock = out_lock,
-        interrupt = interrupt)
+        writer_lock = out_lock)
 
     def read_password(password: String): Boolean =
       try { Byte_Message.read_line(in).map(_.text) == Some(password) }
@@ -275,10 +274,6 @@
       context.notify(JSON.Object(Markup.KIND -> Markup.NODES_STATUS, Markup.NODES_STATUS -> json))
     }
 
-    @volatile private var is_stopped = false
-    override def stopped: Boolean = is_stopped
-    def stop { is_stopped = true }
-
     override def toString: String = context.toString
   }
 
@@ -292,7 +287,7 @@
     val progress: Connection_Progress = context.progress(ident)
     def cancel { progress.stop }
 
-    private lazy val thread = Standard_Thread.fork("server_task")
+    private lazy val thread = Isabelle_Thread.fork(name = "server_task")
     {
       Exn.capture { body(task) } match {
         case Exn.Res(res) =>
@@ -442,7 +437,7 @@
         find(db, name) match {
           case Some(server_info) =>
             using(server_info.connection())(_.write_message("shutdown"))
-            while(server_info.active) { Thread.sleep(50) }
+            while(server_info.active) { Time.seconds(0.05).sleep }
             true
           case None => false
         }
@@ -491,7 +486,7 @@
 
       if (operation_list) {
         for {
-          server_info <- using(SQLite.open_database(Data.database))(list(_))
+          server_info <- using(SQLite.open_database(Data.database))(list)
           if server_info.active
         } Output.writeln(server_info.toString, stdout = true)
       }
@@ -539,7 +534,7 @@
     for ((_, session) <- sessions) {
       try {
         val result = session.stop()
-        if (!result.ok) log("Session shutdown failed: return code " + result.rc)
+        if (!result.ok) log("Session shutdown failed: " + result.print_rc)
       }
       catch { case ERROR(msg) => log("Session shutdown failed: " + msg) }
     }
@@ -601,13 +596,13 @@
   }
 
   private lazy val server_thread: Thread =
-    Standard_Thread.fork("server") {
+    Isabelle_Thread.fork(name = "server") {
       var finished = false
       while (!finished) {
         Exn.capture(server_socket.accept) match {
           case Exn.Res(socket) =>
-            Standard_Thread.fork("server_connection")
-              { using(Server.Connection(socket))(handle(_)) }
+            Isabelle_Thread.fork(name = "server_connection")
+              { using(Server.Connection(socket))(handle) }
           case Exn.Exn(_) => finished = true
         }
       }
--- a/src/Pure/Tools/server_commands.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/server_commands.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -45,11 +45,11 @@
           include_sessions = include_sessions, verbose = verbose)
       }
 
-    def command(args: Args, progress: Progress = No_Progress)
-      : (JSON.Object.T, Build.Results, Sessions.Base_Info) =
+    def command(args: Args, progress: Progress = new Progress)
+      : (JSON.Object.T, Build.Results, Options, Sessions.Base_Info) =
     {
       val options = Options.init(prefs = args.preferences, opts = args.options)
-      val dirs = args.dirs.map(Path.explode(_))
+      val dirs = args.dirs.map(Path.explode)
 
       val base_info =
         Sessions.base_info(options, args.session, progress = progress, dirs = dirs,
@@ -85,8 +85,11 @@
                   "timing" -> result.timing.json)
               }))
 
-      if (results.ok) (results_json, results, base_info)
-      else throw new Server.Error("Session build failed: return code " + results.rc, results_json)
+      if (results.ok) (results_json, results, options, base_info)
+      else {
+        throw new Server.Error("Session build failed: " + Process_Result.print_rc(results.rc),
+          results_json)
+      }
     }
   }
 
@@ -103,14 +106,14 @@
       }
       yield Args(build = build, print_mode = print_mode)
 
-    def command(args: Args, progress: Progress = No_Progress, log: Logger = No_Logger)
+    def command(args: Args, progress: Progress = new Progress, log: Logger = No_Logger)
       : (JSON.Object.T, (UUID.T, Headless.Session)) =
     {
-      val base_info =
-        try { Session_Build.command(args.build, progress = progress)._3 }
+      val (_, _, options, base_info) =
+        try { Session_Build.command(args.build, progress = progress) }
         catch { case exn: Server.Error => error(exn.message) }
 
-      val resources = Headless.Resources(base_info, log = log)
+      val resources = Headless.Resources(options, base_info, log = log)
 
       val session = resources.start_session(print_mode = args.print_mode, progress = progress)
 
@@ -136,7 +139,7 @@
       val result_json = JSON.Object("ok" -> result.ok, "return_code" -> result.rc)
 
       if (result.ok) (result_json, result)
-      else throw new Server.Error("Session shutdown failed: return code " + result.rc, result_json)
+      else throw new Server.Error("Session shutdown failed: " + result.print_rc, result_json)
     }
   }
 
@@ -179,7 +182,7 @@
     def command(args: Args,
       session: Headless.Session,
       id: UUID.T = UUID.random(),
-      progress: Progress = No_Progress): (JSON.Object.T, Headless.Use_Theories_Result) =
+      progress: Progress = new Progress): (JSON.Object.T, Headless.Use_Theories_Result) =
     {
       val result =
         session.use_theories(args.theories, master_dir = args.master_dir,
--- a/src/Pure/Tools/simplifier_trace.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/simplifier_trace.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -123,22 +123,22 @@
   def handle_results(session: Session, id: Document_ID.Command, results: Command.Results): Context =
   {
     val slot = Future.promise[Context]
-    manager.send(Handle_Results(session, id, results, slot))
+    the_manager(session).send(Handle_Results(session, id, results, slot))
     slot.join
   }
 
-  def generate_trace(results: Command.Results): Trace =
+  def generate_trace(session: Session, results: Command.Results): Trace =
   {
     val slot = Future.promise[Trace]
-    manager.send(Generate_Trace(results, slot))
+    the_manager(session).send(Generate_Trace(results, slot))
     slot.join
   }
 
-  def clear_memory() =
-    manager.send(Clear_Memory)
+  def clear_memory(session: Session): Unit =
+    the_manager(session).send(Clear_Memory)
 
-  def send_reply(session: Session, serial: Long, answer: Answer) =
-    manager.send(Reply(session, serial, answer))
+  def send_reply(session: Session, serial: Long, answer: Answer): Unit =
+    the_manager(session).send(Reply(session, serial, answer))
 
   def make_manager: Consumer_Thread[Any] =
   {
@@ -283,10 +283,10 @@
     )
   }
 
-  private val manager_variable = Synchronized[Option[Consumer_Thread[Any]]](None)
+  private val managers = Synchronized(Map.empty[Session, Consumer_Thread[Any]])
 
-  def manager: Consumer_Thread[Any] =
-    manager_variable.value match {
+  def the_manager(session: Session): Consumer_Thread[Any] =
+    managers.value.get(session) match {
       case Some(thread) if thread.is_active => thread
       case _ => error("Bad Simplifier_Trace.manager thread")
     }
@@ -296,25 +296,36 @@
 
   class Handler extends Session.Protocol_Handler
   {
-    try { manager }
-    catch { case ERROR(_) => manager_variable.change(_ => Some(make_manager)) }
+    private var the_session: Session = null
+
+    override def init(session: Session)
+    {
+      try { the_manager(session) }
+      catch { case ERROR(_) => managers.change(map => map + (session -> make_manager)) }
+      the_session = session
+    }
 
-    override def exit() =
+    override def exit()
     {
-      manager.send(Clear_Memory)
-      manager.shutdown()
-      manager_variable.change(_ => None)
+      val session = the_session
+      if (session != null) {
+        val manager = the_manager(session)
+        manager.send(Clear_Memory)
+        manager.shutdown()
+        managers.change(map => map - session)
+        the_session = null
+      }
     }
 
     private def cancel(msg: Prover.Protocol_Output): Boolean =
       msg.properties match {
         case Markup.Simp_Trace_Cancel(serial) =>
-          manager.send(Cancel(serial))
+          the_manager(the_session).send(Cancel(serial))
           true
         case _ =>
           false
       }
 
-    val functions = List(Markup.SIMP_TRACE_CANCEL -> cancel _)
+    val functions = List(Markup.SIMP_TRACE_CANCEL -> cancel)
   }
 }
--- a/src/Pure/Tools/spell_checker.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/spell_checker.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -66,8 +66,8 @@
 
   class Dictionary private[Spell_Checker](val path: Path)
   {
-    val lang = path.drop_ext.file_name
-    val user_path = Path.explode("$ISABELLE_HOME_USER/dictionaries") + Path.basic(lang)
+    val lang: String = path.drop_ext.file_name
+    val user_path: Path = Path.explode("$ISABELLE_HOME_USER/dictionaries") + Path.basic(lang)
     override def toString: String = lang
   }
 
@@ -141,9 +141,9 @@
         permanent_updates
 
     val factory_class = Class.forName("com.inet.jortho.DictionaryFactory")
-    val factory_cons = factory_class.getConstructor()
-    factory_cons.setAccessible(true)
-    val factory = factory_cons.newInstance()
+    val factory_constructor = factory_class.getConstructor()
+    factory_constructor.setAccessible(true)
+    val factory = factory_constructor.newInstance()
 
     val add = Untyped.method(factory_class, "add", classOf[String])
 
--- a/src/Pure/Tools/update.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/update.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -10,7 +10,7 @@
 object Update
 {
   def update(options: Options, logic: String,
-    progress: Progress = No_Progress,
+    progress: Progress = new Progress,
     log: Logger = No_Logger,
     dirs: List[Path] = Nil,
     select_dirs: List[Path] = Nil,
--- a/src/Pure/Tools/update_cartouches.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/Tools/update_cartouches.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -7,13 +7,16 @@
 package isabelle
 
 
+import scala.util.matching.Regex
+
+
 object Update_Cartouches
 {
   /* update cartouches */
 
   private val Verbatim_Body = """(?s)[ \t]*(.*?)[ \t]*""".r
 
-  val Text_Antiq = """(?s)@\{\s*text\b\s*(.+)\}""".r
+  val Text_Antiq: Regex = """(?s)@\{\s*text\b\s*(.+)\}""".r
 
   def update_text(content: String): String =
   {
--- a/src/Pure/build-jars	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/build-jars	Wed Apr 15 14:31:21 2020 +0200
@@ -28,11 +28,12 @@
   src/Pure/Admin/other_isabelle.scala
   src/Pure/Concurrent/consumer_thread.scala
   src/Pure/Concurrent/counter.scala
+  src/Pure/Concurrent/delay.scala
   src/Pure/Concurrent/event_timer.scala
   src/Pure/Concurrent/future.scala
+  src/Pure/Concurrent/isabelle_thread.scala
   src/Pure/Concurrent/mailbox.scala
   src/Pure/Concurrent/par_list.scala
-  src/Pure/Concurrent/standard_thread.scala
   src/Pure/Concurrent/synchronized.scala
   src/Pure/GUI/color_value.scala
   src/Pure/GUI/gui.scala
--- a/src/Pure/context_position.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/context_position.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -17,45 +17,64 @@
   val restore_visible_generic: Context.generic -> Context.generic -> Context.generic
   val restore_visible: Proof.context -> Proof.context -> Proof.context
   val restore_visible_global: theory -> theory -> theory
+  val pide_reports: unit -> bool
+  val reports_enabled_generic: Context.generic -> bool
+  val reports_enabled: Proof.context -> bool
+  val reports_enabled_global: theory -> bool
   val is_reported_generic: Context.generic -> Position.T -> bool
   val is_reported: Proof.context -> Position.T -> bool
+  val is_reported_global: theory -> Position.T -> bool
   val report_generic: Context.generic -> Position.T -> Markup.T -> unit
   val reported_text: Proof.context -> Position.T -> Markup.T -> string -> string
   val report_text: Proof.context -> Position.T -> Markup.T -> string -> unit
   val report: Proof.context -> Position.T -> Markup.T -> unit
   val reports_text: Proof.context -> Position.report_text list -> unit
+  val reports_generic: Context.generic -> Position.report list -> unit
   val reports: Proof.context -> Position.report list -> unit
+  val reports_global: theory -> Position.report list -> unit
 end;
 
 structure Context_Position: CONTEXT_POSITION =
 struct
 
-structure Data = Generic_Data
-(
-  type T = bool option * bool option;  (*really, visible*)
-  val empty: T = (NONE, NONE);
-  val extend = I;
-  fun merge ((a, b), (a', b')) : T = (merge_options (a, a'), merge_options (b, b'));
-);
+(* visible context *)
+
+val really = Config.declare_bool ("really", Position.none) (K true);
+val visible = Config.declare_bool ("visible", Position.none) (K true);
 
-val is_visible_generic = the_default true o snd o Data.get;
-val is_visible = is_visible_generic o Context.Proof;
-val is_visible_global = is_visible_generic o Context.Theory;
+val is_visible_generic = Config.apply_generic visible;
+val is_visible = Config.apply visible;
+val is_visible_global = Config.apply_global visible;
+
+val set_visible_generic = Config.put_generic visible;
+val set_visible = Config.put visible;
+val set_visible_global = Config.put_global visible;
+
+val is_really = Config.apply really;
+fun is_really_visible ctxt = is_really ctxt andalso is_visible ctxt;
+val not_really = Config.put really false;
 
-val set_visible_generic = Data.map o apsnd o K o SOME;
-val set_visible = Context.proof_map o Data.map o apsnd o K o SOME;
-val set_visible_global = Context.theory_map o Data.map o apsnd o K o SOME;
+fun restore_visible_generic context =
+  Config.restore_generic really context #>
+  Config.restore_generic visible context;
+
+val restore_visible = Context.proof_map o restore_visible_generic o Context.Proof;
+val restore_visible_global = Context.theory_map o restore_visible_generic o Context.Theory;
+
+
+(* PIDE reports *)
 
-val is_really = the_default true o fst o Data.get o Context.Proof;
-fun is_really_visible ctxt = is_really ctxt andalso is_visible ctxt;
-val not_really = Context.proof_map (Data.map (apfst (K (SOME false))));
+fun pide_reports () = Options.default_bool "pide_reports";
+
+fun reports_enabled_generic context = pide_reports () andalso is_visible_generic context;
+val reports_enabled = reports_enabled_generic o Context.Proof;
+val reports_enabled_global = reports_enabled_generic o Context.Theory;
 
-val restore_visible_generic = Data.put o Data.get;
-val restore_visible = Context.proof_map o Data.put o Data.get o Context.Proof;
-val restore_visible_global = Context.theory_map o Data.put o Data.get o Context.Theory;
+fun is_reported_generic context pos =
+  reports_enabled_generic context andalso Position.is_reported pos;
 
-fun is_reported_generic context pos = is_visible_generic context andalso Position.is_reported pos;
-fun is_reported ctxt pos = is_visible ctxt andalso Position.is_reported pos;
+val is_reported = is_reported_generic o Context.Proof;
+val is_reported_global = is_reported_generic o Context.Theory;
 
 fun report_generic context pos markup =
   if is_reported_generic context pos then
@@ -68,7 +87,13 @@
 fun report_text ctxt pos markup txt = Output.report [reported_text ctxt pos markup txt];
 fun report ctxt pos markup = report_text ctxt pos markup "";
 
-fun reports_text ctxt reps = if is_visible ctxt then Position.reports_text reps else ();
-fun reports ctxt reps = if is_visible ctxt then Position.reports reps else ();
+fun reports_text ctxt reps =
+  if reports_enabled ctxt then Position.reports_text reps else ();
+
+fun reports_generic context reps =
+  if reports_enabled_generic context then Position.reports reps else ();
+
+val reports = reports_generic o Context.Proof;
+val reports_global = reports_generic o Context.Theory;
 
 end;
--- a/src/Pure/library.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/library.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -140,7 +140,7 @@
     else s
 
   def trim_split_lines(s: String): List[String] =
-    split_lines(trim_line(s)).map(trim_line(_))
+    split_lines(trim_line(s)).map(trim_line)
 
   def isolate_substring(s: String): String = new String(s.toCharArray)
 
@@ -204,7 +204,7 @@
   def is_regex_meta(c: Char): Boolean = """()[]{}\^$|?*+.<>-=!""".contains(c)
 
   def escape_regex(s: String): String =
-    if (s.exists(is_regex_meta(_))) {
+    if (s.exists(is_regex_meta)) {
       (for (c <- s.iterator)
        yield { if (is_regex_meta(c)) "\\" + c.toString else c.toString }).mkString
     }
--- a/src/Pure/skip_proof.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/skip_proof.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -18,7 +18,7 @@
 (* report *)
 
 fun report ctxt =
-  if Context_Position.is_visible ctxt then
+  if Context_Position.reports_enabled ctxt then
     Output.report [Markup.markup (Markup.bad ()) "Skipped proof"]
   else ();
 
--- a/src/Pure/term.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/term.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -50,7 +50,7 @@
     override def toString: String =
       "TVar(" + name + (if (sort.isEmpty) "" else "," + sort) + ")"
   }
-  val dummyT = Type("dummy")
+  val dummyT: Type = Type("dummy")
 
   sealed abstract class Term
   {
@@ -143,7 +143,7 @@
       lookup(x) getOrElse store(Indexname(cache_string(x.name), x.index))
 
     protected def cache_sort(x: Sort): Sort =
-      if (x.isEmpty) Nil else lookup(x) getOrElse store(x.map(cache_string(_)))
+      if (x.isEmpty) Nil else lookup(x) getOrElse store(x.map(cache_string))
 
     protected def cache_typ(x: Typ): Typ =
     {
@@ -166,7 +166,7 @@
       else {
         lookup(x) match {
           case Some(y) => y
-          case None => store(x.map(cache_typ(_)))
+          case None => store(x.map(cache_typ))
         }
       }
     }
--- a/src/Pure/theory.ML	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Pure/theory.ML	Wed Apr 15 14:31:21 2020 +0200
@@ -130,7 +130,7 @@
 fun init_markup (name, pos) thy =
   let
     val id = serial ();
-    val _ = Position.report pos (theory_markup true name id pos);
+    val _ = Context_Position.reports_global thy [(pos, theory_markup true name id pos)];
   in map_thy (fn (_, _, axioms, defs, wrappers) => (pos, id, axioms, defs, wrappers)) thy end;
 
 fun get_markup thy =
--- a/src/Tools/Graphview/graph_file.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/Graphview/graph_file.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -30,8 +30,8 @@
     }
 
     val name = file.getName
-    if (name.endsWith(".png")) Graphics_File.write_png(file, paint _, w, h)
-    else if (name.endsWith(".pdf")) Graphics_File.write_pdf(file, paint _, w, h)
+    if (name.endsWith(".png")) Graphics_File.write_png(file, paint, w, h)
+    else if (name.endsWith(".pdf")) Graphics_File.write_pdf(file, paint, w, h)
     else error("Bad type of file: " + quote(name) + " (.png or .pdf expected)")
   }
 
--- a/src/Tools/Graphview/graphview.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/Graphview/graphview.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -82,7 +82,7 @@
       }
       else node.toString
 
-    _layout = Layout.make(options, metrics, node_text _, visible_graph)
+    _layout = Layout.make(options, metrics, node_text, visible_graph)
   }
 
 
--- a/src/Tools/Graphview/layout.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/Graphview/layout.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -104,7 +104,7 @@
                 val lines = split_lines(node_text(a, content))
                 val w2 = metrics.node_width2(lines)
                 val h2 = metrics.node_height2(lines.length)
-                ((Node(a), Info(0.0, 0.0, w2, h2, lines)), bs.toList.map(Node(_)))
+                ((Node(a), Info(0.0, 0.0, w2, h2, lines)), bs.toList.map(Node))
             }).toList)
 
       val initial_levels: Levels =
--- a/src/Tools/Graphview/model.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/Graphview/model.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -18,7 +18,7 @@
   val events = new Mutator_Event.Bus
 
   private var _mutators : List[Mutator.Info] = Nil
-  def apply() = _mutators
+  def apply(): List[Mutator.Info] = _mutators
   def apply(mutators: List[Mutator.Info])
   {
     _mutators = mutators
--- a/src/Tools/Graphview/mutator.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/Graphview/mutator.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -173,6 +173,6 @@
 
 trait Filter extends Mutator
 {
-  def mutate(full_graph: Graph_Display.Graph, graph: Graph_Display.Graph) = filter(graph)
+  def mutate(full_graph: Graph_Display.Graph, graph: Graph_Display.Graph): Graph_Display.Graph = filter(graph)
   def filter(graph: Graph_Display.Graph): Graph_Display.Graph
 }
--- a/src/Tools/Graphview/mutator_dialog.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/Graphview/mutator_dialog.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -110,7 +110,7 @@
     filter_panel.repaint
   }
 
-  val filter_panel = new BoxPanel(Orientation.Vertical) {}
+  val filter_panel: BoxPanel = new BoxPanel(Orientation.Vertical) {}
 
   private val mutator_box = new ComboBox[Mutator](container.available)
 
--- a/src/Tools/Graphview/tree_panel.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/Graphview/tree_panel.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -112,7 +112,7 @@
   private val selection_field_foreground = selection_field.foreground
 
   private val selection_delay =
-    GUI_Thread.delay_last(graphview.options.seconds("editor_input_delay")) {
+    Delay.last(graphview.options.seconds("editor_input_delay"), gui = true) {
       val (pattern, ok) =
         selection_field.text match {
           case null | "" => (None, true)
--- a/src/Tools/Haskell/Haskell.thy	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/Haskell/Haskell.thy	Wed Apr 15 14:31:21 2020 +0200
@@ -1715,20 +1715,20 @@
         Just n -> fmap trim_line . fst <$> read_block socket n
 \<close>
 
-generate_file "Isabelle/Standard_Thread.hs" = \<open>
-{-  Title:      Isabelle/Standard_Thread.hs
+generate_file "Isabelle/Isabelle_Thread.hs" = \<open>
+{-  Title:      Isabelle/Isabelle_Thread.hs
     Author:     Makarius
     LICENSE:    BSD 3-clause (Isabelle)
 
-Standard thread operations.
-
-See \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/standard_thread.ML\<close>
-and \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/standard_thread.scala\<close>.
+Isabelle-specific thread management.
+
+See \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.ML\<close>
+and \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.scala\<close>.
 -}
 
 {-# LANGUAGE NamedFieldPuns #-}
 
-module Isabelle.Standard_Thread (
+module Isabelle.Isabelle_Thread (
   ThreadId, Result,
   find_id,
   properties, change_properties,
@@ -1912,7 +1912,7 @@
 import Isabelle.Library
 import qualified Isabelle.UUID as UUID
 import qualified Isabelle.Byte_Message as Byte_Message
-import qualified Isabelle.Standard_Thread as Standard_Thread
+import qualified Isabelle.Isabelle_Thread as Isabelle_Thread
 
 
 {- server address -}
@@ -1953,7 +1953,7 @@
     loop :: Socket -> ByteString -> IO ()
     loop server_socket password = forever $ do
       (connection, _) <- Socket.accept server_socket
-      Standard_Thread.fork_finally
+      Isabelle_Thread.fork_finally
         (do
           line <- Byte_Message.read_line connection
           when (line == Just password) $ handle connection)
--- a/src/Tools/VSCode/extension/README.md	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/VSCode/extension/README.md	Wed Apr 15 14:31:21 2020 +0200
@@ -1,15 +1,14 @@
 # Isabelle Prover IDE support
 
 This extension connects VSCode to the Isabelle Prover IDE infrastructure: it
-requires Isabelle2020.
+requires a repository version of Isabelle.
 
 The implementation is centered around the VSCode Language Server protocol, but
 with many add-ons that are specific to VSCode and Isabelle/PIDE.
 
 See also:
 
-  * <https://isabelle.in.tum.de/website-Isabelle2020>
-  * <https://isabelle.in.tum.de/repos/isabelle/file/Isabelle2020/src/Tools/VSCode>
+  * <https://isabelle.in.tum.de/repos/isabelle/file/tip/src/Tools/VSCode>
   * <https://github.com/Microsoft/language-server-protocol>
 
 
@@ -59,8 +58,8 @@
 
 ### Isabelle/VSCode Installation
 
-  * Download Isabelle2020 from <https://isabelle.in.tum.de/website-Isabelle2020>
-    or any of its mirror sites.
+  * Download a recent Isabelle development snapshot from
+  <https://isabelle.in.tum.de/devel/release_snapshot>
 
   * Unpack and run the main Isabelle/jEdit application as usual, to ensure that
   the logic image is built properly and Isabelle works as expected.
@@ -69,7 +68,7 @@
 
   * Open the VSCode *Extensions* view and install the following:
 
-      + *Isabelle2020* (needs to fit to the underlying Isabelle release).
+      + *Isabelle* (needs to fit to the underlying Isabelle release).
 
       + *Prettify Symbols Mode* (important for display of Isabelle symbols).
 
@@ -90,17 +89,17 @@
 
       + Linux:
         ```
-        "isabelle.home": "/home/makarius/Isabelle2020"
+        "isabelle.home": "/home/makarius/Isabelle"
         ```
 
       + Mac OS X:
         ```
-        "isabelle.home": "/Users/makarius/Isabelle.app/Isabelle2020"
+        "isabelle.home": "/Users/makarius/Isabelle.app/Isabelle"
         ```
 
       + Windows:
         ```
-        "isabelle.home": "C:\\Users\\makarius\\Isabelle2020"
+        "isabelle.home": "C:\\Users\\makarius\\Isabelle"
         ```
 
   * Restart the VSCode application to ensure that all extensions are properly
--- a/src/Tools/VSCode/extension/package.json	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/VSCode/extension/package.json	Wed Apr 15 14:31:21 2020 +0200
@@ -1,6 +1,6 @@
 {
-    "name": "Isabelle2020",
-    "displayName": "Isabelle2020",
+    "name": "Isabelle",
+    "displayName": "Isabelle",
     "description": "Isabelle Prover IDE",
     "keywords": [
         "theorem prover",
--- a/src/Tools/VSCode/src/build_vscode.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/VSCode/src/build_vscode.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -17,7 +17,7 @@
 
   /* grammar */
 
-  def build_grammar(options: Options, progress: Progress = No_Progress)
+  def build_grammar(options: Options, progress: Progress = new Progress)
   {
     val logic = Grammar.default_logic
     val keywords = Sessions.base_info(options, logic).check_base.overall_syntax.keywords
@@ -30,7 +30,7 @@
 
   /* extension */
 
-  def build_extension(progress: Progress = No_Progress, publish: Boolean = false)
+  def build_extension(progress: Progress = new Progress, publish: Boolean = false)
   {
     val output_path = extension_dir + Path.explode("out")
     progress.echo(output_path.implode)
--- a/src/Tools/VSCode/src/document_model.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/VSCode/src/document_model.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -60,7 +60,7 @@
 
   def init(session: Session, editor: Server.Editor, node_name: Document.Node.Name): Document_Model =
     Document_Model(session, editor, node_name, Content.empty,
-      node_required = session.resources.file_formats.is_theory(node_name))
+      node_required = File_Format.registry.is_theory(node_name))
 }
 
 sealed case class Document_Model(
--- a/src/Tools/VSCode/src/grammar.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/VSCode/src/grammar.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -23,7 +23,7 @@
   def generate(keywords: Keyword.Keywords): JSON.S =
   {
     val (minor_keywords, operators) =
-      keywords.minor.iterator.toList.partition(Symbol.is_ascii_identifier(_))
+      keywords.minor.iterator.toList.partition(Symbol.is_ascii_identifier)
 
     def major_keywords(pred: String => Boolean): List[String] =
       (for {
@@ -39,7 +39,7 @@
 
 
     def grouped_names(as: List[String]): String =
-      JSON.Format("\\b(" + as.sorted.map(Library.escape_regex(_)).mkString("|") + ")\\b")
+      JSON.Format("\\b(" + as.sorted.map(Library.escape_regex).mkString("|") + ")\\b")
 
     """{
   "name": "Isabelle",
--- a/src/Tools/VSCode/src/server.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/VSCode/src/server.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -127,12 +127,12 @@
 
   /* input from client or file-system */
 
-  private val delay_input: Standard_Thread.Delay =
-    Standard_Thread.delay_last(options.seconds("vscode_input_delay"), channel.Error_Logger)
+  private val delay_input: Delay =
+    Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger)
     { resources.flush_input(session, channel) }
 
-  private val delay_load: Standard_Thread.Delay =
-    Standard_Thread.delay_last(options.seconds("vscode_load_delay"), channel.Error_Logger) {
+  private val delay_load: Delay =
+    Delay.last(options.seconds("vscode_load_delay"), channel.Error_Logger) {
       val (invoke_input, invoke_load) =
         resources.resolve_dependencies(session, editor, file_watcher)
       if (invoke_input) delay_input.invoke()
@@ -140,7 +140,7 @@
     }
 
   private val file_watcher =
-    File_Watcher(sync_documents(_), options.seconds("vscode_load_delay"))
+    File_Watcher(sync_documents, options.seconds("vscode_load_delay"))
 
   private def close_document(file: JFile)
   {
@@ -183,8 +183,8 @@
 
   /* caret handling */
 
-  private val delay_caret_update: Standard_Thread.Delay =
-    Standard_Thread.delay_last(options.seconds("vscode_input_delay"), channel.Error_Logger)
+  private val delay_caret_update: Delay =
+    Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger)
     { session.caret_focus.post(Session.Caret_Focus) }
 
   private def update_caret(caret: Option[(JFile, Line.Position)])
@@ -199,8 +199,8 @@
 
   private lazy val preview_panel = new Preview_Panel(resources)
 
-  private lazy val delay_preview: Standard_Thread.Delay =
-    Standard_Thread.delay_last(options.seconds("vscode_output_delay"), channel.Error_Logger)
+  private lazy val delay_preview: Delay =
+    Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger)
     {
       if (preview_panel.flush(channel)) delay_preview.invoke()
     }
@@ -214,8 +214,8 @@
 
   /* output to client */
 
-  private val delay_output: Standard_Thread.Delay =
-    Standard_Thread.delay_last(options.seconds("vscode_output_delay"), channel.Error_Logger)
+  private val delay_output: Delay =
+    Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger)
     {
       if (resources.flush_output(channel)) delay_output.invoke()
     }
@@ -305,21 +305,12 @@
 
       dynamic_output.init()
 
-      var session_phase: Session.Consumer[Session.Phase] = null
-      session_phase =
-        Session.Consumer(getClass.getName) {
-          case Session.Ready =>
-            session.phase_changed -= session_phase
-            reply_ok("Welcome to Isabelle/" + base_info.session + " (" + Distribution.version + ")")
-          case Session.Terminated(result) if !result.ok =>
-            session.phase_changed -= session_phase
-            reply_error("Prover startup failed: return code " + result.rc)
-          case _ =>
-        }
-      session.phase_changed += session_phase
-
-      Isabelle_Process.start(session, options, modes = modes,
-        sessions_structure = Some(base_info.sessions_structure), logic = base_info.session)
+      try {
+        Isabelle_Process(session, options, base_info.sessions_structure, Sessions.store(options),
+          modes = modes, logic = base_info.session).await_startup
+        reply_ok("Welcome to Isabelle/" + base_info.session + " (" + Distribution.version + ")")
+      }
+      catch { case ERROR(msg) => reply_error(msg) }
     }
   }
 
@@ -342,7 +333,8 @@
         delay_preview.revoke()
 
         val result = session.stop()
-        if (result.ok) reply("") else reply("Prover shutdown failed: return code " + result.rc)
+        if (result.ok) reply("")
+        else reply("Prover shutdown failed: " + result.rc)
         None
       case None =>
         reply("Prover inactive")
@@ -485,7 +477,7 @@
       channel.read() match {
         case Some(json) =>
           json match {
-            case bulk: List[_] => bulk.foreach(handle(_))
+            case bulk: List[_] => bulk.foreach(handle)
             case _ => handle(json)
           }
           loop()
--- a/src/Tools/VSCode/src/vscode_resources.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/VSCode/src/vscode_resources.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -345,7 +345,7 @@
     output_text(XML.content(xml))
 
   def output_pretty(body: XML.Body, margin: Int): String =
-    output_text(Pretty.string_of(body, margin = margin))
+    output_text(Pretty.string_of(body, margin = margin, metric = Symbol.Metric))
   def output_pretty_tooltip(body: XML.Body): String = output_pretty(body, tooltip_margin)
   def output_pretty_message(body: XML.Body): String = output_pretty(body, message_margin)
 
--- a/src/Tools/jEdit/src/active.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/active.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -8,13 +8,22 @@
 
 
 import isabelle._
-
-import org.gjt.sp.jedit.View
-import org.gjt.sp.jedit.browser.VFSBrowser
+import org.gjt.sp.jedit.{ServiceManager, View}
 
 
 object Active
 {
+  abstract class Handler
+  {
+    def handle(
+      view: View, text: String, elem: XML.Elem,
+      doc_view: Document_View, snapshot: Document.Snapshot): Boolean
+  }
+
+  def handlers: List[Handler] =
+    ServiceManager.getServiceNames(classOf[Handler]).toList
+      .map(ServiceManager.getService(classOf[Handler], _))
+
   def action(view: View, text: String, elem: XML.Elem)
   {
     GUI_Thread.require {}
@@ -22,72 +31,80 @@
     Document_View.get(view.getTextArea) match {
       case Some(doc_view) =>
         doc_view.rich_text_area.robust_body() {
-          val text_area = doc_view.text_area
-          val model = doc_view.model
-          val buffer = model.buffer
-          val snapshot = model.snapshot()
-
+          val snapshot = doc_view.model.snapshot()
           if (!snapshot.is_outdated) {
-            // FIXME avoid hard-wired stuff
-            elem match {
-              case XML.Elem(Markup(Markup.BROWSER, _), body) =>
-                Standard_Thread.fork("browser") {
-                  val graph_file = Isabelle_System.tmp_file("graph")
-                  File.write(graph_file, XML.content(body))
-                  Isabelle_System.bash("isabelle browser -c " + File.bash_path(graph_file) + " &")
-                }
-
-              case XML.Elem(Markup(Markup.GRAPHVIEW, _), body) =>
-                Standard_Thread.fork("graphview") {
-                  val graph =
-                    Exn.capture { Graph_Display.decode_graph(body).transitive_reduction_acyclic }
-                  GUI_Thread.later { Graphview_Dockable(view, snapshot, graph) }
-                }
-
-              case XML.Elem(Markup(Markup.THEORY_EXPORTS, props), _) =>
-                GUI_Thread.later {
-                  val name = Markup.Name.unapply(props) getOrElse ""
-                  PIDE.editor.hyperlink_file(true, Isabelle_Export.vfs_prefix + name).follow(view)
-                }
-
-              case XML.Elem(Markup(Markup.JEDIT_ACTION, _), body) =>
-                GUI_Thread.later {
-                  view.getInputHandler.invokeAction(XML.content(body))
-                }
-
-              case XML.Elem(Markup(Markup.SIMP_TRACE_PANEL, props), _) =>
-                val link =
-                  props match {
-                    case Position.Id(id) => PIDE.editor.hyperlink_command(true, snapshot, id)
-                    case _ => None
-                  }
-                GUI_Thread.later {
-                  link.foreach(_.follow(view))
-                  view.getDockableWindowManager.showDockableWindow("isabelle-simplifier-trace")
-                }
-
-              case XML.Elem(Markup(Markup.SENDBACK, props), _) =>
-                if (buffer.isEditable) {
-                  props match {
-                    case Position.Id(id) =>
-                      Isabelle.edit_command(snapshot, text_area,
-                        props.contains(Markup.PADDING_COMMAND), id, text)
-                    case _ =>
-                      if (props.contains(Markup.PADDING_LINE))
-                        Isabelle.insert_line_padding(text_area, text)
-                      else text_area.setSelectedText(text)
-                  }
-                  text_area.requestFocus
-                }
-
-              case Protocol.Dialog(id, serial, result) =>
-                model.session.dialog_result(id, serial, result)
-
-              case _ =>
-            }
+            handlers.find(_.handle(view, text, elem, doc_view, snapshot))
           }
         }
       case None =>
     }
   }
+
+  class Misc_Handler extends Active.Handler
+  {
+    override def handle(
+      view: View, text: String, elem: XML.Elem,
+      doc_view: Document_View, snapshot: Document.Snapshot): Boolean =
+    {
+      val text_area = doc_view.text_area
+      val model = doc_view.model
+      val buffer = model.buffer
+
+      elem match {
+        case XML.Elem(Markup(Markup.BROWSER, _), body) =>
+          Isabelle_Thread.fork(name = "browser") {
+            val graph_file = Isabelle_System.tmp_file("graph")
+            File.write(graph_file, XML.content(body))
+            Isabelle_System.bash("isabelle browser -c " + File.bash_path(graph_file) + " &")
+          }
+          true
+
+        case XML.Elem(Markup(Markup.THEORY_EXPORTS, props), _) =>
+          GUI_Thread.later {
+            val name = Markup.Name.unapply(props) getOrElse ""
+            PIDE.editor.hyperlink_file(true, Isabelle_Export.vfs_prefix + name).follow(view)
+          }
+          true
+
+        case XML.Elem(Markup(Markup.JEDIT_ACTION, _), body) =>
+          GUI_Thread.later {
+            view.getInputHandler.invokeAction(XML.content(body))
+          }
+          true
+
+        case XML.Elem(Markup(Markup.SIMP_TRACE_PANEL, props), _) =>
+          val link =
+            props match {
+              case Position.Id(id) => PIDE.editor.hyperlink_command(true, snapshot, id)
+              case _ => None
+            }
+          GUI_Thread.later {
+            link.foreach(_.follow(view))
+            view.getDockableWindowManager.showDockableWindow("isabelle-simplifier-trace")
+          }
+          true
+
+        case XML.Elem(Markup(Markup.SENDBACK, props), _) =>
+          if (buffer.isEditable) {
+            props match {
+              case Position.Id(id) =>
+                Isabelle.edit_command(snapshot, text_area,
+                  props.contains(Markup.PADDING_COMMAND), id, text)
+              case _ =>
+                if (props.contains(Markup.PADDING_LINE))
+                  Isabelle.insert_line_padding(text_area, text)
+                else text_area.setSelectedText(text)
+            }
+            text_area.requestFocus
+          }
+          true
+
+        case Protocol.Dialog(id, serial, result) =>
+          model.session.dialog_result(id, serial, result)
+          true
+
+        case _ => false
+      }
+    }
+  }
 }
--- a/src/Tools/jEdit/src/completion_popup.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/completion_popup.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -10,7 +10,7 @@
 import isabelle._
 
 import java.awt.{Color, Font, Point, BorderLayout, Dimension}
-import java.awt.event.{KeyEvent, MouseEvent, MouseAdapter, FocusAdapter, FocusEvent}
+import java.awt.event.{KeyEvent, KeyListener, MouseEvent, MouseAdapter, FocusAdapter, FocusEvent}
 import java.io.{File => JFile}
 import javax.swing.{JPanel, JComponent, JLayeredPane, SwingUtilities}
 import javax.swing.border.LineBorder
@@ -18,7 +18,6 @@
 
 import scala.swing.{ListView, ScrollPane}
 import scala.swing.event.MouseClicked
-
 import org.gjt.sp.jedit.View
 import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, Selection}
 import org.gjt.sp.jedit.gui.{HistoryTextField, KeyEventWorkaround}
@@ -374,7 +373,7 @@
     }
 
     private val input_delay =
-      GUI_Thread.delay_last(PIDE.options.seconds("jedit_completion_delay")) {
+      Delay.last(PIDE.options.seconds("jedit_completion_delay"), gui = true) {
         action()
       }
 
@@ -399,7 +398,7 @@
     /* activation */
 
     private val outer_key_listener =
-      JEdit_Lib.key_listener(key_typed = input _)
+      JEdit_Lib.key_listener(key_typed = input)
 
     private def activate()
     {
@@ -531,7 +530,7 @@
     }
 
     private val process_delay =
-      GUI_Thread.delay_last(PIDE.options.seconds("jedit_completion_delay")) {
+      Delay.last(PIDE.options.seconds("jedit_completion_delay"), gui = true) {
         action()
       }
 
@@ -570,7 +569,7 @@
   GUI_Thread.require {}
 
   require(items.nonEmpty)
-  val multi = items.length > 1
+  val multi: Boolean = items.length > 1
 
 
   /* actions */
@@ -621,7 +620,7 @@
 
   /* event handling */
 
-  val inner_key_listener =
+  val inner_key_listener: KeyListener =
     JEdit_Lib.key_listener(
       key_pressed = (e: KeyEvent) =>
         {
@@ -655,7 +654,7 @@
           }
           propagate(e)
         },
-      key_typed = propagate _
+      key_typed = propagate
     )
 
   list_view.peer.addKeyListener(inner_key_listener)
--- a/src/Tools/jEdit/src/debugger_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/debugger_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -74,7 +74,7 @@
 
   val pretty_text_area = new Pretty_Text_Area(view)
 
-  override def detach_operation = pretty_text_area.detach_operation
+  override def detach_operation: Option[() => Unit] = pretty_text_area.detach_operation
 
   private def handle_resize()
   {
@@ -208,22 +208,22 @@
 
   private val continue_button = new Button("Continue") {
     tooltip = "Continue program on current thread, until next breakpoint"
-    reactions += { case ButtonClicked(_) => thread_selection().map(debugger.continue(_)) }
+    reactions += { case ButtonClicked(_) => thread_selection().map(debugger.continue) }
   }
 
   private val step_button = new Button("Step") {
     tooltip = "Single-step in depth-first order"
-    reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step(_)) }
+    reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step) }
   }
 
   private val step_over_button = new Button("Step over") {
     tooltip = "Single-step within this function"
-    reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step_over(_)) }
+    reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step_over) }
   }
 
   private val step_out_button = new Button("Step out") {
     tooltip = "Single-step outside this function"
-    reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step_out(_)) }
+    reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step_out) }
   }
 
   private val context_label = new Label("Context:") {
@@ -318,7 +318,7 @@
 
   /* main panel */
 
-  val main_panel = new SplitPane(Orientation.Vertical) {
+  val main_panel: SplitPane = new SplitPane(Orientation.Vertical) {
     oneTouchExpandable = true
     leftComponent = tree_pane
     rightComponent = Component.wrap(pretty_text_area)
@@ -362,7 +362,7 @@
   /* resize */
 
   private val delay_resize =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() }
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
--- a/src/Tools/jEdit/src/document_model.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/document_model.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -81,7 +81,7 @@
   private val state = Synchronized(State())  // owned by GUI thread
 
   def get_models(): Map[Document.Node.Name, Document_Model] = state.value.models
-  def get(name: Document.Node.Name): Option[Document_Model] = get_models.get(name)
+  def get(name: Document.Node.Name): Option[Document_Model] = get_models().get(name)
   def get(buffer: JEditBuffer): Option[Buffer_Model] = state.value.buffer_models.get(buffer)
 
   def document_blobs(): Document.Blobs = state.value.document_blobs
@@ -313,7 +313,7 @@
     val preview =
       HTTP.get(preview_root, arg =>
         for {
-          query <- Library.try_unprefix(preview_root + "?", arg.uri.toString).map(Url.decode(_))
+          query <- Library.try_unprefix(preview_root + "?", arg.uri.toString).map(Url.decode)
           name = Library.perhaps_unprefix(plain_text_prefix, query)
           model <- get(PIDE.resources.node_name(name))
         }
@@ -348,7 +348,7 @@
         if (hidden) Text.Perspective.empty
         else {
           val view_ranges = document_view_ranges(snapshot)
-          val load_ranges = snapshot.commands_loading_ranges(PIDE.editor.visible_node(_))
+          val load_ranges = snapshot.commands_loading_ranges(PIDE.editor.visible_node)
           Text.Perspective(view_ranges ::: load_ranges)
         }
       val overlays = PIDE.editor.node_overlays(node_name)
@@ -365,7 +365,7 @@
   {
     val snapshot = this.snapshot()
     if (snapshot.is_outdated) {
-      Thread.sleep(PIDE.options.seconds("editor_output_delay").ms)
+      PIDE.options.seconds("editor_output_delay").sleep
       await_stable_snapshot()
     }
     else snapshot
@@ -389,7 +389,7 @@
     file.foreach(PIDE.plugin.file_watcher.register_parent(_))
 
     val content = Document_Model.File_Content(text)
-    val node_required1 = node_required || session.resources.file_formats.is_theory(node_name)
+    val node_required1 = node_required || File_Format.registry.is_theory(node_name)
     File_Model(session, node_name, file, content, node_required1, last_perspective, pending_edits)
   }
 }
@@ -454,7 +454,7 @@
 
   def purge_edits(doc_blobs: Document.Blobs): Option[List[Document.Edit_Text]] =
     if (pending_edits.nonEmpty ||
-        !session.resources.file_formats.is_theory(node_name) &&
+        !File_Format.registry.is_theory(node_name) &&
           (node_required || !Document.Node.is_no_perspective_text(last_perspective))) None
     else {
       val text_edits = List(Text.Edit.remove(0, content.text))
@@ -589,7 +589,7 @@
         if (reparse || edits.nonEmpty || last_perspective != perspective) {
           pending.clear
           last_perspective = perspective
-          node_edits(node_header, edits, perspective)
+          node_edits(node_header(), edits, perspective)
         }
         else Nil
       }
--- a/src/Tools/jEdit/src/document_view.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/document_view.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -137,7 +137,7 @@
         JEdit_Lib.buffer_lock(buffer) {
           val rendering = get_rendering()
 
-          for (i <- 0 until physical_lines.length) {
+          for (i <- physical_lines.indices) {
             if (physical_lines(i) != -1) {
               val line_range = Text.Range(start(i), end(i))
 
@@ -184,7 +184,7 @@
   /* caret handling */
 
   private val delay_caret_update =
-    GUI_Thread.delay_last(PIDE.options.seconds("editor_input_delay")) {
+    Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) {
       session.caret_focus.post(Session.Caret_Focus)
       JEdit_Lib.invalidate(text_area)
     }
@@ -241,7 +241,7 @@
     text_area.revalidate()
     text_area.repaint()
     Isabelle.structure_matchers(JEdit_Lib.buffer_mode(text_area.getBuffer)).
-      foreach(text_area.addStructureMatcher(_))
+      foreach(text_area.addStructureMatcher)
     session.raw_edits += main
     session.commands_changed += main
   }
@@ -253,7 +253,7 @@
     session.raw_edits -= main
     session.commands_changed -= main
     Isabelle.structure_matchers(JEdit_Lib.buffer_mode(text_area.getBuffer)).
-      foreach(text_area.removeStructureMatcher(_))
+      foreach(text_area.removeStructureMatcher)
     text_overview.foreach(_.revoke())
     text_overview.foreach(text_area.removeLeftOfScrollBar(_))
     text_area.removeCaretListener(caret_listener)
--- a/src/Tools/jEdit/src/font_info.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/font_info.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -56,7 +56,7 @@
 
     // owned by GUI thread
     private var steps = 0
-    private val delay = GUI_Thread.delay_last(PIDE.options.seconds("editor_input_delay"))
+    private val delay = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true)
     {
       change_size(size =>
         {
--- a/src/Tools/jEdit/src/graphview_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/graphview_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -39,14 +39,31 @@
   private def reset_implicit(): Unit =
     set_implicit(Document.Snapshot.init, no_graph)
 
-  def apply(view: View, snapshot: Document.Snapshot, graph: Exn.Result[Graph_Display.Graph])
+  class Handler extends Active.Handler
   {
-    set_implicit(snapshot, graph)
-    view.getDockableWindowManager.floatDockableWindow("isabelle-graphview")
+    override def handle(
+      view: View, text: String, elem: XML.Elem,
+      doc_view: Document_View, snapshot: Document.Snapshot): Boolean =
+    {
+      elem match {
+        case XML.Elem(Markup(Markup.GRAPHVIEW, _), body) =>
+          Isabelle_Thread.fork(name = "graphview") {
+            val graph =
+              Exn.capture {
+                Graph_Display.decode_graph(body).transitive_reduction_acyclic
+              }
+            GUI_Thread.later {
+              set_implicit(snapshot, graph)
+              view.getDockableWindowManager.floatDockableWindow("isabelle-graphview")
+            }
+          }
+          true
+        case _ => false
+      }
+    }
   }
 }
 
-
 class Graphview_Dockable(view: View, position: String) extends Dockable(view, position)
 {
   private val snapshot = Graphview_Dockable.implicit_snapshot
--- a/src/Tools/jEdit/src/info_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/info_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -80,7 +80,7 @@
   /* resize */
 
   private val delay_resize =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() }
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
--- a/src/Tools/jEdit/src/isabelle.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/isabelle.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -406,8 +406,8 @@
 
   private def enclose_input(text_area: JEditTextArea, s1: String, s2: String)
   {
-    s1.foreach(text_area.userInput(_))
-    s2.foreach(text_area.userInput(_))
+    s1.foreach(text_area.userInput)
+    s2.foreach(text_area.userInput)
     s2.foreach(_ => text_area.goToPrevCharacter(false))
   }
 
--- a/src/Tools/jEdit/src/isabelle_options.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/isabelle_options.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -39,7 +39,7 @@
 
 class Isabelle_Options1 extends Isabelle_Options("isabelle-general")
 {
-  val options = PIDE.options
+  val options: JEdit_Options = PIDE.options
 
   private val predefined =
     List(JEdit_Sessions.logic_selector(options, false),
--- a/src/Tools/jEdit/src/isabelle_sidekick.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/isabelle_sidekick.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -53,11 +53,11 @@
     }
     override def getLongString: String = _name
     override def getName: String = _name
-    override def setName(name: String) = _name = name
+    override def setName(name: String): Unit = _name = name
     override def getStart: Position = _start
-    override def setStart(start: Position) = _start = start
+    override def setStart(start: Position): Unit = _start = start
     override def getEnd: Position = _end
-    override def setEnd(end: Position) = _end = end
+    override def setEnd(end: Position): Unit = _end = end
     override def toString: String = _name
   }
 
@@ -148,19 +148,19 @@
 
 class Isabelle_Sidekick_Default extends
   Isabelle_Sidekick_Structure("isabelle",
-    PIDE.resources.theory_node_name, Document_Structure.parse_sections _)
+    PIDE.resources.theory_node_name, Document_Structure.parse_sections)
 
 class Isabelle_Sidekick_Context extends
   Isabelle_Sidekick_Structure("isabelle-context",
-    PIDE.resources.theory_node_name, Document_Structure.parse_blocks _)
+    PIDE.resources.theory_node_name, Document_Structure.parse_blocks)
 
 class Isabelle_Sidekick_Options extends
   Isabelle_Sidekick_Structure("isabelle-options",
-    _ => Some(Document.Node.Name("options")), Document_Structure.parse_sections _)
+    _ => Some(Document.Node.Name("options")), Document_Structure.parse_sections)
 
 class Isabelle_Sidekick_Root extends
   Isabelle_Sidekick_Structure("isabelle-root",
-    _ => Some(Document.Node.Name("ROOT")), Document_Structure.parse_sections _)
+    _ => Some(Document.Node.Name("ROOT")), Document_Structure.parse_sections)
 
 class Isabelle_Sidekick_ML extends
   Isabelle_Sidekick_Structure("isabelle-ml",
@@ -257,7 +257,7 @@
         case _ =>
       }
       offset += line.length + 1
-      if (!line.forall(Character.isSpaceChar(_))) end_offset = offset
+      if (!line.forall(Character.isSpaceChar)) end_offset = offset
     }
     if (!stopped) { close2; close1 }
 
--- a/src/Tools/jEdit/src/isabelle_vfs.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/isabelle_vfs.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -77,7 +77,7 @@
     else {
       val files = _listFiles(vfs_session, parent, component)
       if (files == null) null
-      else files.find(_.getPath == url) getOrElse null
+      else files.find(_.getPath == url).orNull
     }
   }
 }
--- a/src/Tools/jEdit/src/jedit_editor.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/jedit_editor.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -32,10 +32,10 @@
   def purge() { flush_edits(purge = true) }
 
   private val delay1_flush =
-    GUI_Thread.delay_last(PIDE.options.seconds("editor_input_delay")) { flush() }
+    Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) { flush() }
 
   private val delay2_flush =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_generated_input_delay")) { flush() }
+    Delay.first(PIDE.options.seconds("editor_generated_input_delay"), gui = true) { flush() }
 
   def invoke(): Unit = delay1_flush.invoke()
   def invoke_generated(): Unit = { delay1_flush.invoke(); delay2_flush.invoke() }
@@ -181,7 +181,7 @@
     if (path.is_file)
       goto_file(true, view, File.platform_path(path))
     else {
-      Standard_Thread.fork("documentation") {
+      Isabelle_Thread.fork(name = "documentation") {
         try { Doc.view(path) }
         catch {
           case exn: Throwable =>
@@ -202,7 +202,7 @@
       case doc: Doc.Text_File if doc.name == name => doc.path
       case doc: Doc.Doc if doc.name == name => doc.path}).map(path =>
         new Hyperlink {
-          override val external = !path.is_file
+          override val external: Boolean = !path.is_file
           def follow(view: View): Unit = goto_doc(view, path)
           override def toString: String = "doc " + quote(name)
         })
@@ -211,7 +211,7 @@
     new Hyperlink {
       override val external = true
       def follow(view: View): Unit =
-        Standard_Thread.fork("hyperlink_url") {
+        Isabelle_Thread.fork(name = "hyperlink_url") {
           try { Isabelle_System.open(Url.escape_name(name)) }
           catch {
             case exn: Throwable =>
--- a/src/Tools/jEdit/src/jedit_lib.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/jedit_lib.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -160,7 +160,7 @@
     else view.getEditPanes().iterator.filter(_ != null).map(_.getTextArea).filter(_ != null)
 
   def jedit_text_areas(): Iterator[JEditTextArea] =
-    jedit_views().flatMap(jedit_text_areas(_))
+    jedit_views().flatMap(jedit_text_areas)
 
   def jedit_text_areas(buffer: JEditBuffer): Iterator[JEditTextArea] =
     jedit_text_areas().filter(_.getBuffer == buffer)
@@ -319,8 +319,8 @@
       def string_width(s: String): Double =
         painter.getFont.getStringBounds(s, painter.getFontRenderContext).getWidth
 
-      val unit = string_width(Symbol.space) max 1.0
-      val average = string_width("mix") / (3 * unit)
+      val unit: Double = string_width(Symbol.space) max 1.0
+      val average: Double = string_width("mix") / (3 * unit)
       def apply(s: String): Double = if (s == "\n") 1.0 else string_width(s) / unit
     }
 
--- a/src/Tools/jEdit/src/jedit_options.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/jedit_options.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -129,7 +129,7 @@
         case None => if (filter(opt.name)) List(make_component(opt)) else Nil
       }
     value.sections.sortBy(_._1).map(
-        { case (a, opts) => (a, opts.sortBy(_.title("jedit")).flatMap(mk_component _)) })
+        { case (a, opts) => (a, opts.sortBy(_.title("jedit")).flatMap(mk_component)) })
       .filterNot(_._2.isEmpty)
   }
 }
--- a/src/Tools/jEdit/src/jedit_rendering.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/jedit_rendering.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -300,11 +300,11 @@
   def tooltip(range: Text.Range, control: Boolean): Option[Text.Info[XML.Body]] =
   {
     val elements = if (control) Rendering.tooltip_elements else Rendering.tooltip_message_elements
-    tooltips(elements, range).map(info => info.map(Pretty.fbreaks(_)))
+    tooltips(elements, range).map(info => info.map(Pretty.fbreaks))
   }
 
-  lazy val tooltip_close_icon = JEdit_Lib.load_icon(options.string("tooltip_close_icon"))
-  lazy val tooltip_detach_icon = JEdit_Lib.load_icon(options.string("tooltip_detach_icon"))
+  lazy val tooltip_close_icon: Icon = JEdit_Lib.load_icon(options.string("tooltip_close_icon"))
+  lazy val tooltip_detach_icon: Icon = JEdit_Lib.load_icon(options.string("tooltip_detach_icon"))
 
 
   /* gutter */
@@ -378,7 +378,7 @@
     else
       snapshot.cumulate(range, current_color, Rendering.text_color_elements, _ =>
         {
-          case (_, Text.Info(_, elem)) => Rendering.text_color.get(elem.name).map(color(_))
+          case (_, Text.Info(_, elem)) => Rendering.text_color.get(elem.name).map(color)
         })
   }
 
--- a/src/Tools/jEdit/src/jedit_resources.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/jedit_resources.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -37,7 +37,7 @@
   /* document node name */
 
   def node_name(path: String): Document.Node.Name =
-    JEdit_Lib.check_file(path).flatMap(find_theory(_)) getOrElse {
+    JEdit_Lib.check_file(path).flatMap(find_theory) getOrElse {
       val vfs = VFSManager.getVFSForPath(path)
       val node = if (vfs.isInstanceOf[FileVFS]) MiscUtilities.resolveSymlinks(path) else path
       val theory = theory_name(Sessions.DRAFT, Thy_Header.theory_name(node))
--- a/src/Tools/jEdit/src/jedit_sessions.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/jedit_sessions.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -126,7 +126,7 @@
       session_requirements = logic_requirements)
 
   def session_build(
-    options: Options, progress: Progress = No_Progress, no_build: Boolean = false): Int =
+    options: Options, progress: Progress = new Progress, no_build: Boolean = false): Int =
   {
     Build.build(session_options(options), progress = progress, build_heap = true,
       no_build = no_build, dirs = session_dirs(), infos = PIDE.resources.session_base_info.infos,
@@ -135,15 +135,17 @@
 
   def session_start(options0: Options)
   {
+    val session = PIDE.session
     val options = session_options(options0)
+    val sessions_structure = PIDE.resources.session_base_info.sessions_structure
+    val store = Sessions.store(options)
 
-    Isabelle_Process.start(PIDE.session, options,
-      sessions_structure = Some(PIDE.resources.session_base_info.sessions_structure),
+    session.phase_changed += PIDE.plugin.session_phase_changed
+
+    Isabelle_Process(session, options, sessions_structure, store,
       logic = PIDE.resources.session_name,
-      store = Some(Sessions.store(options)),
       modes =
         (space_explode(',', options.string("jedit_print_mode")) :::
-         space_explode(',', Isabelle_System.getenv("JEDIT_PRINT_MODE"))).reverse,
-      phase_changed = PIDE.plugin.session_phase_changed)
+         space_explode(',', Isabelle_System.getenv("JEDIT_PRINT_MODE"))).reverse)
   }
 }
--- a/src/Tools/jEdit/src/monitor_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/monitor_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -59,10 +59,10 @@
     }
 
   private val input_delay =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_input_delay")) { update_chart }
+    Delay.first(PIDE.options.seconds("editor_input_delay"), gui = true) { update_chart }
 
   private val update_delay =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_chart_delay")) { update_chart }
+    Delay.first(PIDE.options.seconds("editor_chart_delay"), gui = true) { update_chart }
 
 
   /* controls */
@@ -108,7 +108,7 @@
 
   private val main =
     Session.Consumer[Any](getClass.getName) {
-      case stats: Session.Statistics =>
+      case stats: Session.Runtime_Statistics =>
         add_statistics(stats.props)
         update_delay.invoke()
 
@@ -118,13 +118,13 @@
 
   override def init()
   {
-    PIDE.session.statistics += main
+    PIDE.session.runtime_statistics += main
     PIDE.session.global_options += main
   }
 
   override def exit()
   {
-    PIDE.session.statistics -= main
+    PIDE.session.runtime_statistics -= main
     PIDE.session.global_options -= main
   }
 }
--- a/src/Tools/jEdit/src/output_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/output_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -32,7 +32,7 @@
   val pretty_text_area = new Pretty_Text_Area(view)
   set_content(pretty_text_area)
 
-  override def detach_operation = pretty_text_area.detach_operation
+  override def detach_operation: Option[() => Unit] = pretty_text_area.detach_operation
 
 
   private def handle_resize()
@@ -58,7 +58,7 @@
         }
 
       val new_output =
-        if (!restriction.isDefined || restriction.get.contains(command))
+        if (restriction.isEmpty || restriction.get.contains(command))
           Rendering.output_messages(results)
         else current_output
 
@@ -151,7 +151,7 @@
   /* resize */
 
   private val delay_resize =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() }
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
--- a/src/Tools/jEdit/src/plugin.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/plugin.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -107,7 +107,7 @@
   /* theory files */
 
   lazy val delay_init =
-    GUI_Thread.delay_last(options.seconds("editor_load_delay"))
+    Delay.last(options.seconds("editor_load_delay"), gui = true)
     {
       init_models()
     }
@@ -150,11 +150,11 @@
             }
             else Nil
 
-          (thy_files1 ::: thy_files2 ::: aux_files).filterNot(models.isDefinedAt(_))
+          (thy_files1 ::: thy_files2 ::: aux_files).filterNot(models.isDefinedAt)
         }
         if (required_files.nonEmpty) {
           try {
-            Standard_Thread.fork("resolve_dependencies") {
+            Isabelle_Thread.fork(name = "resolve_dependencies") {
               val loaded_files =
                 for {
                   name <- required_files
@@ -178,18 +178,18 @@
   }
 
   private lazy val delay_load =
-    GUI_Thread.delay_last(options.seconds("editor_load_delay")) { delay_load_action() }
+    Delay.last(options.seconds("editor_load_delay"), gui = true) { delay_load_action() }
 
   private def file_watcher_action(changed: Set[JFile]): Unit =
     if (Document_Model.sync_files(changed)) PIDE.editor.invoke_generated()
 
   lazy val file_watcher: File_Watcher =
-    File_Watcher(file_watcher_action _, options.seconds("editor_load_delay"))
+    File_Watcher(file_watcher_action, options.seconds("editor_load_delay"))
 
 
   /* session phase */
 
-  val session_phase_changed: Session.Phase => Unit =
+  val session_phase_changed: Session.Consumer[Session.Phase] = Session.Consumer("Isabelle/jEdit")
   {
     case Session.Terminated(result) if !result.ok =>
       GUI_Thread.later {
@@ -289,10 +289,10 @@
   @volatile private var startup_failure: Option[Throwable] = None
   @volatile private var startup_notified = false
 
-  private def init_view(view: View)
+  private def init_editor(view: View)
   {
+    Keymap_Merge.check_dialog(view)
     Session_Build.check_dialog(view)
-    Keymap_Merge.check_dialog(view)
   }
 
   private def init_title(view: View)
@@ -341,7 +341,7 @@
           jEdit.propertiesChanged()
 
           val view = jEdit.getActiveView()
-          init_view(view)
+          init_editor(view)
 
           PIDE.editor.hyperlink_position(true, Document.Snapshot.init,
             JEdit_Sessions.logic_root(options.value)).foreach(_.follow(view))
@@ -460,11 +460,11 @@
       completion_history.load()
       spell_checker.update(options.value)
 
-      JEdit_Lib.jedit_views.foreach(init_title(_))
+      JEdit_Lib.jedit_views.foreach(init_title)
 
       isabelle.jedit_base.Syntax_Style.set_style_extender(Syntax_Style.Extender)
       init_mode_provider()
-      JEdit_Lib.jedit_text_areas.foreach(Completion_Popup.Text_Area.init _)
+      JEdit_Lib.jedit_text_areas.foreach(Completion_Popup.Text_Area.init)
 
       http_server.start
 
@@ -480,7 +480,7 @@
     shutting_down.change(_ => false)
 
     val view = jEdit.getActiveView()
-    if (view != null) init_view(view)
+    if (view != null) init_editor(view)
   }
 
   override def stop()
@@ -489,7 +489,7 @@
 
     isabelle.jedit_base.Syntax_Style.dummy_style_extender()
     exit_mode_provider()
-    JEdit_Lib.jedit_text_areas.foreach(Completion_Popup.Text_Area.exit _)
+    JEdit_Lib.jedit_text_areas.foreach(Completion_Popup.Text_Area.exit)
 
     if (startup_failure.isEmpty) {
       options.value.save_prefs()
--- a/src/Tools/jEdit/src/pretty_text_area.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/pretty_text_area.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -191,7 +191,7 @@
     Component.wrap(new Completion_Popup.History_Text_Field("isabelle-search")
       {
         private val input_delay =
-          GUI_Thread.delay_last(PIDE.options.seconds("editor_input_delay")) {
+          Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) {
             search_action(this)
           }
         getDocument.addDocumentListener(new DocumentListener {
--- a/src/Tools/jEdit/src/pretty_tooltip.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/pretty_tooltip.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -77,7 +77,7 @@
   private var active = true
 
   private val pending_delay =
-    GUI_Thread.delay_last(PIDE.options.seconds("jedit_tooltip_delay")) {
+    Delay.last(PIDE.options.seconds("jedit_tooltip_delay"), gui = true) {
       pending match {
         case Some(body) => pending = None; body()
         case None =>
@@ -99,7 +99,7 @@
     }
 
   private lazy val reactivate_delay =
-    GUI_Thread.delay_last(PIDE.options.seconds("jedit_tooltip_delay")) {
+    Delay.last(PIDE.options.seconds("jedit_tooltip_delay"), gui = true) {
       active = true
     }
 
@@ -113,7 +113,7 @@
 
   /* dismiss */
 
-  private lazy val focus_delay = GUI_Thread.delay_last(PIDE.options.seconds("editor_input_delay"))
+  private lazy val focus_delay = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true)
   {
     dismiss_unfocused()
   }
@@ -146,7 +146,7 @@
   }
 
   def dismiss_descendant(parent: JComponent): Unit =
-    descendant(parent).foreach(dismiss(_))
+    descendant(parent).foreach(dismiss)
 
   def dismissed_all(): Boolean =
   {
@@ -203,7 +203,7 @@
 
   /* text area */
 
-  val pretty_text_area =
+  val pretty_text_area: Pretty_Text_Area =
     new Pretty_Text_Area(view, () => Pretty_Tooltip.dismiss(tip), true) {
       override def get_background() = Some(rendering.tooltip_color)
     }
@@ -262,7 +262,7 @@
       val formatted = Pretty.formatted(info.info, margin = margin, metric = metric)
       val lines =
         XML.traverse_text(formatted)(if (XML.text_length(formatted) == 0) 0 else 1)(
-          (n: Int, s: String) => n + s.iterator.filter(_ == '\n').length)
+          (n: Int, s: String) => n + s.iterator.count(_ == '\n'))
 
       val h = painter.getLineHeight * lines + geometry.deco_height
       val margin1 =
--- a/src/Tools/jEdit/src/query_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/query_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -302,7 +302,7 @@
   select_operation()
   set_content(operations_pane)
 
-  override def detach_operation =
+  override def detach_operation: Option[() => Unit] =
     get_operation() match {
       case None => None
       case Some(op) => op.pretty_text_area.detach_operation
@@ -320,7 +320,7 @@
     }
 
   private val delay_resize =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() }
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
--- a/src/Tools/jEdit/src/rich_text_area.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/rich_text_area.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -159,15 +159,15 @@
 
   private val highlight_area =
     new Active_Area[Color](
-      (rendering: JEdit_Rendering) => rendering.highlight _, None)
+      (rendering: JEdit_Rendering) => rendering.highlight, None)
 
   private val hyperlink_area =
     new Active_Area[PIDE.editor.Hyperlink](
-      (rendering: JEdit_Rendering) => rendering.hyperlink _, Some(Cursor.HAND_CURSOR))
+      (rendering: JEdit_Rendering) => rendering.hyperlink, Some(Cursor.HAND_CURSOR))
 
   private val active_area =
     new Active_Area[XML.Elem](
-      (rendering: JEdit_Rendering) => rendering.active _, Some(Cursor.DEFAULT_CURSOR))
+      (rendering: JEdit_Rendering) => rendering.active, Some(Cursor.DEFAULT_CURSOR))
 
   private val active_areas =
     List((highlight_area, true), (hyperlink_area, true), (active_area, false))
@@ -292,7 +292,7 @@
       robust_rendering { rendering =>
         val fm = text_area.getPainter.getFontMetrics
 
-        for (i <- 0 until physical_lines.length) {
+        for (i <- physical_lines.indices) {
           if (physical_lines(i) != -1) {
             val line_range = Text.Range(start(i), end(i) min buffer.getLength)
 
@@ -479,7 +479,7 @@
           ((w - b + 1) / 2, c - b / 2, w - b, line_height - b)
         }
 
-        for (i <- 0 until physical_lines.length) {
+        for (i <- physical_lines.indices) {
           val line = physical_lines(i)
           if (line != -1) {
             val line_range = Text.Range(start(i), end(i) min buffer.getLength)
@@ -525,7 +525,7 @@
     {
       robust_rendering { rendering =>
         val search_pattern = get_search_pattern()
-        for (i <- 0 until physical_lines.length) {
+        for (i <- physical_lines.indices) {
           if (physical_lines(i) != -1) {
             val line_range = Text.Range(start(i), end(i) min buffer.getLength)
 
--- a/src/Tools/jEdit/src/scala_console.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/scala_console.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -30,7 +30,7 @@
     def find_jars(start: String): List[String] =
       if (start != null)
         File.find_files(new JFile(start), file => file.getName.endsWith(".jar")).
-          map(File.absolute_name(_))
+          map(File.absolute_name)
       else Nil
 
     val initial_class_path =
@@ -64,7 +64,7 @@
         if (global_out == null) System.out.print(str)
         else global_out.writeAttrs(null, str)
       }
-      Thread.sleep(10)  // FIXME adhoc delay to avoid loosing output
+      Time.seconds(0.01).sleep  // FIXME adhoc delay to avoid loosing output
     }
 
     override def close() { flush () }
@@ -152,7 +152,7 @@
           }
           finally {
             running.change(_ => None)
-            Thread.interrupted
+            Exn.Interrupt.dispose()
           }
           GUI_Thread.later {
             if (err != null) err.commandDone()
--- a/src/Tools/jEdit/src/services.xml	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/services.xml	Wed Apr 15 14:31:21 2020 +0200
@@ -44,4 +44,10 @@
   <SERVICE CLASS="console.Shell" NAME="Scala">
     new isabelle.jedit.Scala_Console();
   </SERVICE>
+  <SERVICE CLASS="isabelle.jedit.Active$Handler" NAME="misc">
+    new isabelle.jedit.Active$Misc_Handler();
+  </SERVICE>
+  <SERVICE CLASS="isabelle.jedit.Active$Handler" NAME="graphview">
+    new isabelle.jedit.Graphview_Dockable$Handler()
+  </SERVICE>
 </SERVICES>
--- a/src/Tools/jEdit/src/session_build.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/session_build.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -23,24 +23,24 @@
 {
   def check_dialog(view: View)
   {
-    GUI_Thread.require {}
-
     val options = PIDE.options.value
-    try {
-      if (JEdit_Sessions.session_no_build ||
+    Isabelle_Thread.fork() {
+      try {
+        if (JEdit_Sessions.session_no_build ||
           JEdit_Sessions.session_build(options, no_build = true) == 0)
-            JEdit_Sessions.session_start(options)
-      else new Dialog(view)
-    }
-    catch {
-      case exn: Throwable =>
-        GUI.dialog(view, "Isabelle", GUI.scrollable_text(Exn.message(exn)))
+          JEdit_Sessions.session_start(options)
+        else GUI_Thread.later { new Dialog(view) }
+      }
+      catch {
+        case exn: Throwable =>
+          GUI.dialog(view, "Isabelle", GUI.scrollable_text(Exn.message(exn)))
+      }
     }
   }
 
   private class Dialog(view: View) extends JDialog(view)
   {
-    val options = PIDE.options.value
+    val options: Options = PIDE.options.value
 
 
     /* text */
@@ -57,8 +57,6 @@
 
     /* progress */
 
-    @volatile private var is_stopped = false
-
     private val progress = new Progress {
       override def echo(txt: String): Unit =
         GUI_Thread.later {
@@ -68,8 +66,6 @@
         }
 
       override def theory(theory: Progress.Theory): Unit = echo(theory.message)
-
-      override def stopped: Boolean = is_stopped
     }
 
 
@@ -102,7 +98,7 @@
       }
 
     private val delay_exit =
-      GUI_Thread.delay_first(Time.seconds(1.0))
+      Delay.first(Time.seconds(1.0), gui = true)
       {
         if (can_auto_close) conclude()
         else {
@@ -132,7 +128,7 @@
 
     private def stopping()
     {
-      is_stopped = true
+      progress.stop
       set_actions(new Label("Stopping ..."))
     }
 
@@ -165,7 +161,7 @@
     setLocationRelativeTo(view)
     setVisible(true)
 
-    Standard_Thread.fork("session_build") {
+    Isabelle_Thread.fork(name = "session_build") {
       progress.echo("Build started for Isabelle/" + PIDE.resources.session_name + " ...")
 
       val (out, rc) =
@@ -175,7 +171,7 @@
             (Output.error_message_text(Exn.message(exn)) + "\n", Exn.return_code(exn, 2))
         }
 
-      progress.echo(out + (if (rc == 0) "OK\n" else "Return code: " + rc + "\n"))
+      progress.echo(out + (if (rc == 0) "OK" else Process_Result.print_return_code(rc)) + "\n")
 
       if (rc == 0) JEdit_Sessions.session_start(options)
       else progress.echo("Session build failed -- prover process remains inactive!")
--- a/src/Tools/jEdit/src/simplifier_trace_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/simplifier_trace_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -64,7 +64,7 @@
 
   private def show_trace()
   {
-    val trace = Simplifier_Trace.generate_trace(current_results)
+    val trace = Simplifier_Trace.generate_trace(PIDE.session, current_results)
     new Simplifier_Trace_Window(view, current_snapshot, trace)
   }
 
@@ -144,7 +144,7 @@
   /* resize */
 
   private val delay_resize =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() }
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
@@ -181,7 +181,7 @@
         new Button("Clear memory") {
           reactions += {
             case ButtonClicked(_) =>
-              Simplifier_Trace.clear_memory()
+              Simplifier_Trace.clear_memory(PIDE.session)
           }
         }))
 
--- a/src/Tools/jEdit/src/simplifier_trace_window.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/simplifier_trace_window.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -182,7 +182,7 @@
   /* resize */
 
   private val delay_resize =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() }
 
   peer.addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
--- a/src/Tools/jEdit/src/sledgehammer_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/sledgehammer_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -30,7 +30,7 @@
   val pretty_text_area = new Pretty_Text_Area(view)
   set_content(pretty_text_area)
 
-  override def detach_operation = pretty_text_area.detach_operation
+  override def detach_operation: Option[() => Unit] = pretty_text_area.detach_operation
 
 
   /* query operation */
@@ -50,7 +50,7 @@
   }
 
   private val sledgehammer =
-    new Query_Operation(PIDE.editor, view, "sledgehammer", consume_status _,
+    new Query_Operation(PIDE.editor, view, "sledgehammer", consume_status,
       (snapshot, results, body) =>
         pretty_text_area.update(snapshot, results, Pretty.separate(body)))
 
@@ -58,7 +58,7 @@
   /* resize */
 
   private val delay_resize =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() }
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
--- a/src/Tools/jEdit/src/state_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/state_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -29,7 +29,7 @@
   val pretty_text_area = new Pretty_Text_Area(view)
   set_content(pretty_text_area)
 
-  override def detach_operation = pretty_text_area.detach_operation
+  override def detach_operation: Option[() => Unit] = pretty_text_area.detach_operation
 
   private val print_state =
     new Query_Operation(PIDE.editor, view, "print_state", _ => (),
@@ -40,7 +40,7 @@
   /* resize */
 
   private val delay_resize =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() }
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
--- a/src/Tools/jEdit/src/symbols_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/symbols_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -28,7 +28,7 @@
   private val abbrevs_panel = new Abbrevs_Panel
 
   private val abbrevs_refresh_delay =
-    GUI_Thread.delay_last(PIDE.options.seconds("editor_update_delay")) { abbrevs_panel.refresh }
+    Delay.last(PIDE.options.seconds("editor_update_delay"), gui = true) { abbrevs_panel.refresh }
 
   private class Abbrev_Component(txt: String, abbrs: List[String]) extends Button
   {
@@ -129,7 +129,7 @@
 
     val search_space = for ((sym, _) <- Symbol.codes) yield (sym, Word.lowercase(sym))
     val search_delay =
-      GUI_Thread.delay_last(PIDE.options.seconds("editor_input_delay")) {
+      Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) {
         val search_words = Word.explode(Word.lowercase(search_field.text))
         val search_limit = PIDE.options.int("jedit_symbols_search_limit") max 0
         val results =
@@ -174,7 +174,7 @@
     }
 
     for (page <- pages)
-      page.title = Word.implode(Word.explode('_', page.title).map(Word.perhaps_capitalize(_)))
+      page.title = Word.implode(Word.explode('_', page.title).map(Word.perhaps_capitalize))
   }
   set_content(group_tabs)
 
@@ -200,7 +200,7 @@
                 }
               case _ =>
             })
-          comp.revalidate
+          comp.revalidate()
           comp.repaint()
         }
       case _: Session.Commands_Changed => abbrevs_refresh_delay.invoke()
--- a/src/Tools/jEdit/src/syntax_style.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/syntax_style.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -99,9 +99,9 @@
   }
 
   private def control_style(sym: String): Option[Byte => Byte] =
-    if (sym == Symbol.sub_decoded) Some(subscript(_))
-    else if (sym == Symbol.sup_decoded) Some(superscript(_))
-    else if (sym == Symbol.bold_decoded) Some(bold(_))
+    if (sym == Symbol.sub_decoded) Some(subscript)
+    else if (sym == Symbol.sup_decoded) Some(superscript)
+    else if (sym == Symbol.bold_decoded) Some(bold)
     else None
 
   def extended(text: CharSequence): Map[Text.Offset, Byte => Byte] =
--- a/src/Tools/jEdit/src/syslog_dockable.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/syslog_dockable.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -21,7 +21,7 @@
 
   private val syslog = new TextArea()
 
-  private def syslog_delay = GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay"))
+  private def syslog_delay = Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true)
   {
     val text = PIDE.session.syslog_content()
     if (text != syslog.text) syslog.text = text
--- a/src/Tools/jEdit/src/text_overview.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/text_overview.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -75,7 +75,7 @@
   private var current_overview = Overview()
 
   private val delay_repaint =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { repaint() }
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { repaint() }
 
   override def paintComponent(gfx: Graphics)
   {
@@ -116,7 +116,7 @@
   def revoke(): Unit = { delay_refresh.revoke(); future_refresh.change(x => { x.cancel; x }) }
 
   private val delay_refresh =
-    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) {
+    Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) {
       if (!doc_view.rich_text_area.check_robust_body) invoke()
       else {
         JEdit_Lib.buffer_lock(buffer) {
--- a/src/Tools/jEdit/src/text_structure.scala	Wed Apr 15 12:33:42 2020 +0200
+++ b/src/Tools/jEdit/src/text_structure.scala	Wed Apr 15 14:31:21 2020 +0200
@@ -21,7 +21,7 @@
 
   class Navigator(syntax: Outer_Syntax, buffer: JEditBuffer, comments: Boolean)
   {
-    val limit = PIDE.options.value.int("jedit_structure_limit") max 0
+    val limit: Int = PIDE.options.value.int("jedit_structure_limit") max 0
 
     def iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] =
     {
@@ -137,7 +137,7 @@
                   else i })
 
           def indent_extra: Int =
-            if (prev_span.exists(keywords.is_quasi_command(_))) indent_size
+            if (prev_span.exists(keywords.is_quasi_command)) indent_size
             else 0
 
           val indent =