merged
authorwenzelm
Sun, 03 Sep 2023 19:28:59 +0200
changeset 78643 d5a1d64a563d
parent 78642 4f2215cfc3e0 (current diff)
parent 78622 c42f316f0a01 (diff)
child 78644 a7bcd2af7190
child 78646 fff610f1a6f4
merged
Admin/components/components.sha1
etc/settings
src/Pure/Tools/build.scala
src/Pure/Tools/build_cluster.scala
--- a/Admin/components/components.sha1	Sun Sep 03 18:09:00 2023 +0200
+++ b/Admin/components/components.sha1	Sun Sep 03 19:28:59 2023 +0200
@@ -113,6 +113,7 @@
 b1c40ce6c087da7e70e221ddd3fcadfa569acb2f foiltex-2.1.4b.tar.gz
 f339234ec18369679be0095264e0c0af7762f351 gnu-utils-20210414.tar.gz
 71259aa46134e6cf2c6473b4fc408051b3336490 gnu-utils-20211030.tar.gz
+c489cae2a96ce18bec813d2eb1f528c88006382a go-1.21.0-v1.tar.gz
 683acd94761ef460cca1a628f650355370de5afb hol-light-bundle-0.5-126.tar.gz
 511fa8df8be88eb0500032bbd17742d33bdd4636 hugo-0.88.1.tar.gz
 989234b3799fe8750f3c24825d1f717c24fb0214 idea-icons-20210508.tar.gz
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/components/go	Sun Sep 03 19:28:59 2023 +0200
@@ -0,0 +1,1 @@
+go-1.21.0-v1
--- a/CONTRIBUTORS	Sun Sep 03 18:09:00 2023 +0200
+++ b/CONTRIBUTORS	Sun Sep 03 19:28:59 2023 +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 Isabelle2023
 -----------------------------
 
--- a/NEWS	Sun Sep 03 18:09:00 2023 +0200
+++ b/NEWS	Sun Sep 03 19:28:59 2023 +0200
@@ -4,6 +4,17 @@
 (Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.)
 
 
+New in this Isabelle version
+----------------------------
+
+*** System ***
+
+* Isabelle/Scala and derived Scala tools now use the syntax of Scala
+3.3, instead of 3.1. This is the "-old-syntax" variant (Java-like) as
+before, not "-new-syntax" (Python-like). Minor INCOMPATIBILITY.
+
+
+
 New in Isabelle2023 (September 2023)
 ------------------------------------
 
--- a/etc/settings	Sun Sep 03 18:09:00 2023 +0200
+++ b/etc/settings	Sun Sep 03 19:28:59 2023 +0200
@@ -16,8 +16,8 @@
 
 ISABELLE_TOOL_JAVA_OPTIONS="-Djava.awt.headless=true -Xms512m -Xmx4g -Xss16m"
 
-ISABELLE_JAVAC_OPTIONS="-encoding UTF-8 -Xlint:-options -deprecation -source 11 -target 11"
-ISABELLE_SCALAC_OPTIONS="-encoding UTF-8 -feature -deprecation -release 11 -source 3.1 -old-syntax -no-indent -color never -pagewidth 78 -J-Xms512m -J-Xmx4g -J-Xss16m"
+ISABELLE_JAVAC_OPTIONS="-encoding UTF-8 -Xlint:-options -deprecation -source 17 -target 17"
+ISABELLE_SCALAC_OPTIONS="-encoding UTF-8 -feature -java-output-version 17 -source 3.3 -old-syntax -no-indent -color never -pagewidth 78 -J-Xms512m -J-Xmx4g -J-Xss16m"
 
 ISABELLE_SCALA_JAR="$ISABELLE_HOME/lib/classes/isabelle.jar"
 
--- a/lib/Tools/scala_build	Sun Sep 03 18:09:00 2023 +0200
+++ b/lib/Tools/scala_build	Sun Sep 03 19:28:59 2023 +0200
@@ -51,12 +51,6 @@
 
 ## main
 
-#remove historic material
-rm -rf \
-  "$ISABELLE_HOME/lib/classes/Pure.jar" \
-  "$ISABELLE_HOME/lib/classes/Pure.shasum" \
-  "$ISABELLE_HOME/src/Tools/jEdit/dist"
-
 classpath "$CLASSPATH"; export CLASSPATH=""
 
 eval "declare -a JAVA_ARGS=($ISABELLE_TOOL_JAVA_OPTIONS)"
--- a/src/HOL/Algebra/Ring.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Algebra/Ring.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -878,7 +878,7 @@
 lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>"
   by (simp add: inv_char local.ring_axioms ring.r_minus)
 
-lemma (in monoid) inv_eq_imp_eq: "x \<in> Units G \<Longrightarrow> y \<in> Units G \<Longrightarrow> inv x = inv y \<Longrightarrow> x = y"
+lemma (in monoid) inv_eq_imp_eq [dest!]: "inv x = inv y \<Longrightarrow> x \<in> Units G \<Longrightarrow> y \<in> Units G \<Longrightarrow> x = y"
   by (metis Units_inv_inv)
 
 lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> \<in> Units R"
--- a/src/HOL/Analysis/Abstract_Topological_Spaces.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Abstract_Topological_Spaces.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -35,10 +35,6 @@
    "\<lbrakk>connectedin X C; closedin X T; openin X T\<rbrakk> \<Longrightarrow> C \<subseteq> T \<or> disjnt C T"
   by (metis Diff_eq_empty_iff Int_empty_right clopenin_eq_frontier_of connectedin_Int_frontier_of disjnt_def)
 
-lemma connected_space_quotient_map_image:
-   "\<lbrakk>quotient_map X X' q; connected_space X\<rbrakk> \<Longrightarrow> connected_space X'"
-  by (metis connectedin_continuous_map_image connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map)
-
 lemma connected_space_retraction_map_image:
    "\<lbrakk>retraction_map X X' r; connected_space X\<rbrakk> \<Longrightarrow> connected_space X'"
   using connected_space_quotient_map_image retraction_imp_quotient_map by blast
@@ -291,8 +287,6 @@
     by (simp add: Ceq \<open>a \<in> U\<close> \<open>a \<in> topspace X\<close> connected_component_in_connected_components_of)
 qed
 
-thm connected_space_iff_components_eq
-
 lemma open_in_finite_connected_components:
   assumes "finite(connected_components_of X)" "C \<in> connected_components_of X"
   shows "openin X C"
@@ -307,12 +301,7 @@
 lemma connected_components_of_disjoint:
   assumes "C \<in> connected_components_of X" "C' \<in> connected_components_of X"
     shows "(disjnt C C' \<longleftrightarrow> (C \<noteq> C'))"
-proof -
-  have "C \<noteq> {}"
-    using nonempty_connected_components_of assms by blast
-  with assms show ?thesis
-    by (metis disjnt_self_iff_empty pairwiseD pairwise_disjoint_connected_components_of)
-qed
+  using assms nonempty_connected_components_of pairwiseD pairwise_disjoint_connected_components_of by fastforce
 
 lemma connected_components_of_overlap:
    "\<lbrakk>C \<in> connected_components_of X; C' \<in> connected_components_of X\<rbrakk> \<Longrightarrow> C \<inter> C' \<noteq> {} \<longleftrightarrow> C = C'"
@@ -439,8 +428,7 @@
 next
   case False
   then show ?thesis
-    apply (simp add: PiE_iff)
-    by (smt (verit) Collect_empty_eq False PiE_eq_empty_iff PiE_iff connected_component_of_eq_empty)
+    by (smt (verit, best) PiE_eq_empty_iff PiE_iff connected_component_of_eq_empty topspace_product_topology)
 qed
 
 
@@ -490,11 +478,14 @@
 
 lemma monotone_map_from_subtopology:
   assumes "monotone_map X Y f" 
-    "\<And>x y. x \<in> topspace X \<and> y \<in> topspace X \<and> x \<in> S \<and> f x = f y \<Longrightarrow> y \<in> S"
+    "\<And>x y. \<lbrakk>x \<in> topspace X; y \<in> topspace X; x \<in> S; f x = f y\<rbrakk> \<Longrightarrow> y \<in> S"
   shows "monotone_map (subtopology X S) Y f"
-  using assms
-  unfolding monotone_map_def connectedin_subtopology
-  by (smt (verit, del_insts) Collect_cong Collect_empty_eq IntE IntI connectedin_empty image_subset_iff mem_Collect_eq subsetI topspace_subtopology)
+proof -
+  have "\<And>y. y \<in> topspace Y \<Longrightarrow> connectedin X {x \<in> topspace X. x \<in> S \<and> f x = y}"
+    by (smt (verit) Collect_cong assms connectedin_empty empty_def monotone_map_def)
+  then show ?thesis
+    using assms by (auto simp: monotone_map_def connectedin_subtopology)
+qed
 
 lemma monotone_map_restriction:
   "monotone_map X Y f \<and> {x \<in> topspace X. f x \<in> v} = u
@@ -5115,17 +5106,10 @@
           locally_compact_space_prod_topology by blast
 qed
 
-text \<open>Essentially the same proof\<close>
 lemma k_space_prod_topology_right:
   assumes "k_space X" and Y: "locally_compact_space Y" "Hausdorff_space Y \<or> regular_space Y"
   shows "k_space (prod_topology X Y)"
-proof -
-  obtain q and Z :: "('a set * 'a) topology" where "locally_compact_space Z" and q: "quotient_map Z X q"
-    using \<open>k_space X\<close> k_space_as_quotient by blast
-  then show ?thesis
-    using quotient_map_prod_left [OF Y q] using Y k_space_quotient_map_image locally_compact_imp_k_space 
-          locally_compact_space_prod_topology by blast
-qed
+  using assms homeomorphic_k_space homeomorphic_space_prod_topology_swap k_space_prod_topology_left by blast
 
 
 lemma continuous_map_from_k_space:
@@ -5258,13 +5242,7 @@
          f \<in> (topspace X) \<rightarrow> topspace Y \<and>
          (\<forall>k. compactin Y k
               \<longrightarrow> open_map (subtopology X {x \<in> topspace X. f x \<in> k}) (subtopology Y k) f)"
-       (is "?lhs=?rhs")
-proof
-  show "?lhs \<Longrightarrow> ?rhs"
-    by (simp add: open_map_imp_subset_topspace open_map_restriction)
-  show "?rhs \<Longrightarrow> ?lhs"
-    by (simp add: assms open_map_into_k_space)
-qed
+  using assms open_map_imp_subset_topspace open_map_into_k_space open_map_restriction by fastforce
 
 lemma closed_map_into_k_space_eq:
   assumes "k_space Y"
--- a/src/HOL/Analysis/Affine.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Affine.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -5,7 +5,7 @@
 begin
 
 lemma if_smult: "(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)"
-  by (fact if_distrib)
+  by simp
 
 lemma sum_delta_notmem:
   assumes "x \<notin> s"
@@ -13,12 +13,7 @@
     and "sum (\<lambda>y. if (x = y) then P x else Q y) s = sum Q s"
     and "sum (\<lambda>y. if (y = x) then P y else Q y) s = sum Q s"
     and "sum (\<lambda>y. if (x = y) then P y else Q y) s = sum Q s"
-  apply (rule_tac [!] sum.cong)
-  using assms
-  apply auto
-  done
-
-lemmas independent_finite = independent_imp_finite
+  by (smt (verit, best) assms sum.cong)+
 
 lemma span_substd_basis:
   assumes d: "d \<subseteq> Basis"
@@ -32,13 +27,11 @@
   ultimately have "span d \<subseteq> ?B"
     using span_mono[of d "?B"] span_eq_iff[of "?B"] by blast
   moreover have *: "card d \<le> dim (span d)"
-    using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms]
-      span_superset[of d]
-    by auto
+    by (simp add: d dim_eq_card_independent independent_substdbasis)
   moreover from * have "dim ?B \<le> dim (span d)"
     using dim_substandard[OF assms] by auto
   ultimately show ?thesis
-    using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto
+    by (simp add: s subspace_dim_equal)
 qed
 
 lemma basis_to_substdbasis_subspace_isomorphism:
@@ -51,7 +44,8 @@
     using dim_unique[of B B "card B"] assms span_superset[of B] by auto
   have "dim B \<le> card (Basis :: 'a set)"
     using dim_subset_UNIV[of B] by simp
-  from obtain_subset_with_card_n[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B"
+  from obtain_subset_with_card_n[OF this] 
+  obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B"
     by auto
   let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
   have "\<exists>f. linear f \<and> f ` B = d \<and> f ` span B = ?t \<and> inj_on f (span B)"
@@ -65,9 +59,9 @@
 subsection \<open>Affine set and affine hull\<close>
 
 definition\<^marker>\<open>tag important\<close> affine :: "'a::real_vector set \<Rightarrow> bool"
-  where "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
+  where "affine S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> S)"
 
-lemma affine_alt: "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)"
+lemma affine_alt: "affine S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S)"
   unfolding affine_def by (metis eq_diff_eq')
 
 lemma affine_empty [iff]: "affine {}"
@@ -79,21 +73,21 @@
 lemma affine_UNIV [iff]: "affine UNIV"
   unfolding affine_def by auto
 
-lemma affine_Inter [intro]: "(\<And>s. s\<in>f \<Longrightarrow> affine s) \<Longrightarrow> affine (\<Inter>f)"
+lemma affine_Inter [intro]: "(\<And>S. S\<in>\<F> \<Longrightarrow> affine S) \<Longrightarrow> affine (\<Inter>\<F>)"
   unfolding affine_def by auto
 
-lemma affine_Int[intro]: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)"
+lemma affine_Int[intro]: "affine S \<Longrightarrow> affine T \<Longrightarrow> affine (S \<inter> T)"
   unfolding affine_def by auto
 
-lemma affine_scaling: "affine s \<Longrightarrow> affine (image (\<lambda>x. c *\<^sub>R x) s)"
-  apply (clarsimp simp add: affine_def)
+lemma affine_scaling: "affine S \<Longrightarrow> affine ((*\<^sub>R) c ` S)"
+  apply (clarsimp simp: affine_def)
   apply (rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in image_eqI)
   apply (auto simp: algebra_simps)
   done
 
-lemma affine_affine_hull [simp]: "affine(affine hull s)"
+lemma affine_affine_hull [simp]: "affine(affine hull S)"
   unfolding hull_def
-  using affine_Inter[of "{t. affine t \<and> s \<subseteq> t}"] by auto
+  using affine_Inter[of "{T. affine T \<and> S \<subseteq> T}"] by auto
 
 lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s"
   by (metis affine_affine_hull hull_same)
@@ -198,6 +192,8 @@
   "affine hull p = {y. \<exists>S u. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p \<and> sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}"
   (is "_ = ?rhs")
 proof (rule hull_unique)
+  have "\<And>x. sum (\<lambda>z. 1) {x} = 1"
+      by auto
   show "p \<subseteq> ?rhs"
   proof (intro subsetI CollectI exI conjI)
     show "\<And>x. sum (\<lambda>z. 1) {x} = 1"
@@ -452,31 +448,16 @@
   unfolding affine_parallel_def
   using image_add_0 by blast
 
-lemma affine_parallel_commut:
+lemma affine_parallel_commute:
   assumes "affine_parallel A B"
   shows "affine_parallel B A"
-proof -
-  from assms obtain a where B: "B = (\<lambda>x. a + x) ` A"
-    unfolding affine_parallel_def by auto
-  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
-  from B show ?thesis
-    using translation_galois [of B a A]
-    unfolding affine_parallel_def by blast
-qed
+  by (metis affine_parallel_def assms translation_galois)
 
 lemma affine_parallel_assoc:
   assumes "affine_parallel A B"
     and "affine_parallel B C"
   shows "affine_parallel A C"
-proof -
-  from assms obtain ab where "B = (\<lambda>x. ab + x) ` A"
-    unfolding affine_parallel_def by auto
-  moreover
-  from assms obtain bc where "C = (\<lambda>x. bc + x) ` B"
-    unfolding affine_parallel_def by auto
-  ultimately show ?thesis
-    using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto
-qed
+  by (metis affine_parallel_def assms translation_assoc)
 
 lemma affine_translation_aux:
   fixes a :: "'a::real_vector"
@@ -503,24 +484,13 @@
 
 lemma affine_translation:
   "affine S \<longleftrightarrow> affine ((+) a ` S)" for a :: "'a::real_vector"
-proof
-  show "affine ((+) a ` S)" if "affine S"
-    using that translation_assoc [of "- a" a S]
-    by (auto intro: affine_translation_aux [of "- a" "((+) a ` S)"])
-  show "affine S" if "affine ((+) a ` S)"
-    using that by (rule affine_translation_aux)
-qed
+  by (metis affine_translation_aux translation_galois)
 
 lemma parallel_is_affine:
   fixes S T :: "'a::real_vector set"
   assumes "affine S" "affine_parallel S T"
   shows "affine T"
-proof -
-  from assms obtain a where "T = (\<lambda>x. a + x) ` S"
-    unfolding affine_parallel_def by auto
-  then show ?thesis
-    using affine_translation assms by auto
-qed
+  by (metis affine_parallel_def affine_translation assms)
 
 lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s"
   unfolding subspace_def affine_def by auto
@@ -532,64 +502,12 @@
 subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Subspace parallel to an affine set\<close>
 
 lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S"
-proof -
-  have h0: "subspace S \<Longrightarrow> affine S \<and> 0 \<in> S"
-    using subspace_imp_affine[of S] subspace_0 by auto
-  {
-    assume assm: "affine S \<and> 0 \<in> S"
-    {
-      fix c :: real
-      fix x
-      assume x: "x \<in> S"
-      have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto
-      moreover
-      have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \<in> S"
-        using affine_alt[of S] assm x by auto
-      ultimately have "c *\<^sub>R x \<in> S" by auto
-    }
-    then have h1: "\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S" by auto
-
-    {
-      fix x y
-      assume xy: "x \<in> S" "y \<in> S"
-      define u where "u = (1 :: real)/2"
-      have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)"
-        by auto
-      moreover
-      have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y"
-        by (simp add: algebra_simps)
-      moreover
-      have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S"
-        using affine_alt[of S] assm xy by auto
-      ultimately
-      have "(1/2) *\<^sub>R (x+y) \<in> S"
-        using u_def by auto
-      moreover
-      have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))"
-        by auto
-      ultimately
-      have "x + y \<in> S"
-        using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto
-    }
-    then have "\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S"
-      by auto
-    then have "subspace S"
-      using h1 assm unfolding subspace_def by auto
-  }
-  then show ?thesis using h0 by metis
-qed
+  by (metis add_cancel_right_left affine_alt diff_add_cancel mem_affine_3 scaleR_eq_0_iff subspace_def vector_space_assms(4))
 
 lemma affine_diffs_subspace:
   assumes "affine S" "a \<in> S"
   shows "subspace ((\<lambda>x. (-a)+x) ` S)"
-proof -
-  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
-  have "affine ((\<lambda>x. (-a)+x) ` S)"
-    using affine_translation assms by blast
-  moreover have "0 \<in> ((\<lambda>x. (-a)+x) ` S)"
-    using assms exI[of "(\<lambda>x. x\<in>S \<and> -a+x = 0)" a] by auto
-  ultimately show ?thesis using subspace_affine by auto
-qed
+  by (metis ab_left_minus affine_translation assms image_eqI subspace_affine)
 
 lemma affine_diffs_subspace_subtract:
   "subspace ((\<lambda>x. x - a) ` S)" if "affine S" "a \<in> S"
@@ -600,61 +518,26 @@
     and "a \<in> S"
   assumes "L \<equiv> {y. \<exists>x \<in> S. (-a) + x = y}"
   shows "subspace L \<and> affine_parallel S L"
-proof -
-  from assms have "L = plus (- a) ` S" by auto
-  then have par: "affine_parallel S L"
-    unfolding affine_parallel_def ..
-  then have "affine L" using assms parallel_is_affine by auto
-  moreover have "0 \<in> L"
-    using assms by auto
-  ultimately show ?thesis
-    using subspace_affine par by auto
-qed
+  by (smt (verit) Collect_cong ab_left_minus affine_parallel_def assms image_def mem_Collect_eq parallel_is_affine subspace_affine)
 
 lemma parallel_subspace_aux:
   assumes "subspace A"
     and "subspace B"
     and "affine_parallel A B"
   shows "A \<supseteq> B"
-proof -
-  from assms obtain a where a: "\<forall>x. x \<in> A \<longleftrightarrow> a + x \<in> B"
-    using affine_parallel_expl[of A B] by auto
-  then have "-a \<in> A"
-    using assms subspace_0[of B] by auto
-  then have "a \<in> A"
-    using assms subspace_neg[of A "-a"] by auto
-  then show ?thesis
-    using assms a unfolding subspace_def by auto
-qed
+  by (metis add.right_neutral affine_parallel_expl assms subsetI subspace_def)
 
 lemma parallel_subspace:
   assumes "subspace A"
     and "subspace B"
     and "affine_parallel A B"
   shows "A = B"
-proof
-  show "A \<supseteq> B"
-    using assms parallel_subspace_aux by auto
-  show "A \<subseteq> B"
-    using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto
-qed
+  by (simp add: affine_parallel_commute assms parallel_subspace_aux subset_antisym)
 
 lemma affine_parallel_subspace:
   assumes "affine S" "S \<noteq> {}"
   shows "\<exists>!L. subspace L \<and> affine_parallel S L"
-proof -
-  have ex: "\<exists>L. subspace L \<and> affine_parallel S L"
-    using assms parallel_subspace_explicit by auto
-  {
-    fix L1 L2
-    assume ass: "subspace L1 \<and> affine_parallel S L1" "subspace L2 \<and> affine_parallel S L2"
-    then have "affine_parallel L1 L2"
-      using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto
-    then have "L1 = L2"
-      using ass parallel_subspace by auto
-  }
-  then show ?thesis using ex by auto
-qed
+  by (meson affine_parallel_assoc affine_parallel_commute assms equals0I parallel_subspace parallel_subspace_explicit)
 
 
 subsection \<open>Affine Dependence\<close>
@@ -662,50 +545,49 @@
 text "Formalized by Lars Schewe."
 
 definition\<^marker>\<open>tag important\<close> affine_dependent :: "'a::real_vector set \<Rightarrow> bool"
-  where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
+  where "affine_dependent S \<longleftrightarrow> (\<exists>x\<in>S. x \<in> affine hull (S - {x}))"
 
-lemma affine_dependent_imp_dependent: "affine_dependent s \<Longrightarrow> dependent s"
+lemma affine_dependent_imp_dependent: "affine_dependent S \<Longrightarrow> dependent S"
   unfolding affine_dependent_def dependent_def
   using affine_hull_subset_span by auto
 
 lemma affine_dependent_subset:
-   "\<lbrakk>affine_dependent s; s \<subseteq> t\<rbrakk> \<Longrightarrow> affine_dependent t"
-apply (simp add: affine_dependent_def Bex_def)
-apply (blast dest: hull_mono [OF Diff_mono [OF _ subset_refl]])
-done
+   "\<lbrakk>affine_dependent S; S \<subseteq> T\<rbrakk> \<Longrightarrow> affine_dependent T"
+  using hull_mono [OF Diff_mono [OF _ subset_refl]]
+  by (smt (verit) affine_dependent_def subsetD)
 
 lemma affine_independent_subset:
-  shows "\<lbrakk>\<not> affine_dependent t; s \<subseteq> t\<rbrakk> \<Longrightarrow> \<not> affine_dependent s"
-by (metis affine_dependent_subset)
+  shows "\<lbrakk>\<not> affine_dependent T; S \<subseteq> T\<rbrakk> \<Longrightarrow> \<not> affine_dependent S"
+  by (metis affine_dependent_subset)
 
 lemma affine_independent_Diff:
-   "\<not> affine_dependent s \<Longrightarrow> \<not> affine_dependent(s - t)"
+   "\<not> affine_dependent S \<Longrightarrow> \<not> affine_dependent(S - T)"
 by (meson Diff_subset affine_dependent_subset)
 
 proposition affine_dependent_explicit:
   "affine_dependent p \<longleftrightarrow>
-    (\<exists>S u. finite S \<and> S \<subseteq> p \<and> sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) S = 0)"
+    (\<exists>S U. finite S \<and> S \<subseteq> p \<and> sum U S = 0 \<and> (\<exists>v\<in>S. U v \<noteq> 0) \<and> sum (\<lambda>v. U v *\<^sub>R v) S = 0)"
 proof -
-  have "\<exists>S u. finite S \<and> S \<subseteq> p \<and> sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> (\<Sum>w\<in>S. u w *\<^sub>R w) = 0"
-    if "(\<Sum>w\<in>S. u w *\<^sub>R w) = x" "x \<in> p" "finite S" "S \<noteq> {}" "S \<subseteq> p - {x}" "sum u S = 1" for x S u
+  have "\<exists>S U. finite S \<and> S \<subseteq> p \<and> sum U S = 0 \<and> (\<exists>v\<in>S. U v \<noteq> 0) \<and> (\<Sum>w\<in>S. U w *\<^sub>R w) = 0"
+    if "(\<Sum>w\<in>S. U w *\<^sub>R w) = x" "x \<in> p" "finite S" "S \<noteq> {}" "S \<subseteq> p - {x}" "sum U S = 1" for x S U
   proof (intro exI conjI)
     have "x \<notin> S" 
       using that by auto
-    then show "(\<Sum>v \<in> insert x S. if v = x then - 1 else u v) = 0"
+    then show "(\<Sum>v \<in> insert x S. if v = x then - 1 else U v) = 0"
       using that by (simp add: sum_delta_notmem)
-    show "(\<Sum>w \<in> insert x S. (if w = x then - 1 else u w) *\<^sub>R w) = 0"
+    show "(\<Sum>w \<in> insert x S. (if w = x then - 1 else U w) *\<^sub>R w) = 0"
       using that \<open>x \<notin> S\<close> by (simp add: if_smult sum_delta_notmem cong: if_cong)
   qed (use that in auto)
-  moreover have "\<exists>x\<in>p. \<exists>S u. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p - {x} \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = x"
-    if "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0" "finite S" "S \<subseteq> p" "sum u S = 0" "v \<in> S" "u v \<noteq> 0" for S u v
+  moreover have "\<exists>x\<in>p. \<exists>S U. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p - {x} \<and> sum U S = 1 \<and> (\<Sum>v\<in>S. U v *\<^sub>R v) = x"
+    if "(\<Sum>v\<in>S. U v *\<^sub>R v) = 0" "finite S" "S \<subseteq> p" "sum U S = 0" "v \<in> S" "U v \<noteq> 0" for S U v
   proof (intro bexI exI conjI)
     have "S \<noteq> {v}"
       using that by auto
     then show "S - {v} \<noteq> {}"
       using that by auto
-    show "(\<Sum>x \<in> S - {v}. - (1 / u v) * u x) = 1"
+    show "(\<Sum>x \<in> S - {v}. - (1 / U v) * U x) = 1"
       unfolding sum_distrib_left[symmetric] sum_diff1[OF \<open>finite S\<close>] by (simp add: that)
-    show "(\<Sum>x\<in>S - {v}. (- (1 / u v) * u x) *\<^sub>R x) = v"
+    show "(\<Sum>x\<in>S - {v}. (- (1 / U v) * U x) *\<^sub>R x) = v"
       unfolding sum_distrib_left [symmetric] scaleR_scaleR[symmetric]
                 scaleR_right.sum [symmetric] sum_diff1[OF \<open>finite S\<close>] 
       using that by auto
@@ -720,90 +602,82 @@
   fixes S :: "'a::real_vector set"
   assumes "finite S"
   shows "affine_dependent S \<longleftrightarrow>
-    (\<exists>u. sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) S = 0)"
+    (\<exists>U. sum U S = 0 \<and> (\<exists>v\<in>S. U v \<noteq> 0) \<and> sum (\<lambda>v. U v *\<^sub>R v) S = 0)"
   (is "?lhs = ?rhs")
 proof
-  have *: "\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)"
+  have *: "\<And>vt U v. (if vt then U v else 0) *\<^sub>R v = (if vt then (U v) *\<^sub>R v else 0::'a)"
     by auto
   assume ?lhs
-  then obtain t u v where
-    "finite t" "t \<subseteq> S" "sum u t = 0" "v\<in>t" "u v \<noteq> 0"  "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
+  then obtain T U v where
+    "finite T" "T \<subseteq> S" "sum U T = 0" "v\<in>T" "U v \<noteq> 0"  "(\<Sum>v\<in>T. U v *\<^sub>R v) = 0"
     unfolding affine_dependent_explicit by auto
   then show ?rhs
-    apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
-    apply (auto simp: * sum.inter_restrict[OF assms, symmetric] Int_absorb1[OF \<open>t\<subseteq>S\<close>])
+    apply (rule_tac x="\<lambda>x. if x\<in>T then U x else 0" in exI)
+    apply (auto simp: * sum.inter_restrict[OF assms, symmetric] Int_absorb1[OF \<open>T\<subseteq>S\<close>])
     done
 next
   assume ?rhs
-  then obtain u v where "sum u S = 0"  "v\<in>S" "u v \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
+  then obtain U v where "sum U S = 0"  "v\<in>S" "U v \<noteq> 0" "(\<Sum>v\<in>S. U v *\<^sub>R v) = 0"
     by auto
   then show ?lhs unfolding affine_dependent_explicit
     using assms by auto
 qed
 
 lemma dependent_imp_affine_dependent:
-  assumes "dependent {x - a| x . x \<in> s}"
-    and "a \<notin> s"
-  shows "affine_dependent (insert a s)"
+  assumes "dependent {x - a| x . x \<in> S}"
+    and "a \<notin> S"
+  shows "affine_dependent (insert a S)"
 proof -
-  from assms(1)[unfolded dependent_explicit] obtain S u v
-    where obt: "finite S" "S \<subseteq> {x - a |x. x \<in> s}" "v\<in>S" "u v  \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
+  from assms(1)[unfolded dependent_explicit] obtain S' U v
+    where S: "finite S'" "S' \<subseteq> {x - a |x. x \<in> S}" "v\<in>S'" "U v  \<noteq> 0" "(\<Sum>v\<in>S'. U v *\<^sub>R v) = 0"
     by auto
-  define t where "t = (\<lambda>x. x + a) ` S"
-
-  have inj: "inj_on (\<lambda>x. x + a) S"
+  define T where "T = (\<lambda>x. x + a) ` S'"
+  have inj: "inj_on (\<lambda>x. x + a) S'"
     unfolding inj_on_def by auto
-  have "0 \<notin> S"
-    using obt(2) assms(2) unfolding subset_eq by auto
-  have fin: "finite t" and "t \<subseteq> s"
-    unfolding t_def using obt(1,2) by auto
-  then have "finite (insert a t)" and "insert a t \<subseteq> insert a s"
+  have "0 \<notin> S'"
+    using S(2) assms(2) unfolding subset_eq by auto
+  have fin: "finite T" and "T \<subseteq> S"
+    unfolding T_def using S(1,2) by auto
+  then have "finite (insert a T)" and "insert a T \<subseteq> insert a S"
     by auto
-  moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
-    apply (rule sum.cong)
-    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
-    apply auto
-    done
-  have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
-    unfolding sum_clauses(2)[OF fin] * using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> by auto
-  moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0"
-    using obt(3,4) \<open>0\<notin>S\<close>
-    by (rule_tac x="v + a" in bexI) (auto simp: t_def)
-  moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)"
-    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> by (auto intro!: sum.cong)
-  have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
+  moreover have *: "\<And>P Q. (\<Sum>x\<in>T. (if x = a then P x else Q x)) = (\<Sum>x\<in>T. Q x)"
+    by (smt (verit, best) \<open>T \<subseteq> S\<close> assms(2) subsetD sum.cong)
+  have "(\<Sum>x\<in>insert a T. if x = a then - (\<Sum>x\<in>T. U (x - a)) else U (x - a)) = 0"
+    by (smt (verit) \<open>T \<subseteq> S\<close> assms(2) fin insert_absorb insert_subset sum.insert sum_mono)
+  moreover have "\<exists>v\<in>insert a T. (if v = a then - (\<Sum>x\<in>T. U (x - a)) else U (v - a)) \<noteq> 0"
+    using S(3,4) \<open>0\<notin>S'\<close>
+    by (rule_tac x="v + a" in bexI) (auto simp: T_def)
+  moreover have *: "\<And>P Q. (\<Sum>x\<in>T. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>T. Q x *\<^sub>R x)"
+    using \<open>a\<notin>S\<close> \<open>T\<subseteq>S\<close> by (auto intro!: sum.cong)
+  have "(\<Sum>x\<in>T. U (x - a)) *\<^sub>R a = (\<Sum>v\<in>T. U (v - a) *\<^sub>R v)"
     unfolding scaleR_left.sum
-    unfolding t_def and sum.reindex[OF inj] and o_def
-    using obt(5)
+    unfolding T_def and sum.reindex[OF inj] and o_def
+    using S(5)
     by (auto simp: sum.distrib scaleR_right_distrib)
-  then have "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0"
-    unfolding sum_clauses(2)[OF fin]
-    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
-    by (auto simp: *)
+  then have "(\<Sum>v\<in>insert a T. (if v = a then - (\<Sum>x\<in>T. U (x - a)) else U (v - a)) *\<^sub>R v) = 0"
+    unfolding sum_clauses(2)[OF fin] using \<open>a\<notin>S\<close> \<open>T\<subseteq>S\<close> by (auto simp: *)
   ultimately show ?thesis
     unfolding affine_dependent_explicit
-    apply (rule_tac x="insert a t" in exI, auto)
-    done
+    by (force intro!: exI[where x="insert a T"])
 qed
 
 lemma affine_dependent_biggerset:
-  fixes s :: "'a::euclidean_space set"
-  assumes "finite s" "card s \<ge> DIM('a) + 2"
-  shows "affine_dependent s"
+  fixes S :: "'a::euclidean_space set"
+  assumes "finite S" "card S \<ge> DIM('a) + 2"
+  shows "affine_dependent S"
 proof -
-  have "s \<noteq> {}" using assms by auto
-  then obtain a where "a\<in>s" by auto
-  have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
+  have "S \<noteq> {}" using assms by auto
+  then obtain a where "a\<in>S" by auto
+  have *: "{x - a |x. x \<in> S - {a}} = (\<lambda>x. x - a) ` (S - {a})"
     by auto
-  have "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
+  have "card {x - a |x. x \<in> S - {a}} = card (S - {a})"
     unfolding * by (simp add: card_image inj_on_def)
   also have "\<dots> > DIM('a)" using assms(2)
-    unfolding card_Diff_singleton[OF \<open>a\<in>s\<close>] by auto
-  finally show ?thesis
-    apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
-    apply (rule dependent_imp_affine_dependent)
-    apply (rule dependent_biggerset, auto)
-    done
+    unfolding card_Diff_singleton[OF \<open>a\<in>S\<close>] by auto
+  finally  have "affine_dependent (insert a (S - {a}))"
+    using dependent_biggerset dependent_imp_affine_dependent by blast
+  then show ?thesis
+    by (simp add: \<open>a \<in> S\<close> insert_absorb)
 qed
 
 lemma affine_dependent_biggerset_general:
@@ -822,13 +696,10 @@
   also have "\<dots> < dim S + 1" by auto
   also have "\<dots> \<le> card (S - {a})"
     using assms card_Diff_singleton[OF \<open>a\<in>S\<close>] by auto
-  finally show ?thesis
-    apply (subst insert_Diff[OF \<open>a\<in>S\<close>, symmetric])
-    apply (rule dependent_imp_affine_dependent)
-    apply (rule dependent_biggerset_general)
-    unfolding **
-    apply auto
-    done
+  finally have "affine_dependent (insert a (S - {a}))"
+    by (smt (verit) Collect_cong dependent_imp_affine_dependent dependent_biggerset_general ** Diff_iff insertCI)
+  then show ?thesis
+    by (simp add: \<open>a \<in> S\<close> insert_absorb)
 qed
 
 
@@ -882,16 +753,7 @@
 
 lemma affine_dependent_translation_eq:
   "affine_dependent S \<longleftrightarrow> affine_dependent ((\<lambda>x. a + x) ` S)"
-proof -
-  {
-    assume "affine_dependent ((\<lambda>x. a + x) ` S)"
-    then have "affine_dependent S"
-      using affine_dependent_translation[of "((\<lambda>x. a + x) ` S)" "-a"] translation_assoc[of "-a" a]
-      by auto
-  }
-  then show ?thesis
-    using affine_dependent_translation by auto
-qed
+  by (metis affine_dependent_translation translation_galois)
 
 lemma affine_hull_0_dependent:
   assumes "0 \<in> affine hull S"
@@ -919,14 +781,8 @@
   ultimately have "x \<in> span (S - {x})" by auto
   then have "x \<noteq> 0 \<Longrightarrow> dependent S"
     using x dependent_def by auto
-  moreover
-  {
-    assume "x = 0"
-    then have "0 \<in> affine hull S"
-      using x hull_mono[of "S - {0}" S] by auto
-    then have "dependent S"
-      using affine_hull_0_dependent by auto
-  }
+  moreover have "dependent S" if "x = 0"
+    by (metis that affine_hull_0_dependent Diff_insert_absorb dependent_zero x)
   ultimately show ?thesis by auto
 qed
 
@@ -945,60 +801,45 @@
 lemma affine_dependent_iff_dependent2:
   assumes "a \<in> S"
   shows "affine_dependent S \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` (S-{a}))"
-proof -
-  have "insert a (S - {a}) = S"
-    using assms by auto
-  then show ?thesis
-    using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto
-qed
+  by (metis Diff_iff affine_dependent_iff_dependent assms insert_Diff singletonI)
 
 lemma affine_hull_insert_span_gen:
-  "affine hull (insert a s) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` s)"
+  "affine hull (insert a S) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` S)"
 proof -
-  have h1: "{x - a |x. x \<in> s} = ((\<lambda>x. -a+x) ` s)"
+  have h1: "{x - a |x. x \<in> S} = ((\<lambda>x. -a+x) ` S)"
     by auto
   {
-    assume "a \<notin> s"
+    assume "a \<notin> S"
     then have ?thesis
-      using affine_hull_insert_span[of a s] h1 by auto
+      using affine_hull_insert_span[of a S] h1 by auto
   }
   moreover
   {
-    assume a1: "a \<in> s"
-    have "\<exists>x. x \<in> s \<and> -a+x=0"
-      apply (rule exI[of _ a])
-      using a1
-      apply auto
-      done
-    then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
+    assume a1: "a \<in> S"
+    then have "insert 0 ((\<lambda>x. -a+x) ` (S - {a})) = (\<lambda>x. -a+x) ` S"
       by auto
-    then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)"
-      using span_insert_0[of "(+) (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
-    moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
+    then have "span ((\<lambda>x. -a+x) ` (S - {a})) = span ((\<lambda>x. -a+x) ` S)"
+      using span_insert_0[of "(+) (- a) ` (S - {a})"]
+      by presburger
+    moreover have "{x - a |x. x \<in> (S - {a})} = ((\<lambda>x. -a+x) ` (S - {a}))"
       by auto
-    moreover have "insert a (s - {a}) = insert a s"
+    moreover have "insert a (S - {a}) = insert a S"
       by auto
     ultimately have ?thesis
-      using affine_hull_insert_span[of "a" "s-{a}"] by auto
+      using affine_hull_insert_span[of "a" "S-{a}"] by auto
   }
   ultimately show ?thesis by auto
 qed
 
 lemma affine_hull_span2:
-  assumes "a \<in> s"
-  shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (s-{a}))"
-  using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]]
-  by auto
+  assumes "a \<in> S"
+  shows "affine hull S = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (S-{a}))"
+  by (metis affine_hull_insert_span_gen assms insert_Diff)
 
 lemma affine_hull_span_gen:
-  assumes "a \<in> affine hull s"
-  shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` s)"
-proof -
-  have "affine hull (insert a s) = affine hull s"
-    using hull_redundant[of a affine s] assms by auto
-  then show ?thesis
-    using affine_hull_insert_span_gen[of a "s"] by auto
-qed
+  assumes "a \<in> affine hull S"
+  shows "affine hull S = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` S)"
+  by (metis affine_hull_insert_span_gen assms hull_redundant)
 
 lemma affine_hull_span_0:
   assumes "0 \<in> affine hull S"
@@ -1044,29 +885,13 @@
 lemma affine_basis_exists:
   fixes V :: "'n::real_vector set"
   shows "\<exists>B. B \<subseteq> V \<and> \<not> affine_dependent B \<and> affine hull V = affine hull B"
-proof (cases "V = {}")
-  case True
-  then show ?thesis
-    using affine_independent_0 by auto
-next
-  case False
-  then obtain x where "x \<in> V" by auto
-  then show ?thesis
-    using affine_dependent_def[of "{x}"] extend_to_affine_basis_nonempty[of "{x}" V]
-    by auto
-qed
+  by (metis affine_dependent_def affine_independent_1 empty_subsetI extend_to_affine_basis_nonempty insert_subset order_refl)
 
 proposition extend_to_affine_basis:
   fixes S V :: "'n::real_vector set"
   assumes "\<not> affine_dependent S" "S \<subseteq> V"
   obtains T where "\<not> affine_dependent T" "S \<subseteq> T" "T \<subseteq> V" "affine hull T = affine hull V"
-proof (cases "S = {}")
-  case True then show ?thesis
-    using affine_basis_exists by (metis empty_subsetI that)
-next
-  case False
-  then show ?thesis by (metis assms extend_to_affine_basis_nonempty that)
-qed
+  by (metis affine_basis_exists assms(1) assms(2) bot.extremum extend_to_affine_basis_nonempty)
 
 
 subsection \<open>Affine Dimension of a Set\<close>
@@ -1095,7 +920,7 @@
 by (metis affine_empty subset_empty subset_hull)
 
 lemma empty_eq_affine_hull[simp]: "{} = affine hull S \<longleftrightarrow> S = {}"
-by (metis affine_hull_eq_empty)
+  by (metis affine_hull_eq_empty)
 
 lemma aff_dim_parallel_subspace_aux:
   fixes B :: "'n::euclidean_space set"
@@ -1146,7 +971,7 @@
   define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
   moreover have "affine_parallel (affine hull B) Lb"
     using Lb_def B assms affine_hull_span2[of a B] a
-      affine_parallel_commut[of "Lb" "(affine hull B)"]
+      affine_parallel_commute[of "Lb" "(affine hull B)"]
     unfolding affine_parallel_def
     by auto
   moreover have "subspace Lb"
@@ -1168,15 +993,7 @@
   fixes B :: "'n::euclidean_space set"
   assumes "\<not> affine_dependent B"
   shows "finite B"
-proof -
-  {
-    assume "B \<noteq> {}"
-    then obtain a where "a \<in> B" by auto
-    then have ?thesis
-      using aff_dim_parallel_subspace_aux assms by auto
-  }
-  then show ?thesis by auto
-qed
+  using aff_dim_parallel_subspace_aux assms finite.simps by fastforce
 
 
 lemma aff_dim_empty:
@@ -1195,7 +1012,7 @@
 qed
 
 lemma aff_dim_empty_eq [simp]: "aff_dim ({}::'a::euclidean_space set) = -1"
-  by (simp add: aff_dim_empty [symmetric])
+  using aff_dim_empty by blast
 
 lemma aff_dim_affine_hull [simp]: "aff_dim (affine hull S) = aff_dim S"
   unfolding aff_dim_def using hull_hull[of _ S] by auto
@@ -1224,7 +1041,7 @@
   define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
   have "affine_parallel (affine hull B) Lb"
     using Lb_def affine_hull_span2[of a B] a
-      affine_parallel_commut[of "Lb" "(affine hull B)"]
+      affine_parallel_commute[of "Lb" "(affine hull B)"]
     unfolding affine_parallel_def by auto
   moreover have "subspace Lb"
     using Lb_def subspace_span by auto
@@ -1245,11 +1062,14 @@
   using aff_dim_unique[of B B] assms by auto
 
 lemma affine_independent_iff_card:
-    fixes s :: "'a::euclidean_space set"
-    shows "\<not> affine_dependent s \<longleftrightarrow> finite s \<and> aff_dim s = int(card s) - 1"
-  apply (rule iffI)
-  apply (simp add: aff_dim_affine_independent aff_independent_finite)
-  by (metis affine_basis_exists [of s] aff_dim_unique card_subset_eq diff_add_cancel of_nat_eq_iff)
+    fixes S :: "'a::euclidean_space set"
+    shows "\<not> affine_dependent S \<longleftrightarrow> finite S \<and> aff_dim S = int(card S) - 1" (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs" 
+    by (simp add: aff_dim_affine_independent aff_independent_finite)
+  show "?rhs \<Longrightarrow> ?lhs" 
+    by (metis of_nat_eq_iff affine_basis_exists aff_dim_unique card_subset_eq diff_add_cancel)
+qed
 
 lemma aff_dim_sing [simp]:
   fixes a :: "'n::euclidean_space"
@@ -1272,78 +1092,39 @@
   fixes V :: "('n::euclidean_space) set"
   shows "\<exists>B. B \<subseteq> V \<and> affine hull B = affine hull V \<and>
     \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1"
-proof -
-  obtain B where B: "\<not> affine_dependent B" "B \<subseteq> V" "affine hull B = affine hull V"
-    using affine_basis_exists[of V] by auto
-  then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto
-  with B show ?thesis by auto
-qed
+  by (metis aff_dim_unique affine_basis_exists)
 
 lemma aff_dim_le_card:
   fixes V :: "'n::euclidean_space set"
   assumes "finite V"
   shows "aff_dim V \<le> of_nat (card V) - 1"
-proof -
-  obtain B where B: "B \<subseteq> V" "of_nat (card B) = aff_dim V + 1"
-    using aff_dim_inner_basis_exists[of V] by auto
-  then have "card B \<le> card V"
-    using assms card_mono by auto
-  with B show ?thesis by auto
+  by (metis aff_dim_inner_basis_exists assms card_mono le_diff_eq of_nat_le_iff)
+
+lemma aff_dim_parallel_le:
+  fixes S T :: "'n::euclidean_space set"
+  assumes "affine_parallel (affine hull S) (affine hull T)"
+  shows "aff_dim S \<le> aff_dim T"
+proof (cases "S={} \<or> T={}")
+  case True
+  then show ?thesis
+    by (smt (verit, best) aff_dim_affine_hull2 affine_hull_empty affine_parallel_def assms empty_is_image)
+next
+  case False
+    then obtain L where L: "subspace L" "affine_parallel (affine hull T) L"
+      by (metis affine_affine_hull affine_hull_eq_empty affine_parallel_subspace)
+    with False show ?thesis
+      by (metis aff_dim_parallel_subspace affine_parallel_assoc assms dual_order.refl)
 qed
 
 lemma aff_dim_parallel_eq:
   fixes S T :: "'n::euclidean_space set"
   assumes "affine_parallel (affine hull S) (affine hull T)"
   shows "aff_dim S = aff_dim T"
-proof -
-  {
-    assume "T \<noteq> {}" "S \<noteq> {}"
-    then obtain L where L: "subspace L \<and> affine_parallel (affine hull T) L"
-      using affine_parallel_subspace[of "affine hull T"]
-        affine_affine_hull[of T]
-      by auto
-    then have "aff_dim T = int (dim L)"
-      using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto
-    moreover have *: "subspace L \<and> affine_parallel (affine hull S) L"
-       using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto
-    moreover from * have "aff_dim S = int (dim L)"
-      using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto
-    ultimately have ?thesis by auto
-  }
-  moreover
-  {
-    assume "S = {}"
-    then have "S = {}" and "T = {}"
-      using assms
-      unfolding affine_parallel_def
-      by auto
-    then have ?thesis using aff_dim_empty by auto
-  }
-  moreover
-  {
-    assume "T = {}"
-    then have "S = {}" and "T = {}"
-      using assms
-      unfolding affine_parallel_def
-      by auto
-    then have ?thesis
-      using aff_dim_empty by auto
-  }
-  ultimately show ?thesis by blast
-qed
+  by (smt (verit, del_insts) aff_dim_parallel_le affine_parallel_commute assms)
 
 lemma aff_dim_translation_eq:
   "aff_dim ((+) a ` S) = aff_dim S" for a :: "'n::euclidean_space"
-proof -
-  have "affine_parallel (affine hull S) (affine hull ((\<lambda>x. a + x) ` S))"
-    unfolding affine_parallel_def
-    apply (rule exI[of _ "a"])
-    using affine_hull_translation[of a S]
-    apply auto
-    done
-  then show ?thesis
-    using aff_dim_parallel_eq[of S "(\<lambda>x. a + x) ` S"] by auto
-qed
+  by (metis aff_dim_parallel_eq affine_hull_translation affine_parallel_def)
 
 lemma aff_dim_translation_eq_subtract:
   "aff_dim ((\<lambda>x. x - a) ` S) = aff_dim S" for a :: "'n::euclidean_space"
@@ -1351,97 +1132,51 @@
 
 lemma aff_dim_affine:
   fixes S L :: "'n::euclidean_space set"
-  assumes "S \<noteq> {}"
-    and "affine S"
-    and "subspace L"
-    and "affine_parallel S L"
+  assumes "affine S" "subspace L" "affine_parallel S L" "S \<noteq> {}"
   shows "aff_dim S = int (dim L)"
-proof -
-  have *: "affine hull S = S"
-    using assms affine_hull_eq[of S] by auto
-  then have "affine_parallel (affine hull S) L"
-    using assms by (simp add: *)
-  then show ?thesis
-    using assms aff_dim_parallel_subspace[of S L] by blast
-qed
+  by (simp add: aff_dim_parallel_subspace assms hull_same)
 
-lemma dim_affine_hull:
+lemma dim_affine_hull [simp]:
   fixes S :: "'n::euclidean_space set"
   shows "dim (affine hull S) = dim S"
-proof -
-  have "dim (affine hull S) \<ge> dim S"
-    using dim_subset by auto
-  moreover have "dim (span S) \<ge> dim (affine hull S)"
-    using dim_subset affine_hull_subset_span by blast
-  moreover have "dim (span S) = dim S"
-    using dim_span by auto
-  ultimately show ?thesis by auto
-qed
+  by (metis affine_hull_subset_span dim_eq_span dim_mono hull_subset span_eq_dim)
 
 lemma aff_dim_subspace:
   fixes S :: "'n::euclidean_space set"
   assumes "subspace S"
   shows "aff_dim S = int (dim S)"
-proof (cases "S={}")
-  case True with assms show ?thesis
-    by (simp add: subspace_affine)
-next
-  case False
-  with aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S] subspace_affine
-  show ?thesis by auto
-qed
+  by (metis aff_dim_affine affine_parallel_subspace assms empty_iff parallel_subspace subspace_affine)
 
 lemma aff_dim_zero:
   fixes S :: "'n::euclidean_space set"
   assumes "0 \<in> affine hull S"
   shows "aff_dim S = int (dim S)"
-proof -
-  have "subspace (affine hull S)"
-    using subspace_affine[of "affine hull S"] affine_affine_hull assms
-    by auto
-  then have "aff_dim (affine hull S) = int (dim (affine hull S))"
-    using assms aff_dim_subspace[of "affine hull S"] by auto
-  then show ?thesis
-    using aff_dim_affine_hull[of S] dim_affine_hull[of S]
-    by auto
-qed
+  by (metis aff_dim_affine_hull aff_dim_subspace affine_hull_span_0 assms dim_span subspace_span)
 
 lemma aff_dim_eq_dim:
-  "aff_dim S = int (dim ((+) (- a) ` S))" if "a \<in> affine hull S"
-    for S :: "'n::euclidean_space set"
-proof -
-  have "0 \<in> affine hull (+) (- a) ` S"
-    unfolding affine_hull_translation
-    using that by (simp add: ac_simps)
-  with aff_dim_zero show ?thesis
-    by (metis aff_dim_translation_eq)
-qed
+  fixes S :: "'n::euclidean_space set"
+  assumes "a \<in> affine hull S"
+  shows "aff_dim S = int (dim ((+) (- a) ` S))" 
+  by (metis ab_left_minus aff_dim_translation_eq aff_dim_zero affine_hull_translation image_eqI assms)
 
 lemma aff_dim_eq_dim_subtract:
-  "aff_dim S = int (dim ((\<lambda>x. x - a) ` S))" if "a \<in> affine hull S"
-    for S :: "'n::euclidean_space set"
-  using aff_dim_eq_dim [of a] that by (simp cong: image_cong_simp)
+  fixes S :: "'n::euclidean_space set"
+  assumes "a \<in> affine hull S"
+  shows "aff_dim S = int (dim ((\<lambda>x. x - a) ` S))"
+  using aff_dim_eq_dim assms by auto
 
 lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
-  using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"]
-    dim_UNIV[where 'a="'n::euclidean_space"]
-  by auto
+  by (simp add: aff_dim_subspace)
 
 lemma aff_dim_geq:
   fixes V :: "'n::euclidean_space set"
   shows "aff_dim V \<ge> -1"
-proof -
-  obtain B where "affine hull B = affine hull V"
-    and "\<not> affine_dependent B"
-    and "int (card B) = aff_dim V + 1"
-    using aff_dim_basis_exists by auto
-  then show ?thesis by auto
-qed
+  by (metis add_le_cancel_right aff_dim_basis_exists diff_self of_nat_0_le_iff uminus_add_conv_diff)
 
 lemma aff_dim_negative_iff [simp]:
   fixes S :: "'n::euclidean_space set"
-  shows "aff_dim S < 0 \<longleftrightarrow>S = {}"
-by (metis aff_dim_empty aff_dim_geq diff_0 eq_iff zle_diff1_eq)
+  shows "aff_dim S < 0 \<longleftrightarrow> S = {}"
+  by (metis aff_dim_empty aff_dim_geq diff_0 eq_iff zle_diff1_eq)
 
 lemma aff_lowdim_subset_hyperplane:
   fixes S :: "'a::euclidean_space set"
@@ -1482,83 +1217,48 @@
   fixes S :: "'a :: euclidean_space set"
   assumes "\<not> affine_dependent S" "a \<in> S"
     shows "card S = dim ((\<lambda>x. x - a) ` S) + 1"
-proof -
-  have non: "\<not> affine_dependent (insert a S)"
-    by (simp add: assms insert_absorb)
-  have "finite S"
-    by (meson assms aff_independent_finite)
-  with \<open>a \<in> S\<close> have "card S \<noteq> 0" by auto
-  moreover have "dim ((\<lambda>x. x - a) ` S) = card S - 1"
-    using aff_dim_eq_dim_subtract aff_dim_unique \<open>a \<in> S\<close> hull_inc insert_absorb non by fastforce
-  ultimately show ?thesis
-    by auto
-qed
+  using aff_dim_affine_independent aff_dim_eq_dim_subtract assms hull_subset by fastforce
 
 lemma independent_card_le_aff_dim:
   fixes B :: "'n::euclidean_space set"
   assumes "B \<subseteq> V"
   assumes "\<not> affine_dependent B"
   shows "int (card B) \<le> aff_dim V + 1"
-proof -
-  obtain T where T: "\<not> affine_dependent T \<and> B \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V"
-    by (metis assms extend_to_affine_basis[of B V])
-  then have "of_nat (card T) = aff_dim V + 1"
-    using aff_dim_unique by auto
-  then show ?thesis
-    using T card_mono[of T B] aff_independent_finite[of T] by auto
-qed
+  by (metis aff_dim_unique aff_independent_finite assms card_mono extend_to_affine_basis of_nat_mono)
 
 lemma aff_dim_subset:
   fixes S T :: "'n::euclidean_space set"
   assumes "S \<subseteq> T"
   shows "aff_dim S \<le> aff_dim T"
-proof -
-  obtain B where B: "\<not> affine_dependent B" "B \<subseteq> S" "affine hull B = affine hull S"
-    "of_nat (card B) = aff_dim S + 1"
-    using aff_dim_inner_basis_exists[of S] by auto
-  then have "int (card B) \<le> aff_dim T + 1"
-    using assms independent_card_le_aff_dim[of B T] by auto
-  with B show ?thesis by auto
-qed
+  by (metis add_le_cancel_right aff_dim_inner_basis_exists assms dual_order.trans independent_card_le_aff_dim)
 
 lemma aff_dim_le_DIM:
   fixes S :: "'n::euclidean_space set"
   shows "aff_dim S \<le> int (DIM('n))"
-proof -
-  have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
-    using aff_dim_UNIV by auto
-  then show "aff_dim (S:: 'n::euclidean_space set) \<le> int(DIM('n))"
-    using aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
-qed
+  by (metis aff_dim_UNIV aff_dim_subset top_greatest)
 
 lemma affine_dim_equal:
   fixes S :: "'n::euclidean_space set"
   assumes "affine S" "affine T" "S \<noteq> {}" "S \<subseteq> T" "aff_dim S = aff_dim T"
   shows "S = T"
 proof -
-  obtain a where "a \<in> S" using assms by auto
-  then have "a \<in> T" using assms by auto
+  obtain a where "a \<in> S" "a \<in> T" "T \<noteq> {}" using assms by auto
   define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}"
   then have ls: "subspace LS" "affine_parallel S LS"
     using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto
   then have h1: "int(dim LS) = aff_dim S"
     using assms aff_dim_affine[of S LS] by auto
-  have "T \<noteq> {}" using assms by auto
   define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}"
   then have lt: "subspace LT \<and> affine_parallel T LT"
     using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto
-  then have "int(dim LT) = aff_dim T"
-    using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto
   then have "dim LS = dim LT"
-    using h1 assms by auto
+    using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close>  h1 by auto
   moreover have "LS \<le> LT"
     using LS_def LT_def assms by auto
   ultimately have "LS = LT"
     using subspace_dim_equal[of LS LT] ls lt by auto
-  moreover have "S = {x. \<exists>y \<in> LS. a+y=x}"
-    using LS_def by auto
-  moreover have "T = {x. \<exists>y \<in> LT. a+y=x}"
-    using LT_def by auto
+  moreover have "S = {x. \<exists>y \<in> LS. a+y=x}" "T = {x. \<exists>y \<in> LT. a+y=x}"
+    using LS_def LT_def by auto
   ultimately show ?thesis by auto
 qed
 
@@ -1566,10 +1266,6 @@
   fixes S :: "'a::euclidean_space set"
   shows "aff_dim S = 0 \<longleftrightarrow> (\<exists>a. S = {a})"
 proof (cases "S = {}")
-  case True
-  then show ?thesis
-    by auto
-next
   case False
   then obtain a where "a \<in> S" by auto
   show ?thesis
@@ -1580,58 +1276,39 @@
     then show "\<exists>a. S = {a}"
       using \<open>a \<in> S\<close> by blast
   qed auto
-qed
+qed auto
 
 lemma affine_hull_UNIV:
   fixes S :: "'n::euclidean_space set"
   assumes "aff_dim S = int(DIM('n))"
   shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
-proof -
-  have "S \<noteq> {}"
-    using assms aff_dim_empty[of S] by auto
-  have h0: "S \<subseteq> affine hull S"
-    using hull_subset[of S _] by auto
-  have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S"
-    using aff_dim_UNIV assms by auto
-  then have h2: "aff_dim (affine hull S) \<le> aff_dim (UNIV :: ('n::euclidean_space) set)"
-    using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto
-  have h3: "aff_dim S \<le> aff_dim (affine hull S)"
-    using h0 aff_dim_subset[of S "affine hull S"] assms by auto
-  then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)"
-    using h0 h1 h2 by auto
-  then show ?thesis
-    using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"]
-      affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close>
-    by auto
-qed
+  by (simp add: aff_dim_empty affine_dim_equal assms)
 
 lemma disjoint_affine_hull:
-  fixes s :: "'n::euclidean_space set"
-  assumes "\<not> affine_dependent s" "t \<subseteq> s" "u \<subseteq> s" "t \<inter> u = {}"
-    shows "(affine hull t) \<inter> (affine hull u) = {}"
+  fixes S :: "'n::euclidean_space set"
+  assumes "\<not> affine_dependent S" "T \<subseteq> S" "U \<subseteq> S" "T \<inter> U = {}"
+    shows "(affine hull T) \<inter> (affine hull U) = {}"
 proof -
-  from assms(1) have "finite s"
-    by (simp add: aff_independent_finite)
-  with assms(2,3) have "finite t" "finite u"
-    by (blast intro: finite_subset)+
-  have False if "y \<in> affine hull t" and "y \<in> affine hull u" for y
+  obtain "finite S" "finite T" "finite U"
+    using assms by (simp add: aff_independent_finite finite_subset)
+  have False if "y \<in> affine hull T" and "y \<in> affine hull U" for y
   proof -
     from that obtain a b
-      where a1 [simp]: "sum a t = 1"
-        and [simp]: "sum (\<lambda>v. a v *\<^sub>R v) t = y"
-        and [simp]: "sum b u = 1" "sum (\<lambda>v. b v *\<^sub>R v) u = y"
-      by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>)
-    define c where "c x = (if x \<in> t then a x else if x \<in> u then -(b x) else 0)" for x
-    from assms(2,3,4) have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u"
-      by auto
-    have "sum c s = 0"
-      by (simp add: c_def comm_monoid_add_class.sum.If_cases \<open>finite s\<close> sum_negf)
-    moreover have "\<not> (\<forall>v\<in>s. c v = 0)"
-      by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def sum.neutral zero_neq_one)
-    moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0"
-      by (simp add: c_def if_smult sum_negf comm_monoid_add_class.sum.If_cases \<open>finite s\<close>)
+      where a1 [simp]: "sum a T = 1"
+        and [simp]: "sum (\<lambda>v. a v *\<^sub>R v) T = y"
+        and [simp]: "sum b U = 1" "sum (\<lambda>v. b v *\<^sub>R v) U = y"
+      by (auto simp: affine_hull_finite \<open>finite T\<close> \<open>finite U\<close>)
+    define c where "c x = (if x \<in> T then a x else if x \<in> U then -(b x) else 0)" for x
+    have [simp]: "S \<inter> T = T" "S \<inter> - T \<inter> U = U"
+      using assms by auto
+    have "sum c S = 0"
+      by (simp add: c_def comm_monoid_add_class.sum.If_cases \<open>finite S\<close> sum_negf)
+    moreover have "\<not> (\<forall>v\<in>S. c v = 0)"
+      by (metis (no_types) IntD1 \<open>S \<inter> T = T\<close> a1 c_def sum.neutral zero_neq_one)
+    moreover have "(\<Sum>v\<in>S. c v *\<^sub>R v) = 0"
+      by (simp add: c_def if_smult sum_negf comm_monoid_add_class.sum.If_cases \<open>finite S\<close>)
     ultimately show ?thesis
-      using assms(1) \<open>finite s\<close> by (auto simp: affine_dependent_explicit)
+      using assms(1) \<open>finite S\<close> by (auto simp: affine_dependent_explicit)
   qed
   then show ?thesis by blast
 qed
--- a/src/HOL/Analysis/Borel_Space.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Borel_Space.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -77,13 +77,7 @@
         using of_rat_dense by blast
       assume * [rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> l < f x"
       from q2 have "real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2)"
-      proof auto
-        fix x assume "x \<in> A" "x < a"
-        with q2 *[of "a - x"] show "f x < real_of_rat q2"
-          apply (auto simp add: dist_real_def not_less)
-          apply (subgoal_tac "f x \<le> f xa")
-          by (auto intro: mono)
-      qed
+        using q2 *[of "a - _"] dist_real_def mono by fastforce
       thus ?thesis by auto
     next
       fix u assume "u > f a"
@@ -91,13 +85,7 @@
         using of_rat_dense by blast
       assume *[rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> u > f x"
       from q2 have "real_of_rat q2 > f a \<and> (\<forall>x\<in>A. x > a \<longrightarrow> f x > real_of_rat q2)"
-      proof auto
-        fix x assume "x \<in> A" "x > a"
-        with q2 *[of "x - a"] show "f x > real_of_rat q2"
-          apply (auto simp add: dist_real_def)
-          apply (subgoal_tac "f x \<ge> f xa")
-          by (auto intro: mono)
-      qed
+        using q2 *[of "_ - a"] dist_real_def mono by fastforce
       thus ?thesis by auto
     qed
   qed
@@ -127,13 +115,8 @@
   fixes A :: "real set"
   assumes "open A" "mono_on A f"
   shows "countable {a\<in>A. \<not>isCont f a}"
-proof -
-  have "{a\<in>A. \<not>isCont f a} = {a\<in>A. \<not>(continuous (at a within A) f)}"
-    by (auto simp add: continuous_within_open [OF _ \<open>open A\<close>])
-  thus ?thesis
-    apply (elim ssubst)
-    by (rule mono_on_ctble_discont, rule assms)
-qed
+  using continuous_within_open [OF _ \<open>open A\<close>] \<open>mono_on A f\<close>
+  by (smt (verit, ccfv_threshold) Collect_cong mono_on_ctble_discont)
 
 lemma mono_ctble_discont:
   fixes f :: "real \<Rightarrow> real"
@@ -144,8 +127,7 @@
 lemma has_real_derivative_imp_continuous_on:
   assumes "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)"
   shows "continuous_on A f"
-  apply (intro differentiable_imp_continuous_on, unfold differentiable_on_def)
-  using assms differentiable_at_withinI real_differentiable_def by blast
+  by (meson DERIV_isCont assms continuous_at_imp_continuous_on)
 
 lemma continuous_interval_vimage_Int:
   assumes "continuous_on {a::real..b} g" and mono: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
@@ -173,14 +155,11 @@
              intro!: mono)
   moreover have "c' \<le> d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto
   moreover have "g c' \<le> c" "g d' \<ge> d"
-    apply (insert c'' d'' c'd'_in_set)
-    apply (subst c''(2)[symmetric])
-    apply (auto simp: c'_def intro!: mono cInf_lower c'') []
-    apply (subst d''(2)[symmetric])
-    apply (auto simp: d'_def intro!: mono cSup_upper d'') []
-    done
-  with c'd'_in_set have "g c' = c" "g d' = d" by auto
-  ultimately show ?thesis using that by blast
+    using c'' d'' calculation by (metis IntE atLeastAtMost_iff mono order_class.order_eq_iff)+
+  with c'd'_in_set have "g c' = c" "g d' = d" 
+    by auto
+  ultimately show ?thesis 
+    using that by blast
 qed
 
 subsection \<open>Generic Borel spaces\<close>
@@ -196,9 +175,7 @@
   by (auto simp add: measurable_def borel_def)
 
 lemma in_borel_measurable_borel:
-   "f \<in> borel_measurable M \<longleftrightarrow>
-    (\<forall>S \<in> sets borel.
-      f -` S \<inter> space M \<in> sets M)"
+   "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>S \<in> sets borel. f -` S \<inter> space M \<in> sets M)"
   by (auto simp add: measurable_def borel_def)
 
 lemma space_borel[simp]: "space borel = UNIV"
@@ -219,10 +196,7 @@
 
 lemma borel_open[measurable (raw generic)]:
   assumes "open A" shows "A \<in> sets borel"
-proof -
-  have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
-  thus ?thesis unfolding borel_def by auto
-qed
+  by (simp add: assms sets_borel)
 
 lemma borel_closed[measurable (raw generic)]:
   assumes "closed A" shows "A \<in> sets borel"
@@ -237,12 +211,7 @@
   unfolding insert_def by (rule sets.Un) auto
 
 lemma sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
-proof -
-  have "(\<Union>a\<in>A. {a}) \<in> sets borel" for A :: "'a set"
-    by (intro sets.countable_UN') auto
-  then show ?thesis
-    by auto
-qed
+  by (simp add: set_eq_iff sets.countable)
 
 lemma borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
   unfolding Compl_eq_Diff_UNIV by simp
@@ -312,7 +281,7 @@
   shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
     (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M"
   by (subst measurable_restrict_space_iff)
-     (auto simp: indicator_def of_bool_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong)
+     (auto simp: indicator_def of_bool_def if_distrib[where f="\<lambda>x. a * x" for a] cong: if_cong)
 
 lemma borel_measurable_restrict_space_iff_ennreal:
   fixes f :: "'a \<Rightarrow> ennreal"
@@ -320,7 +289,7 @@
   shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
     (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M"
   by (subst measurable_restrict_space_iff)
-     (auto simp: indicator_def of_bool_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong)
+     (auto simp: indicator_def of_bool_def if_distrib[where f="\<lambda>x. a * x" for a] cong: if_cong)
 
 lemma borel_measurable_restrict_space_iff:
   fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
@@ -329,7 +298,7 @@
     (\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> borel_measurable M"
   by (subst measurable_restrict_space_iff)
      (auto simp: indicator_def of_bool_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps
-       cong del: if_weak_cong)
+       cong: if_cong)
 
 lemma cbox_borel[measurable]: "cbox a b \<in> sets borel"
   by (auto intro: borel_closed)
@@ -338,7 +307,7 @@
   by (auto intro: borel_open)
 
 lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
-  by (auto intro: borel_closed dest!: compact_imp_closed)
+  by (simp add: borel_closed compact_imp_closed)
 
 lemma borel_sigma_sets_subset:
   "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
@@ -360,7 +329,7 @@
     using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto
   finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (F`A)" .
   show "sigma_sets UNIV (F`A) \<subseteq> sigma_sets UNIV {S. open S}"
-    unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
+    by (metis F image_subset_iff sets_borel sigma_sets_mono)
 qed auto
 
 lemma borel_eq_sigmaI2:
@@ -371,7 +340,7 @@
   assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"
   shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
   using assms
-  by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` B" and F="(\<lambda>(i, j). F i j)"]) auto
+  by (smt (verit, del_insts) borel_eq_sigmaI1 image_iff prod.collapse split_beta)
 
 lemma borel_eq_sigmaI3:
   fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
@@ -443,20 +412,15 @@
 qed (auto simp: eq intro: generate_topology.Basis)
 
 lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)"
-  unfolding borel_def
-proof (intro sigma_eqI sigma_sets_eqI, safe)
-  fix x :: "'a set" assume "open x"
-  hence "x = UNIV - (UNIV - x)" by auto
-  also have "\<dots> \<in> sigma_sets UNIV (Collect closed)"
-    by (force intro: sigma_sets.Compl simp: \<open>open x\<close>)
-  finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp
-next
-  fix x :: "'a set" assume "closed x"
-  hence "x = UNIV - (UNIV - x)" by auto
-  also have "\<dots> \<in> sigma_sets UNIV (Collect open)"
-    by (force intro: sigma_sets.Compl simp: \<open>closed x\<close>)
-  finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
-qed simp_all
+proof -
+  have "x \<in> sigma_sets UNIV (Collect closed)" 
+     if  "open x" for x :: "'a set"
+    by (metis that Compl_eq_Diff_UNIV closed_Compl double_complement mem_Collect_eq 
+        sigma_sets.Basic sigma_sets.Compl)
+  then show ?thesis
+    unfolding borel_def
+    by (metis Pow_UNIV borel_closed mem_Collect_eq sets_borel sigma_eqI sigma_sets_eqI top_greatest)
+qed
 
 proposition borel_eq_countable_basis:
   fixes B::"'a::topological_space set set"
@@ -517,8 +481,7 @@
 lemma borel_measurable_continuous_on_indicator:
   fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
   shows "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel"
-  by (subst borel_measurable_restrict_space_iff[symmetric])
-     (auto intro: borel_measurable_continuous_on_restrict)
+  using borel_measurable_continuous_on_restrict borel_measurable_restrict_space_iff inf_top.right_neutral by blast
 
 lemma borel_measurable_Pair[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
@@ -720,7 +683,7 @@
   shows "(\<lambda>x. SUP i\<in>I. F i x) \<in> borel_measurable M"
 proof cases
   assume "I = {}" then show ?thesis
-    unfolding \<open>I = {}\<close> image_empty by simp
+    by (simp add: borel_measurable_const)
 next
   assume "I \<noteq> {}"
   show ?thesis
@@ -742,7 +705,7 @@
   shows "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable M"
 proof cases
   assume "I = {}" then show ?thesis
-    unfolding \<open>I = {}\<close> image_empty by simp
+    by (simp add: borel_measurable_const)
 next
   assume "I \<noteq> {}"
   show ?thesis
@@ -1045,25 +1008,23 @@
 proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
   fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
   then have i: "i \<in> Basis" by auto
-  have "{x::'a. x\<bullet>i \<le> a} = UNIV - {x::'a. a < x\<bullet>i}" by auto
-  also have *: "{x::'a. a < x\<bullet>i} =
-      (\<Union>k::nat. {x. (\<Sum>n\<in>Basis. (if n = i then a else -real k) *\<^sub>R n) <e x})" using i
-  proof (safe, simp_all add: eucl_less_def split: if_split_asm)
-    fix x :: 'a
+  have **: "\<exists>y. \<forall>j\<in>Basis. j \<noteq> i \<longrightarrow> - real y < x \<bullet> j" if "a < x \<bullet> i" for x
+  proof -
     obtain k where k: "Max ((\<bullet>) (- x) ` Basis) < real k"
       using reals_Archimedean2 by blast
     { fix i :: 'a assume "i \<in> Basis"
       then have "-x\<bullet>i < real k"
         using k by (subst (asm) Max_less_iff) auto
       then have "- real k < x\<bullet>i" by simp }
-    then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> -real k < x \<bullet> ia"
+    then show ?thesis
       by (auto intro!: exI[of _ k])
   qed
-  finally show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA"
-    apply (simp only:)
-    apply (intro sets.countable_UN sets.Diff)
-    apply (auto intro: sigma_sets_top)
-    done
+  have "{x::'a. x\<bullet>i \<le> a} = UNIV - {x::'a. a < x\<bullet>i}" by auto
+  also have *: "{x::'a. a < x\<bullet>i} = (\<Union>k::nat. {x. (\<Sum>n\<in>Basis. (if n = i then a else -k) *\<^sub>R n) <e x})" 
+    using i ** by (force simp add: eucl_less_def split: if_split_asm)
+  finally have eq: "{x. x \<bullet> i \<le> a} = UNIV - (\<Union>x. {xa. (\<Sum>n\<in>Basis. (if n = i then a else - real x) *\<^sub>R n) <e xa})" .
+  show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA"
+    unfolding eq by (fastforce intro!: sigma_sets_top sets.Diff)
 qed auto
 
 lemma borel_eq_lessThan:
@@ -1072,24 +1033,26 @@
 proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
   fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
   then have i: "i \<in> Basis" by auto
-  have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto
-  also have *: "{x::'a. x\<bullet>i < a} = (\<Union>k::nat. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" using \<open>i\<in> Basis\<close>
-  proof (safe, simp_all add: eucl_less_def split: if_split_asm)
-    fix x :: 'a
+  have **: "\<exists>y. \<forall>j\<in>Basis. j \<noteq> i \<longrightarrow> real y > x \<bullet> j" if "a > x \<bullet> i" for x
+  proof -
     obtain k where k: "Max ((\<bullet>) x ` Basis) < real k"
       using reals_Archimedean2 by blast
     { fix i :: 'a assume "i \<in> Basis"
       then have "x\<bullet>i < real k"
         using k by (subst (asm) Max_less_iff) auto
       then have "x\<bullet>i < real k" by simp }
-    then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia < real k"
+    then show ?thesis
       by (auto intro!: exI[of _ k])
   qed
-  finally show "{x. a \<le> x\<bullet>i} \<in> ?SIGMA"
-    apply (simp only:)
-    apply (intro sets.countable_UN sets.Diff)
-    apply (auto intro: sigma_sets_top )
-    done
+  have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto
+  also have *: "{x::'a. x\<bullet>i < a} = (\<Union>k::nat. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" using \<open>i\<in> Basis\<close>
+    using i ** by (force simp add: eucl_less_def split: if_split_asm)
+  finally
+  have eq: "{x. a \<le> x \<bullet> i} =
+            UNIV - (\<Union>k. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" .
+
+  show "{x. a \<le> x\<bullet>i} \<in> ?SIGMA"
+    unfolding eq by (fastforce intro!: sigma_sets_top sets.Diff)
 qed auto
 
 lemma borel_eq_atLeastAtMost:
@@ -1276,9 +1239,7 @@
 lemma borel_measurable_uminus_eq [simp]:
   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
   shows "(\<lambda>x. - f x) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
-proof
-  assume ?l from borel_measurable_uminus[OF this] show ?r by simp
-qed auto
+  by (smt (verit, ccfv_SIG) borel_measurable_uminus equation_minus_iff measurable_cong)
 
 lemma affine_borel_measurable_vector:
   fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
@@ -1294,8 +1255,10 @@
       by (auto simp: algebra_simps)
     hence "?S \<in> sets borel" by auto
     moreover
-    from \<open>b \<noteq> 0\<close> have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
-      apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all)
+    have "\<And>x. \<lbrakk>a + b *\<^sub>R f x \<in> S\<rbrakk> \<Longrightarrow> f x \<in> (\<lambda>x. (x - a) /\<^sub>R b) ` S"
+      using \<open>b \<noteq> 0\<close> image_iff by fastforce
+    with \<open>b \<noteq> 0\<close> have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
+      by auto
     ultimately show ?thesis using assms unfolding in_borel_measurable_borel
       by auto
   qed simp
@@ -1313,10 +1276,12 @@
   fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra"
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
-  apply (rule measurable_compose[OF f])
-  apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"])
-  apply (auto intro!: continuous_on_inverse continuous_on_id)
-  done
+proof -
+  have "countable {0::'b}" "continuous_on (- {0::'b}) inverse"
+    by (auto intro!: continuous_on_inverse continuous_on_id)
+  then show ?thesis
+    by (metis borel_measurable_continuous_countable_exceptions f measurable_compose)
+qed
 
 lemma borel_measurable_divide[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
@@ -1341,10 +1306,8 @@
 lemma borel_measurable_ln[measurable (raw)]:
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. ln (f x :: real)) \<in> borel_measurable M"
-  apply (rule measurable_compose[OF f])
-  apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"])
-  apply (auto intro!: continuous_on_ln continuous_on_id)
-  done
+  using borel_measurable_continuous_countable_exceptions[of "{0}"] measurable_compose[OF f]
+  by (auto intro!: continuous_on_ln continuous_on_id)
 
 lemma borel_measurable_log[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M"
@@ -1402,12 +1365,16 @@
 
 lemma\<^marker>\<open>tag important\<close> borel_measurable_complex_iff:
   "f \<in> borel_measurable M \<longleftrightarrow>
-    (\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M"
-  apply auto
-  apply (subst fun_complex_eq)
-  apply (intro borel_measurable_add)
-  apply auto
-  done
+    (\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    using borel_measurable_Im borel_measurable_Re measurable_compose by blast
+  assume R: ?rhs
+  then have "(\<lambda>x. complex_of_real (Re (f x)) + \<i> * complex_of_real (Im (f x))) \<in> borel_measurable M"
+    by (intro borel_measurable_add) auto
+  then show ?lhs
+    using complex_eq by force
+qed
 
 lemma powr_real_measurable [measurable]:
   assumes "f \<in> measurable M borel" "g \<in> measurable M borel"
@@ -1427,10 +1394,8 @@
   fixes f :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M"
-  apply (rule measurable_compose[OF f])
-  apply (rule borel_measurable_continuous_countable_exceptions[of "{\<infinity>, -\<infinity> }"])
-  apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV)
-  done
+  using measurable_compose[OF f] borel_measurable_continuous_countable_exceptions[of "{\<infinity>, -\<infinity> }"]
+  by (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV)
 
 lemma borel_measurable_ereal_cases:
   fixes f :: "'a \<Rightarrow> ereal"
@@ -1451,10 +1416,8 @@
   by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If)
 
 lemma borel_measurable_uminus_eq_ereal[simp]:
-  "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
-proof
-  assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp
-qed auto
+  "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
+  by (smt (verit, ccfv_SIG) borel_measurable_uminus_ereal ereal_uminus_uminus measurable_cong)
 
 lemma set_Collect_ereal2:
   fixes f g :: "'a \<Rightarrow> ereal"
@@ -1511,13 +1474,7 @@
 
 lemma vimage_sets_compl_iff:
   "f -` A \<inter> space M \<in> sets M \<longleftrightarrow> f -` (- A) \<inter> space M \<in> sets M"
-proof -
-  { fix A assume "f -` A \<inter> space M \<in> sets M"
-    moreover have "f -` (- A) \<inter> space M = space M - f -` A \<inter> space M" by auto
-    ultimately have "f -` (- A) \<inter> space M \<in> sets M" by auto }
-  from this[of A] this[of "-A"] show ?thesis
-    by (metis double_complement)
-qed
+  by (metis Diff_Compl Diff_Diff_Int diff_eq inf_aci(1) sets.Diff sets.top vimage_Compl)
 
 lemma borel_measurable_iff_Iic_ereal:
   "(f::'a\<Rightarrow>ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
@@ -1741,7 +1698,7 @@
 lemma Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
   by (simp add: pred_def)
 
-(* Proof by Jeremy Avigad and Luke Serafin *)
+text \<open>Proof by Jeremy Avigad and Luke Serafin\<close>
 lemma isCont_borel_pred[measurable]:
   fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
   shows "Measurable.pred borel (isCont f)"
@@ -1821,12 +1778,19 @@
   fixes f :: "real \<Rightarrow> real" and A :: "real set"
   assumes "mono_on A f"
   shows "f \<in> borel_measurable (restrict_space borel A)"
-  apply (rule measurable_restrict_countable[OF mono_on_ctble_discont[OF assms]])
-  apply (auto intro!: image_eqI[where x="{x}" for x] simp: sets_restrict_space)
-  apply (auto simp add: sets_restrict_restrict_space continuous_on_eq_continuous_within
-              cong: measurable_cong_sets
-              intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset)
-  done
+proof -
+  have "\<And>x. x \<in> A \<Longrightarrow> {x} \<in> sets (restrict_space borel A)"
+    using sets_restrict_space by fastforce
+  moreover
+  have "continuous_on (A \<inter> - {a \<in> A. \<not> continuous (at a within A) f}) f"
+    by (force simp: continuous_on_eq_continuous_within intro: continuous_within_subset)
+  then have "f \<in> borel_measurable (restrict_space (restrict_space borel A) 
+              (- {a \<in> A. \<not> continuous (at a within A) f}))"
+    by (smt (verit, best) borel_measurable_continuous_on_restrict measurable_cong_sets sets_restrict_restrict_space)
+  ultimately show ?thesis
+    using measurable_restrict_countable[OF mono_on_ctble_discont[OF assms]]
+    by (smt (verit, del_insts) UNIV_I mem_Collect_eq space_borel)
+qed
 
 lemma borel_measurable_piecewise_mono:
   fixes f::"real \<Rightarrow> real" and C::"real set set"
@@ -1920,7 +1884,6 @@
   fixes f g::"_\<Rightarrow> 'a::{second_countable_topology, t2_space}"
   assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "{x \<in> space M. f x = g x} \<in> sets M"
-
 proof -
   define A where "A = {x \<in> space M. f x = g x}"
   define B where "B = {y. \<exists>x::'a. y = (x,x)}"
@@ -1931,33 +1894,16 @@
   then show ?thesis unfolding A_def by simp
 qed
 
-lemma measurable_inequality_set [measurable]:
+text \<open>Logically equivalent to those with the opposite orientation, still these are needed\<close>
+lemma measurable_inequality_set_flipped:
   fixes f g::"_ \<Rightarrow> 'a::{second_countable_topology, linorder_topology}"
   assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
-        "{x \<in> space M. f x < g x} \<in> sets M"
-        "{x \<in> space M. f x \<ge> g x} \<in> sets M"
+  shows "{x \<in> space M. f x \<ge> g x} \<in> sets M"
         "{x \<in> space M. f x > g x} \<in> sets M"
-proof -
-  define F where "F = (\<lambda>x. (f x, g x))"
-  have * [measurable]: "F \<in> borel_measurable M" unfolding F_def by simp
-
-  have "{x \<in> space M. f x \<le> g x} = F-`{(x, y) | x y. x \<le> y} \<inter> space M" unfolding F_def by auto
-  moreover have "{(x, y) | x y. x \<le> (y::'a)} \<in> sets borel" using closed_subdiagonal borel_closed by blast
-  ultimately show "{x \<in> space M. f x \<le> g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
+  by auto
 
-  have "{x \<in> space M. f x < g x} = F-`{(x, y) | x y. x < y} \<inter> space M" unfolding F_def by auto
-  moreover have "{(x, y) | x y. x < (y::'a)} \<in> sets borel" using open_subdiagonal borel_open by blast
-  ultimately show "{x \<in> space M. f x < g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
-
-  have "{x \<in> space M. f x \<ge> g x} = F-`{(x, y) | x y. x \<ge> y} \<inter> space M" unfolding F_def by auto
-  moreover have "{(x, y) | x y. x \<ge> (y::'a)} \<in> sets borel" using closed_superdiagonal borel_closed by blast
-  ultimately show "{x \<in> space M. f x \<ge> g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
-
-  have "{x \<in> space M. f x > g x} = F-`{(x, y) | x y. x > y} \<inter> space M" unfolding F_def by auto
-  moreover have "{(x, y) | x y. x > (y::'a)} \<in> sets borel" using open_superdiagonal borel_open by blast
-  ultimately show "{x \<in> space M. f x > g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
-qed
+lemmas measurable_inequality_set [measurable] =
+  borel_measurable_le borel_measurable_less measurable_inequality_set_flipped
 
 proposition measurable_limit [measurable]:
   fixes f::"nat \<Rightarrow> 'a \<Rightarrow> 'b::first_countable_topology"
@@ -2051,8 +1997,8 @@
 lemma measurable_restrict_mono:
   assumes f: "f \<in> restrict_space M A \<rightarrow>\<^sub>M N" and "B \<subseteq> A"
   shows "f \<in> restrict_space M B \<rightarrow>\<^sub>M N"
-by (rule measurable_compose[OF measurable_restrict_space3 f])
-   (insert \<open>B \<subseteq> A\<close>, auto)
+  by (rule measurable_compose[OF measurable_restrict_space3 f]) (use \<open>B \<subseteq> A\<close> in auto)
+
 
 text \<open>The next one is a variation around \<open>measurable_piecewise_restrict\<close>.\<close>
 
@@ -2065,22 +2011,27 @@
   fix B assume [measurable]: "B \<in> sets N"
   {
     fix n::nat
-    obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" using assms(3) by blast
+    obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" 
+      using assms(3) by blast
     then have *: "f-`B \<inter> A n = h-`B \<inter> A n" by auto
-    have "h-`B \<inter> A n = h-`B \<inter> space M \<inter> A n" using assms(2) sets.sets_into_space by auto
-    then have "h-`B \<inter> A n \<in> sets M" by simp
-    then have "f-`B \<inter> A n \<in> sets M" using * by simp
+    have "h-`B \<inter> A n = h-`B \<inter> space M \<inter> A n" 
+      using assms(2) sets.sets_into_space by auto
+    then have "f-`B \<inter> A n \<in> sets M"
+      by (simp add: "*")
   }
-  then have "(\<Union>n. f-`B \<inter> A n) \<in> sets M" by measurable
-  moreover have "f-`B \<inter> space M = (\<Union>n. f-`B \<inter> A n)" using assms(2) by blast
+  then have "(\<Union>n. f-`B \<inter> A n) \<in> sets M" 
+    by measurable
+  moreover have "f-`B \<inter> space M = (\<Union>n. f-`B \<inter> A n)" 
+    using assms(2) by blast
   ultimately show "f-`B \<inter> space M \<in> sets M" by simp
 next
   fix x assume "x \<in> space M"
-  then obtain n where "x \<in> A n" using assms(2) by blast
-  obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" using assms(3) by blast
-  then have "f x = h x" using \<open>x \<in> A n\<close> by blast
-  moreover have "h x \<in> space N" by (metis measurable_space \<open>x \<in> space M\<close> \<open>h \<in> measurable M N\<close>)
-  ultimately show "f x \<in> space N" by simp
+  then obtain n where "x \<in> A n" 
+    using assms(2) by blast
+  obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" 
+    using assms(3) by blast
+  then show "f x \<in> space N"
+    by (metis \<open>x \<in> A n\<close> \<open>x \<in> space M\<close> measurable_space)
 qed
 
 end
--- a/src/HOL/Analysis/Complex_Analysis_Basics.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -138,25 +138,23 @@
   fixes l::complex
   assumes "(f \<longlongrightarrow> l) F" and "\<not> trivial_limit F" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>"
   shows  "l \<in> \<real>"
-proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)])
-  show "eventually (\<lambda>x. f x \<in> \<real>) F"
-    using assms(3, 4) by (auto intro: eventually_mono)
-qed
+  using Lim_in_closed_set[OF closed_complex_Reals] assms
+  by (smt (verit) eventually_mono)
 
 lemma real_lim_sequentially:
   fixes l::complex
   shows "(f \<longlongrightarrow> l) sequentially \<Longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"
-by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially)
+  by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially)
 
 lemma real_series:
   fixes l::complex
   shows "f sums l \<Longrightarrow> (\<And>n. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"
-unfolding sums_def
-by (metis real_lim_sequentially sum_in_Reals)
+  unfolding sums_def
+  by (metis real_lim_sequentially sum_in_Reals)
 
 lemma Lim_null_comparison_Re:
   assumes "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F" "(g \<longlongrightarrow> 0) F" shows "(f \<longlongrightarrow> 0) F"
-  by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp
+  using Lim_null_comparison assms tendsto_Re by fastforce
 
 subsection\<open>Holomorphic functions\<close>
 
@@ -191,25 +189,11 @@
 lemma holomorphic_on_UN_open:
   assumes "\<And>n. n \<in> I \<Longrightarrow> f holomorphic_on A n" "\<And>n. n \<in> I \<Longrightarrow> open (A n)"
   shows   "f holomorphic_on (\<Union>n\<in>I. A n)"
-proof -
-  have "f field_differentiable at z within (\<Union>n\<in>I. A n)" if "z \<in> (\<Union>n\<in>I. A n)" for z
-  proof -
-    from that obtain n where "n \<in> I" "z \<in> A n"
-      by blast
-    hence "f holomorphic_on A n" "open (A n)"
-      by (simp add: assms)+
-    with \<open>z \<in> A n\<close> have "f field_differentiable at z"
-      by (auto simp: holomorphic_on_open field_differentiable_def)
-    thus ?thesis
-      by (meson field_differentiable_at_within)
-  qed
-  thus ?thesis
-    by (auto simp: holomorphic_on_def)
-qed
+  by (metis UN_E assms holomorphic_on_open open_UN)
 
 lemma holomorphic_on_imp_continuous_on:
     "f holomorphic_on s \<Longrightarrow> continuous_on s f"
-  by (metis field_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def)
+  using differentiable_imp_continuous_on holomorphic_on_imp_differentiable_on by blast
 
 lemma holomorphic_closedin_preimage_constant:
   assumes "f holomorphic_on D" 
@@ -247,22 +231,15 @@
 lemma constant_on_imp_holomorphic_on:
   assumes "f constant_on A"
   shows   "f holomorphic_on A"
-proof -
-  from assms obtain c where c: "\<forall>x\<in>A. f x = c"
-    unfolding constant_on_def by blast
-  have "f holomorphic_on A \<longleftrightarrow> (\<lambda>_. c) holomorphic_on A"
-    by (intro holomorphic_cong) (use c in auto)
-  thus ?thesis
-    by simp
-qed
+  by (metis assms constant_on_def holomorphic_on_const holomorphic_transform)
 
 lemma holomorphic_on_compose:
-  "f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s) \<Longrightarrow> (g o f) holomorphic_on s"
+  "f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s) \<Longrightarrow> (g \<circ> f) holomorphic_on s"
   using field_differentiable_compose_within[of f _ s g]
   by (auto simp: holomorphic_on_def)
 
 lemma holomorphic_on_compose_gen:
-  "f holomorphic_on s \<Longrightarrow> g holomorphic_on t \<Longrightarrow> f ` s \<subseteq> t \<Longrightarrow> (g o f) holomorphic_on s"
+  "f holomorphic_on s \<Longrightarrow> g holomorphic_on t \<Longrightarrow> f ` s \<subseteq> t \<Longrightarrow> (g \<circ> f) holomorphic_on s"
   by (metis holomorphic_on_compose holomorphic_on_subset)
 
 lemma holomorphic_on_balls_imp_entire:
@@ -284,15 +261,10 @@
 lemma holomorphic_on_balls_imp_entire':
   assumes "\<And>r. r > 0 \<Longrightarrow> f holomorphic_on ball c r"
   shows   "f holomorphic_on B"
-proof (rule holomorphic_on_balls_imp_entire)
-  {
-    fix M :: real
-    have "\<exists>x. x > max M 0" by (intro gt_ex)
-    hence "\<exists>x>0. x > M" by auto
-  }
-  thus "\<not>bdd_above {(0::real)<..}" unfolding bdd_above_def
-    by (auto simp: not_le)
-qed (insert assms, auto)
+proof (rule holomorphic_on_balls_imp_entire)  
+  show "\<not>bdd_above {(0::real)<..}" unfolding bdd_above_def
+    by (meson greaterThan_iff gt_ex less_le_not_le order_le_less_trans)
+qed (use assms in auto)
 
 lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on A \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on A"
   by (metis field_differentiable_minus holomorphic_on_def)
@@ -357,8 +329,7 @@
 lemma holomorphic_on_Un [holomorphic_intros]:
   assumes "f holomorphic_on A" "f holomorphic_on B" "open A" "open B"
   shows   "f holomorphic_on (A \<union> B)"
-  using assms by (auto simp: holomorphic_on_def  at_within_open[of _ A]
-                             at_within_open[of _ B]  at_within_open[of _ "A \<union> B"] open_Un)
+  by (metis Un_iff assms holomorphic_on_open open_Un)
 
 lemma holomorphic_on_If_Un [holomorphic_intros]:
   assumes "f holomorphic_on A" "g holomorphic_on B" "open A" "open B"
@@ -374,19 +345,16 @@
   also have "g holomorphic_on B \<longleftrightarrow> ?h holomorphic_on B"
     using assms by (intro holomorphic_cong) auto
   finally show \<dots> .
-qed (insert assms, auto)
+qed (use assms in auto)
 
 lemma holomorphic_derivI:
-     "\<lbrakk>f holomorphic_on S; open S; x \<in> S\<rbrakk>
-      \<Longrightarrow> (f has_field_derivative deriv f x) (at x within T)"
-by (metis DERIV_deriv_iff_field_differentiable at_within_open  holomorphic_on_def has_field_derivative_at_within)
+     "\<lbrakk>f holomorphic_on S; open S; x \<in> S\<rbrakk> \<Longrightarrow> (f has_field_derivative deriv f x) (at x within T)"
+  by (metis DERIV_deriv_iff_field_differentiable at_within_open  holomorphic_on_def has_field_derivative_at_within)
 
 lemma complex_derivative_transform_within_open:
   "\<lbrakk>f holomorphic_on s; g holomorphic_on s; open s; z \<in> s; \<And>w. w \<in> s \<Longrightarrow> f w = g w\<rbrakk>
    \<Longrightarrow> deriv f z = deriv g z"
-  unfolding holomorphic_on_def
-  by (rule DERIV_imp_deriv)
-     (metis DERIV_deriv_iff_field_differentiable has_field_derivative_transform_within_open at_within_open)
+  by (smt (verit) DERIV_imp_deriv has_field_derivative_transform_within_open holomorphic_on_open)
 
 lemma holomorphic_on_compose_cnj_cnj:
   assumes "f holomorphic_on cnj ` A" "open A"
@@ -408,13 +376,13 @@
 subsection\<open>Analyticity on a set\<close>
 
 definition\<^marker>\<open>tag important\<close> analytic_on (infixl "(analytic'_on)" 50)
-  where "f analytic_on S \<equiv> \<forall>x \<in> S. \<exists>e. 0 < e \<and> f holomorphic_on (ball x e)"
+  where "f analytic_on S \<equiv> \<forall>x \<in> S. \<exists>\<epsilon>. 0 < \<epsilon> \<and> f holomorphic_on (ball x \<epsilon>)"
 
 named_theorems\<^marker>\<open>tag important\<close> analytic_intros "introduction rules for proving analyticity"
 
 lemma analytic_imp_holomorphic: "f analytic_on S \<Longrightarrow> f holomorphic_on S"
-  by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def)
-     (metis centre_in_ball field_differentiable_at_within)
+  unfolding analytic_on_def holomorphic_on_def
+  using centre_in_ball field_differentiable_at_within field_differentiable_within_open by blast
 
 lemma analytic_on_open: "open S \<Longrightarrow> f analytic_on S \<longleftrightarrow> f holomorphic_on S"
   by (meson analytic_imp_holomorphic analytic_on_def holomorphic_on_subset openE)
@@ -426,15 +394,12 @@
 lemma analytic_at_imp_isCont:
   assumes "f analytic_on {z}"
   shows   "isCont f z"
-  using assms by (meson analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at insertI1)
+  by (meson analytic_on_imp_differentiable_at assms field_differentiable_imp_continuous_at insertCI)
 
 lemma analytic_at_neq_imp_eventually_neq:
   assumes "f analytic_on {x}" "f x \<noteq> c"
   shows   "eventually (\<lambda>y. f y \<noteq> c) (at x)"
-proof (intro tendsto_imp_eventually_ne)
-  show "f \<midarrow>x\<rightarrow> f x"
-    using assms by (simp add: analytic_at_imp_isCont isContD)
-qed (use assms in auto)
+  using analytic_at_imp_isCont assms isContD tendsto_imp_eventually_ne by blast
 
 lemma analytic_on_subset: "f analytic_on S \<Longrightarrow> T \<subseteq> S \<Longrightarrow> f analytic_on T"
   by (auto simp: analytic_on_def)
@@ -455,18 +420,21 @@
   have "?lhs \<longleftrightarrow> (\<exists>T. open T \<and> S \<subseteq> T \<and> f analytic_on T)"
   proof safe
     assume "f analytic_on S"
-    then show "\<exists>T. open T \<and> S \<subseteq> T \<and> f analytic_on T"
-      apply (simp add: analytic_on_def)
-      apply (rule exI [where x="\<Union>{U. open U \<and> f analytic_on U}"], auto)
-      apply (metis open_ball analytic_on_open centre_in_ball)
-      by (metis analytic_on_def)
+    then have "\<forall>x \<in> \<Union>{U. open U \<and> f analytic_on U}. \<exists>\<epsilon>>0. f holomorphic_on ball x \<epsilon>"
+      using analytic_on_def by force
+    moreover have "S \<subseteq> \<Union>{U. open U \<and> f analytic_on U}"
+      using \<open>f analytic_on S\<close>
+      by (smt (verit, best) open_ball Union_iff analytic_on_def analytic_on_open centre_in_ball mem_Collect_eq subsetI)
+    ultimately show "\<exists>T. open T \<and> S \<subseteq> T \<and> f analytic_on T"
+      unfolding analytic_on_def
+      by (metis (mono_tags, lifting) mem_Collect_eq open_Union)
   next
     fix T
     assume "open T" "S \<subseteq> T" "f analytic_on T"
     then show "f analytic_on S"
         by (metis analytic_on_subset)
   qed
-  also have "... \<longleftrightarrow> ?rhs"
+  also have "\<dots> \<longleftrightarrow> ?rhs"
     by (auto simp: analytic_on_open)
   finally show ?thesis .
 qed
@@ -486,7 +454,7 @@
 lemma analytic_on_compose:
   assumes f: "f analytic_on S"
       and g: "g analytic_on (f ` S)"
-    shows "(g o f) analytic_on S"
+    shows "(g \<circ> f) analytic_on S"
 unfolding analytic_on_def
 proof (intro ballI)
   fix x
@@ -500,21 +468,19 @@
   with e' obtain d where d: "0 < d" and fd: "f ` ball x d \<subseteq> ball (f x) e'"
      by (auto simp: continuous_at_ball)
   have "g \<circ> f holomorphic_on ball x (min d e)"
-    apply (rule holomorphic_on_compose)
-    apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
-    by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball)
+    by (meson fd fh gh holomorphic_on_compose_gen holomorphic_on_subset image_mono min.cobounded1 min.cobounded2 subset_ball)
   then show "\<exists>e>0. g \<circ> f holomorphic_on ball x e"
     by (metis d e min_less_iff_conj)
 qed
 
 lemma analytic_on_compose_gen:
   "f analytic_on S \<Longrightarrow> g analytic_on T \<Longrightarrow> (\<And>z. z \<in> S \<Longrightarrow> f z \<in> T)
-             \<Longrightarrow> g o f analytic_on S"
-by (metis analytic_on_compose analytic_on_subset image_subset_iff)
+             \<Longrightarrow> g \<circ> f analytic_on S"
+  by (metis analytic_on_compose analytic_on_subset image_subset_iff)
 
 lemma analytic_on_neg [analytic_intros]:
   "f analytic_on S \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on S"
-by (metis analytic_on_holomorphic holomorphic_on_minus)
+  by (metis analytic_on_holomorphic holomorphic_on_minus)
 
 lemma analytic_on_add [analytic_intros]:
   assumes f: "f analytic_on S"
@@ -529,33 +495,11 @@
   obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
     by (metis analytic_on_def g z)
   have "(\<lambda>z. f z + g z) holomorphic_on ball z (min e e')"
-    apply (rule holomorphic_on_add)
-    apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
-    by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+    by (metis fh gh holomorphic_on_add holomorphic_on_subset linorder_linear min_def subset_ball)
   then show "\<exists>e>0. (\<lambda>z. f z + g z) holomorphic_on ball z e"
     by (metis e e' min_less_iff_conj)
 qed
 
-lemma analytic_on_diff [analytic_intros]:
-  assumes f: "f analytic_on S"
-      and g: "g analytic_on S"
-    shows "(\<lambda>z. f z - g z) analytic_on S"
-unfolding analytic_on_def
-proof (intro ballI)
-  fix z
-  assume z: "z \<in> S"
-  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
-    by (metis analytic_on_def)
-  obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
-    by (metis analytic_on_def g z)
-  have "(\<lambda>z. f z - g z) holomorphic_on ball z (min e e')"
-    apply (rule holomorphic_on_diff)
-    apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
-    by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
-  then show "\<exists>e>0. (\<lambda>z. f z - g z) holomorphic_on ball z e"
-    by (metis e e' min_less_iff_conj)
-qed
-
 lemma analytic_on_mult [analytic_intros]:
   assumes f: "f analytic_on S"
       and g: "g analytic_on S"
@@ -569,13 +513,23 @@
   obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
     by (metis analytic_on_def g z)
   have "(\<lambda>z. f z * g z) holomorphic_on ball z (min e e')"
-    apply (rule holomorphic_on_mult)
-    apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
-    by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+    by (metis fh gh holomorphic_on_mult holomorphic_on_subset min.absorb_iff2 min_def subset_ball)
   then show "\<exists>e>0. (\<lambda>z. f z * g z) holomorphic_on ball z e"
     by (metis e e' min_less_iff_conj)
 qed
 
+lemma analytic_on_diff [analytic_intros]:
+  assumes f: "f analytic_on S" and g: "g analytic_on S"
+  shows "(\<lambda>z. f z - g z) analytic_on S"
+proof -
+  have "(\<lambda>z. - g z) analytic_on S"
+    by (simp add: analytic_on_neg g)
+  then have "(\<lambda>z. f z + - g z) analytic_on S"
+    using analytic_on_add f by blast
+  then show ?thesis
+    by fastforce
+qed
+
 lemma analytic_on_inverse [analytic_intros]:
   assumes f: "f analytic_on S"
       and nz: "(\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0)"
@@ -591,24 +545,20 @@
   then obtain e' where e': "0 < e'" and nz': "\<And>y. dist z y < e' \<Longrightarrow> f y \<noteq> 0"
     by (metis open_ball centre_in_ball continuous_on_open_avoid e z nz)
   have "(\<lambda>z. inverse (f z)) holomorphic_on ball z (min e e')"
-    apply (rule holomorphic_on_inverse)
-    apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball)
-    by (metis nz' mem_ball min_less_iff_conj)
+    using fh holomorphic_on_inverse holomorphic_on_open nz' by fastforce
   then show "\<exists>e>0. (\<lambda>z. inverse (f z)) holomorphic_on ball z e"
     by (metis e e' min_less_iff_conj)
 qed
 
 lemma analytic_on_divide [analytic_intros]:
-  assumes f: "f analytic_on S"
-      and g: "g analytic_on S"
-      and nz: "(\<And>z. z \<in> S \<Longrightarrow> g z \<noteq> 0)"
-    shows "(\<lambda>z. f z / g z) analytic_on S"
-unfolding divide_inverse
-by (metis analytic_on_inverse analytic_on_mult f g nz)
+  assumes f: "f analytic_on S" and g: "g analytic_on S"
+    and nz: "(\<And>z. z \<in> S \<Longrightarrow> g z \<noteq> 0)"
+  shows "(\<lambda>z. f z / g z) analytic_on S"
+  unfolding divide_inverse by (metis analytic_on_inverse analytic_on_mult f g nz)
 
 lemma analytic_on_power [analytic_intros]:
   "f analytic_on S \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on S"
-by (induct n) (auto simp: analytic_on_mult)
+  by (induct n) (auto simp: analytic_on_mult)
 
 lemma analytic_on_power_int [analytic_intros]:
   assumes nz: "n \<ge> 0 \<or> (\<forall>x\<in>A. f x \<noteq> 0)" and f: "f analytic_on A"
@@ -645,15 +595,15 @@
 proof -
   have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w"
     by (simp add: algebra_simps)
-  also have "... = deriv (g o f) w"
+  also have "\<dots> = deriv (g \<circ> f) w"
     using assms
     by (metis analytic_on_imp_differentiable_at analytic_on_open deriv_chain image_subset_iff)
-  also have "... = deriv id w"
+  also have "\<dots> = deriv id w"
   proof (rule complex_derivative_transform_within_open [where s=S])
     show "g \<circ> f holomorphic_on S"
       by (rule assms holomorphic_on_compose_gen holomorphic_intros)+
   qed (use assms in auto)
-  also have "... = 1"
+  also have "\<dots> = 1"
     by simp
   finally show ?thesis .
 qed
@@ -679,19 +629,16 @@
 
 lemma analytic_at_two:
   "f analytic_on {z} \<and> g analytic_on {z} \<longleftrightarrow>
-   (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s \<and> g holomorphic_on s)"
+   (\<exists>S. open S \<and> z \<in> S \<and> f holomorphic_on S \<and> g holomorphic_on S)"
   (is "?lhs = ?rhs")
 proof
   assume ?lhs
-  then obtain s t
-    where st: "open s" "z \<in> s" "f holomorphic_on s"
-              "open t" "z \<in> t" "g holomorphic_on t"
+  then obtain S T
+    where st: "open S" "z \<in> S" "f holomorphic_on S"
+              "open T" "z \<in> T" "g holomorphic_on T"
     by (auto simp: analytic_at)
-  show ?rhs
-    apply (rule_tac x="s \<inter> t" in exI)
-    using st
-    apply (auto simp: holomorphic_on_subset)
-    done
+  then show ?rhs
+    by (metis Int_iff holomorphic_on_subset inf_le1 inf_le2 open_Int)
 next
   assume ?rhs
   then show ?lhs
@@ -707,32 +654,23 @@
     and complex_derivative_mult_at: "deriv (\<lambda>w. f w * g w) z =
            f z * deriv g z + deriv f z * g z"
 proof -
-  obtain s where s: "open s" "z \<in> s" "f holomorphic_on s" "g holomorphic_on s"
-    using assms by (metis analytic_at_two)
   show "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z"
-    apply (rule DERIV_imp_deriv [OF DERIV_add])
-    using s
-    apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable)
-    done
+    using analytic_on_imp_differentiable_at assms by auto
   show "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z"
-    apply (rule DERIV_imp_deriv [OF DERIV_diff])
-    using s
-    apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable)
-    done
-  show "deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z"
-    apply (rule DERIV_imp_deriv [OF DERIV_mult'])
-    using s
-    apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable)
-    done
+    using analytic_on_imp_differentiable_at assms by force
+  obtain S where "open S" "z \<in> S" "f holomorphic_on S" "g holomorphic_on S"
+    using assms by (metis analytic_at_two)
+  then show "deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z"
+    by (simp add: DERIV_imp_deriv [OF DERIV_mult'] holomorphic_derivI)
 qed
 
 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)
+  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)
+  by (auto simp: complex_derivative_mult_at)
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Complex differentiation of sequences and series\<close>
 
@@ -748,27 +686,16 @@
 proof -
   from assms obtain x l where x: "x \<in> S" and tf: "((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially"
     by blast
-  { fix e::real assume e: "e > 0"
-    then obtain N where N: "\<forall>n\<ge>N. \<forall>x. x \<in> S \<longrightarrow> cmod (f' n x - g' x) \<le> e"
-      by (metis conv)
-    have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
-    proof (rule exI [of _ N], clarify)
-      fix n y h
-      assume "N \<le> n" "y \<in> S"
-      then have "cmod (f' n y - g' y) \<le> e"
-        by (metis N)
-      then have "cmod h * cmod (f' n y - g' y) \<le> cmod h * e"
-        by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)
-      then show "cmod (f' n y * h - g' y * h) \<le> e * cmod h"
-        by (simp add: norm_mult [symmetric] field_simps)
-    qed
-  } note ** = this
   show ?thesis
     unfolding has_field_derivative_def
   proof (rule has_derivative_sequence [OF cvs _ _ x])
     show "(\<lambda>n. f n x) \<longlonglongrightarrow> l"
       by (rule tf)
-  next show "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
+  next 
+    have **: "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. cmod (f' n x * h - g' x * h) \<le> \<epsilon> * cmod h"
+      if "\<epsilon> > 0" for \<epsilon>::real 
+      by (metis that left_diff_distrib mult_right_mono norm_ge_zero norm_mult conv)
+    show "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
       unfolding eventually_sequentially by (blast intro: **)
   qed (metis has_field_derivative_def df)
 qed
@@ -784,19 +711,17 @@
 proof -
   from assms obtain x l where x: "x \<in> S" and sf: "((\<lambda>n. f n x) sums l)"
     by blast
-  { fix e::real assume e: "e > 0"
+  { fix \<epsilon>::real assume e: "\<epsilon> > 0"
     then obtain N where N: "\<forall>n x. n \<ge> N \<longrightarrow> x \<in> S
-            \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e"
+            \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> \<epsilon>"
       by (metis conv)
-    have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"
+    have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> \<epsilon> * cmod h"
     proof (rule exI [of _ N], clarify)
       fix n y h
       assume "N \<le> n" "y \<in> S"
-      then have "cmod ((\<Sum>i<n. f' i y) - g' y) \<le> e"
-        by (metis N)
-      then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e"
-        by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)
-      then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h"
+      have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * \<epsilon>"
+        by (simp add: N \<open>N \<le> n\<close> \<open>y \<in> S\<close> mult_le_cancel_left)
+      then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> \<epsilon> * cmod h"
         by (simp add: norm_mult [symmetric] field_simps sum_distrib_left)
     qed
   } note ** = this
@@ -818,8 +743,8 @@
 
 lemma sum_Suc_reindex:
   fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
-    shows  "sum f {0..n} = f 0 - f (Suc n) + sum (\<lambda>i. f (Suc i)) {0..n}"
-by (induct n) auto
+  shows "sum f {0..n} = f 0 - f (Suc n) + sum (\<lambda>i. f (Suc i)) {0..n}"
+  by (induct n) auto
 
 lemma field_Taylor:
   assumes S: "convex S"
@@ -836,7 +761,7 @@
     assume "u \<in> closed_segment w z"
     then have "u \<in> S"
       by (metis wzs subsetD)
-    have "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) +
+    have *: "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) +
                       f (Suc i) u * (z-u)^i / (fact i)) =
               f (Suc n) u * (z-u) ^ n / (fact n)"
     proof (induction n)
@@ -849,7 +774,7 @@
            f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) -
            f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))"
         using Suc by simp
-      also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))"
+      also have "\<dots> = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))"
       proof -
         have "(fact(Suc n)) *
              (f(Suc n) u *(z-u) ^ n / (fact n) +
@@ -859,29 +784,26 @@
             ((fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / (fact(Suc n))) -
             ((fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / (fact(Suc n))"
           by (simp add: algebra_simps del: fact_Suc)
-        also have "... = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) +
+        also have "\<dots> = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) +
                          (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -
                          (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"
           by (simp del: fact_Suc)
-        also have "... = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) +
+        also have "\<dots> = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) +
                          (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -
                          (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"
           by (simp only: fact_Suc of_nat_mult ac_simps) simp
-        also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)"
+        also have "\<dots> = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)"
           by (simp add: algebra_simps)
         finally show ?thesis
         by (simp add: mult_left_cancel [where c = "(fact (Suc n))", THEN iffD1] del: fact_Suc)
       qed
       finally show ?case .
     qed
-    then have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / (fact i)))
+    have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / (fact i)))
                 has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n))
                (at u within S)"
-      apply (intro derivative_eq_intros)
-      apply (blast intro: assms \<open>u \<in> S\<close>)
-      apply (rule refl)+
-      apply (auto simp: field_simps)
-      done
+      unfolding * [symmetric]
+      by (rule derivative_eq_intros assms \<open>u \<in> S\<close> refl | auto simp: field_simps)+
   } note sum_deriv = this
   { fix u
     assume u: "u \<in> closed_segment w z"
@@ -889,9 +811,9 @@
       by (metis wzs subsetD)
     have "norm (f (Suc n) u) * norm (z - u) ^ n \<le> norm (f (Suc n) u) * norm (u - z) ^ n"
       by (metis norm_minus_commute order_refl)
-    also have "... \<le> norm (f (Suc n) u) * norm (z - w) ^ n"
+    also have "\<dots> \<le> norm (f (Suc n) u) * norm (z - w) ^ n"
       by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u])
-    also have "... \<le> B * norm (z - w) ^ n"
+    also have "\<dots> \<le> B * norm (z - w) ^ n"
       by (metis norm_ge_zero zero_le_power mult_right_mono  B [OF us])
     finally have "norm (f (Suc n) u) * norm (z - u) ^ n \<le> B * norm (z - w) ^ n" .
   } note cmod_bound = this
@@ -903,14 +825,14 @@
                 \<le> norm ((\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)) -
                         (\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)))"
     by (simp add: norm_minus_commute)
-  also have "... \<le> B * norm (z - w) ^ n / (fact n) * norm (w - z)"
-    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: 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)"
+  also have "\<dots> \<le> B * norm (z - w) ^ n / (fact n) * norm (w - z)"
+  proof (rule field_differentiable_bound)
+    show "\<And>x. x \<in> closed_segment w z \<Longrightarrow>
+          ((\<lambda>\<xi>. \<Sum>i\<le>n. f i \<xi> * (z - \<xi>) ^ i / fact i) has_field_derivative f (Suc n) x * (z - x) ^ n / fact n)
+           (at x within closed_segment w z)"
+      using DERIV_subset sum_deriv wzs by blast
+  qed (auto simp: norm_divide norm_mult norm_power divide_le_cancel cmod_bound)
+  also have "\<dots>  \<le> B * norm (z - w) ^ Suc n / (fact n)"
     by (simp add: algebra_simps norm_minus_commute)
   finally show ?thesis .
 qed
@@ -921,8 +843,7 @@
       and B: "\<And>x. x \<in> S \<Longrightarrow> cmod (f (Suc n) x) \<le> B"
       and w: "w \<in> S"
       and z: "z \<in> S"
-    shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i)))
-          \<le> B * cmod(z - w)^(Suc n) / fact n"
+    shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i))) \<le> B * cmod(z - w)^(Suc n) / fact n"
   using assms by (rule field_Taylor)
 
 
@@ -932,20 +853,22 @@
   assumes "\<And>u. u \<in> closed_segment w z \<Longrightarrow> (f has_field_derivative f'(u)) (at u)"
     shows "\<exists>u. u \<in> closed_segment w z \<and> Re(f z) - Re(f w) = Re(f'(u) * (z - w))"
 proof -
-  have twz: "\<And>t. (1 - t) *\<^sub>R w + t *\<^sub>R z = w + t *\<^sub>R (z - w)"
-    by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib)
+  define \<phi> where "\<phi> \<equiv> \<lambda>t. (1 - t) *\<^sub>R w + t *\<^sub>R z"
+  have twz: "\<And>t. \<phi> t = w + t *\<^sub>R (z - w)"
+    by (simp add: \<phi>_def real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib)
   note assms[unfolded has_field_derivative_def, derivative_intros]
-  show ?thesis
-    apply (cut_tac mvt_simple
-                     [of 0 1 "Re o f o (\<lambda>t. (1 - t) *\<^sub>R w +  t *\<^sub>R z)"
-                      "\<lambda>u. Re o (\<lambda>h. f'((1 - u) *\<^sub>R w + u *\<^sub>R z) * h) o (\<lambda>t. t *\<^sub>R (z - w))"])
-    apply auto
-    apply (rule_tac x="(1 - x) *\<^sub>R w + x *\<^sub>R z" in exI)
-    apply (auto simp: closed_segment_def twz) []
-    apply (intro derivative_eq_intros has_derivative_at_withinI, simp_all)
-    apply (simp add: fun_eq_iff real_vector.scale_right_diff_distrib)
-    apply (force simp: twz closed_segment_def)
-    done
+  have *: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
+        \<Longrightarrow> (Re \<circ> f \<circ> \<phi> has_derivative Re \<circ> (*) (f' (\<phi> x)) \<circ> (\<lambda>t. t *\<^sub>R (z - w)))
+            (at x within {0..1})"
+    unfolding \<phi>_def
+    by (intro derivative_eq_intros has_derivative_at_withinI) 
+       (auto simp: in_segment scaleR_right_diff_distrib)
+  obtain x where "0<x" "x<1" "(Re \<circ> f \<circ> \<phi>) 1 -
+       (Re \<circ> f \<circ> \<phi>) 0 = (Re \<circ> (*) (f' (\<phi> x)) \<circ> (\<lambda>t. t *\<^sub>R (z - w))) (1 - 0)"
+    using mvt_simple [OF zero_less_one *] by force
+  then show ?thesis
+    unfolding \<phi>_def
+    by (smt (verit) comp_apply in_segment(1) scaleR_left_distrib scaleR_one scaleR_zero_left)
 qed
 
 lemma complex_Taylor_mvt:
@@ -967,30 +890,27 @@
                  (f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) /
                  (fact (Suc i)))"
        by (subst sum_Suc_reindex) simp
-    also have "... = f (Suc 0) u -
+    also have "\<dots> = f (Suc 0) u -
              (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) /
              (fact (Suc n)) +
              (\<Sum>i = 0..n.
                  f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i))  -
                  f (Suc i) u * (z-u) ^ i / (fact i))"
       by (simp only: diff_divide_distrib fact_cancel ac_simps)
-    also have "... = f (Suc 0) u -
+    also have "\<dots> = f (Suc 0) u -
              (f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) /
              (fact (Suc n)) +
              f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u"
       by (subst sum_Suc_diff) auto
-    also have "... = f (Suc n) u * (z-u) ^ n / (fact n)"
+    also have "\<dots> = f (Suc n) u * (z-u) ^ n / (fact n)"
       by (simp only: algebra_simps diff_divide_distrib fact_cancel)
-    finally have "(\<Sum>i = 0..n. (f (Suc i) u * (z - u) ^ i
+    finally have *: "(\<Sum>i = 0..n. (f (Suc i) u * (z - u) ^ i
                              - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) =
                   f (Suc n) u * (z - u) ^ n / (fact n)" .
-    then have "((\<lambda>u. \<Sum>i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative
+    have "((\<lambda>u. \<Sum>i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative
                 f (Suc n) u * (z - u) ^ n / (fact n))  (at u)"
-      apply (intro derivative_eq_intros)+
-      apply (force intro: u assms)
-      apply (rule refl)+
-      apply (auto simp: ac_simps)
-      done
+      unfolding * [symmetric]
+      by (rule derivative_eq_intros assms u refl | auto simp: field_simps)+
   }
   then show ?thesis
     apply (cut_tac complex_mvt_line [of w z "\<lambda>u. \<Sum>i = 0..n. f i u * (z-u) ^ i / (fact i)"
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -437,7 +437,7 @@
     then obtain L where L: "subspace L" "affine_parallel S L"
       using assms affine_parallel_subspace[of S] by auto
     then obtain a where a: "S = ((+) a ` L)"
-      using affine_parallel_def[of L S] affine_parallel_commut by auto
+      using affine_parallel_def[of L S] affine_parallel_commute by auto
     from L have "closed L" using closed_subspace by auto
     then have "closed S"
       using closed_translation a by auto
--- a/src/HOL/Analysis/Elementary_Topology.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Elementary_Topology.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -1692,20 +1692,15 @@
     (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l))"
 
 lemma seq_compactI:
-  assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
+  assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
   shows "seq_compact S"
   unfolding seq_compact_def using assms by fast
 
 lemma seq_compactE:
   assumes "seq_compact S" "\<forall>n. f n \<in> S"
-  obtains l r where "l \<in> S" "strict_mono (r :: nat \<Rightarrow> nat)" "((f \<circ> r) \<longlongrightarrow> l) sequentially"
+  obtains l r where "l \<in> S" "strict_mono (r :: nat \<Rightarrow> nat)" "(f \<circ> r) \<longlonglongrightarrow> l"
   using assms unfolding seq_compact_def by fast
 
-lemma closed_sequentially: (* TODO: move upwards *)
-  assumes "closed S" and "\<And>n. f n \<in> S" and "f \<longlonglongrightarrow> l"
-  shows "l \<in> S"
-  by (metis Lim_in_closed_set assms eventually_sequentially trivial_limit_sequentially)
-
 lemma seq_compact_Int_closed:
   assumes "seq_compact S" and "closed T"
   shows "seq_compact (S \<inter> T)"
--- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -1373,7 +1373,7 @@
       obtain a where "a \<in> interior (convex hull insert 0 B)"
       proof (rule interior_simplex_nonempty [OF indB])
         show "finite B"
-          by (simp add: indB independent_finite)
+          by (simp add: indB independent_imp_finite)
         show "card B = DIM('N)"
           by (simp add: cardB 2)
       qed
--- a/src/HOL/Analysis/Finite_Product_Measure.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Finite_Product_Measure.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -8,9 +8,11 @@
 imports Binary_Product_Measure Function_Topology
 begin
 
-lemma PiE_choice: "(\<exists>f\<in>Pi\<^sub>E I F. \<forall>i\<in>I. P i (f i)) \<longleftrightarrow> (\<forall>i\<in>I. \<exists>x\<in>F i. P i x)"
-  by (auto simp: Bex_def PiE_iff Ball_def dest!: choice_iff'[THEN iffD1])
-     (force intro: exI[of _ "restrict f I" for f])
+lemma Pi_choice: "(\<forall>i\<in>I. \<exists>x\<in>F i. P i x) \<longleftrightarrow> (\<exists>f\<in>Pi I F. \<forall>i\<in>I. P i (f i))"
+  by (metis Pi_iff)
+
+lemma PiE_choice: "(\<forall>i\<in>I. \<exists>x\<in>F i. P i x) \<longleftrightarrow>(\<exists>f\<in>Pi\<^sub>E I F. \<forall>i\<in>I. P i (f i))"
+  unfolding Pi_choice by (metis Int_iff PiE_def restrict_PiE restrict_apply)
 
 lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
   by auto
@@ -97,7 +99,7 @@
     then show "x \<in> A \<longleftrightarrow> x \<in> B"
       using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S]
       by (auto simp del: PiE_cancel_merge simp add: Un_absorb1 eq)
-  qed (insert sets, auto)
+  qed (use sets in auto)
 qed
 
 lemma restrict_vimage:
@@ -178,7 +180,7 @@
 
 lemma Pi_cong_sets:
     "\<lbrakk>I = J; \<And>x. x \<in> I \<Longrightarrow> M x = N x\<rbrakk> \<Longrightarrow> Pi I M = Pi J N"
-  unfolding Pi_def by auto
+  by auto
 
 lemma PiM_cong:
   assumes "I = J" "\<And>x. x \<in> I \<Longrightarrow> M x = N x"
@@ -327,17 +329,11 @@
     done
 qed
 proposition prod_algebra_cong:
-  assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
+  assumes "I = J" and "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
   shows "prod_algebra I M = prod_algebra J N"
-proof -
-  have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)"
-    using sets_eq_imp_space_eq[OF sets] by auto
-  with sets show ?thesis unfolding \<open>I = J\<close>
-    by (intro antisym prod_algebra_mono) auto
-qed
+  by (metis assms prod_algebra_mono sets_eq_imp_space_eq subsetI subset_antisym)
 
-lemma space_in_prod_algebra:
-  "(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
+lemma space_in_prod_algebra: "(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
 proof cases
   assume "I = {}" then show ?thesis
     by (auto simp add: prod_algebra_def image_iff prod_emb_def)
@@ -346,9 +342,8 @@
   then obtain i where "i \<in> I" by auto
   then have "(\<Pi>\<^sub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
     by (auto simp: prod_emb_def)
-  also have "\<dots> \<in> prod_algebra I M"
-    using \<open>i \<in> I\<close> by (intro prod_algebraI) auto
-  finally show ?thesis .
+  then show ?thesis
+    by (simp add: \<open>i \<in> I\<close> prod_algebraI) 
 qed
 
 lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
@@ -383,8 +378,8 @@
   show "A \<in> sigma_sets ?\<Omega> ?R"
   proof cases
     assume "I = {}"
-    with X have "A = {\<lambda>x. undefined}" by (auto simp: prod_emb_def)
-    with \<open>I = {}\<close> show ?thesis by (auto intro!: sigma_sets_top)
+    with X show ?thesis
+      by (metis (no_types, lifting) PiE_cong R.top empty_iff prod_emb_PiE subset_eq)
   next
     assume "I \<noteq> {}"
     with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^sub>E i\<in>I. space (M i)). f j \<in> X j})"
@@ -405,16 +400,25 @@
 qed
 
 lemma sets_PiM_eq_proj:
-  "I \<noteq> {} \<Longrightarrow> sets (PiM I M) = sets (SUP i\<in>I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i))"
-  apply (simp add: sets_PiM_single)
-  apply (subst sets_Sup_eq[where X="\<Pi>\<^sub>E i\<in>I. space (M i)"])
-  apply auto []
-  apply auto []
-  apply simp
-  apply (subst arg_cong [of _ _ Sup, OF image_cong, OF refl])
-  apply (rule sets_vimage_algebra2)
-  apply (auto intro!: arg_cong2[where f=sigma_sets])
-  done
+  assumes "I \<noteq> {}"
+  shows "sets (PiM I M) = sets (SUP i\<in>I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i))"
+        (is "?lhs = ?rhs")
+proof -
+  have "?lhs = 
+        sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
+    by (simp add: sets_PiM_single)
+  also have "\<dots> = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i))
+                (\<Union>x\<in>I. sets (vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>xa. xa x) (M x)))"
+    apply (subst arg_cong [of _ _ Sup, OF image_cong, OF refl])
+     apply (rule sets_vimage_algebra2)
+    by (auto intro!: arg_cong2[where f=sigma_sets])
+  also have "... = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i))
+     (\<Union> (sets ` (\<lambda>i. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i)) ` I))"
+    by simp
+  also have "... = ?rhs"
+    by (subst sets_Sup_eq[where X="\<Pi>\<^sub>E i\<in>I. space (M i)"]) (use assms in auto)
+  finally show ?thesis .
+qed
 
 lemma
   shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\<lambda>k. undefined}"
@@ -616,7 +620,7 @@
     by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: if_split_asm)
   then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) \<in> sets (Pi\<^sub>M I M)"
     using \<open>A \<in> sets (M i)\<close> \<open>i \<in> I\<close> by (auto intro!: sets_PiM_I)
-qed (insert \<open>i \<in> I\<close>, auto simp: space_PiM)
+qed (use \<open>i \<in> I\<close> in \<open>auto simp: space_PiM\<close>)
 
 lemma measurable_component_singleton'[measurable_dest]:
   assumes f: "f \<in> measurable N (Pi\<^sub>M I M)"
@@ -696,7 +700,10 @@
     by auto
   then show "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} \<in> sets N"
     using A X by (auto intro!: measurable_sets)
-qed (insert X, auto simp add: PiE_def dest: measurable_space)
+next
+  show "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
+    using X by (auto simp add: PiE_def dest: measurable_space)
+qed 
 
 lemma measurable_abs_UNIV:
   "(\<And>n. (\<lambda>\<omega>. f n \<omega>) \<in> measurable M (N n)) \<Longrightarrow> (\<lambda>\<omega> n. f n \<omega>) \<in> measurable M (PiM UNIV N)"
@@ -708,13 +715,7 @@
 lemma measurable_restrict_subset':
   assumes "J \<subseteq> L" "\<And>x. x \<in> J \<Longrightarrow> sets (M x) = sets (N x)"
   shows "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)"
-proof-
-  from assms(1) have "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
-    by (rule measurable_restrict_subset)
-  also from assms(2) have "measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M) = measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)"
-    by (intro sets_PiM_cong measurable_cong_sets) simp_all
-  finally show ?thesis .
-qed
+  by (metis (no_types) assms measurable_cong_sets measurable_restrict_subset sets_PiM_cong)
 
 lemma measurable_prod_emb[intro, simp]:
   "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^sub>M L M)"
@@ -754,11 +755,7 @@
 
 lemma sets_in_extensional_aux:
   "{x\<in>space (PiM I M). x \<in> extensional I} \<in> sets (PiM I M)"
-proof -
-  have "{x\<in>space (PiM I M). x \<in> extensional I} = space (PiM I M)"
-    by (auto simp add: extensional_def space_PiM)
-  then show ?thesis by simp
-qed
+  by (smt (verit) PiE_iff mem_Collect_eq sets.top space_PiM subsetI subset_antisym)
 
 lemma sets_in_extensional[measurable (raw)]:
   "f \<in> measurable N (PiM I M) \<Longrightarrow> Measurable.pred N (\<lambda>x. f x \<in> extensional I)"
@@ -772,7 +769,7 @@
   then have "Pi\<^sub>E I E = (\<Inter>i\<in>I. prod_emb I M {i} (Pi\<^sub>E {i} E))"
     using E[THEN sets.sets_into_space] by (auto simp: PiE_iff prod_emb_def fun_eq_iff)
   also have "\<dots> \<in> sets (PiM I M)"
-    using I \<open>I \<noteq> {}\<close> by (safe intro!: sets.countable_INT' measurable_prod_emb sets_PiM_I_finite E)
+    using I \<open>I \<noteq> {}\<close> by (simp add: E sets.countable_INT' sets_PiM_I subset_eq)
   finally show ?thesis .
 qed (simp add: sets_PiM_empty)
 
@@ -805,8 +802,9 @@
     and K: "\<And>i. K i = prod_emb I M (J i) (X i)"
     by (metis Union.IH)
   show ?case
-  proof (intro exI[of _ "\<Union>i. J i"] bexI[of _ "\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i)"] conjI)
-    show "(\<Union>i. J i) \<subseteq> I" "countable (\<Union>i. J i)" using J by auto
+  proof (intro exI bexI conjI)
+    show "(\<Union>i. J i) \<subseteq> I" "countable (\<Union>i. J i)" 
+      using J by auto
     with J show "\<Union>(K ` UNIV) = prod_emb I M (\<Union>i. J i) (\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i))"
       by (simp add: K[abs_def] SUP_upper)
   qed(auto intro: X)
@@ -1010,7 +1008,7 @@
     in_space: "\<And>j. space (M j) = \<Union>(F j)"
     using sigma_finite_countable by (metis subset_eq)
   moreover have "(\<Union>(Pi\<^sub>E I ` Pi\<^sub>E I F)) = space (Pi\<^sub>M I M)"
-    using in_space by (auto simp: space_PiM PiE_iff intro!: PiE_choice[THEN iffD2])
+    using in_space by (auto simp: space_PiM PiE_iff intro!: PiE_choice[THEN iffD1])
   ultimately show "\<exists>A. countable A \<and> A \<subseteq> sets (Pi\<^sub>M I M) \<and> \<Union>A = space (Pi\<^sub>M I M) \<and> (\<forall>a\<in>A. emeasure (Pi\<^sub>M I M) a \<noteq> \<infinity>)"
     by (intro exI[of _ "Pi\<^sub>E I ` Pi\<^sub>E I F"])
        (auto intro!: countable_PiE sets_PiM_I_finite
@@ -1042,26 +1040,26 @@
       (\<Prod>i\<in>I \<union> J. emeasure (M i) (A i))"
     by (subst emeasure_distr)
        (auto simp: * J.emeasure_pair_measure_Times I.measure_times J.measure_times prod.union_disjoint)
-qed (insert fin, simp_all)
+qed (use fin in simp_all)
 
 proposition (in product_sigma_finite) product_nn_integral_fold:
   assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
   and f[measurable]: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
-  shows "integral\<^sup>N (Pi\<^sub>M (I \<union> J) M) f =
-    (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^sub>M J M)) \<partial>(Pi\<^sub>M I M))"
+shows "integral\<^sup>N (Pi\<^sub>M (I \<union> J) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^sub>M J M)) \<partial>(Pi\<^sub>M I M))"
+        (is "?lhs = ?rhs")
 proof -
   interpret I: finite_product_sigma_finite M I by standard fact
   interpret J: finite_product_sigma_finite M J by standard fact
   interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by standard
   have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
     using measurable_comp[OF measurable_merge f] by (simp add: comp_def)
-  show ?thesis
-    apply (subst distr_merge[OF IJ, symmetric])
-    apply (subst nn_integral_distr[OF measurable_merge])
-    apply measurable []
-    apply (subst J.nn_integral_fst[symmetric, OF P_borel])
-    apply simp
-    done
+  have "?lhs = integral\<^sup>N (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) f"
+    by (simp add: I.finite_index J.finite_index assms(1) distr_merge)
+  also have "... = \<integral>\<^sup>+ x. f (merge I J x) \<partial>(Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
+    by (simp add: nn_integral_distr)
+  also have "... = ?rhs"
+    using P.Fubini P.nn_integral_snd by force
+  finally show ?thesis .
 qed
 
 lemma (in product_sigma_finite) distr_singleton:
@@ -1082,10 +1080,7 @@
 proof -
   interpret I: finite_product_sigma_finite M "{i}" by standard simp
   from f show ?thesis
-    apply (subst distr_singleton[symmetric])
-    apply (subst nn_integral_distr[OF measurable_component_singleton])
-    apply simp_all
-    done
+    by (metis distr_singleton insert_iff measurable_component_singleton nn_integral_distr)
 qed
 
 proposition (in product_sigma_finite) product_nn_integral_insert:
@@ -1118,8 +1113,7 @@
   shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x(i := y)) \<partial>(Pi\<^sub>M I M)) \<partial>(M i))"
   apply (subst product_nn_integral_insert[OF assms])
   apply (rule pair_sigma_finite.Fubini')
-  apply intro_locales []
-  apply (rule sigma_finite[OF I(1)])
+  apply (simp add: local.sigma_finite pair_sigma_finite.intro sigma_finite_measures)
   apply measurable
   done
 
@@ -1139,11 +1133,8 @@
               measurable_comp[OF measurable_component_singleton, unfolded comp_def])
        auto
   then show ?case
-    apply (simp add: product_nn_integral_insert[OF insert(1,2)])
-    apply (simp add: insert(2-) * nn_integral_multc)
-    apply (subst nn_integral_cmult)
-    apply (auto simp add: insert(2-))
-    done
+    using product_nn_integral_insert[OF insert(1,2)]
+    by (simp add: insert(2-) * nn_integral_multc nn_integral_cmult)
 qed (simp add: space_PiM)
 
 proposition (in product_sigma_finite) product_nn_integral_pair:
@@ -1176,7 +1167,8 @@
 lemma (in product_sigma_finite)
   assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)"
   shows emeasure_fold_integral:
-    "emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" (is ?I)
+    "emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" 
+         (is "?lhs = ?rhs")
     and emeasure_fold_measurable:
     "(\<lambda>x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M))) \<in> borel_measurable (Pi\<^sub>M I M)" (is ?B)
 proof -
@@ -1185,13 +1177,14 @@
   interpret IJ: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" ..
   have merge: "merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) \<in> sets (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
     by (intro measurable_sets[OF _ A] measurable_merge assms)
-
-  show ?I
-    apply (subst distr_merge[symmetric, OF IJ])
-    apply (subst emeasure_distr[OF measurable_merge A])
+  have "?lhs = emeasure (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) A"
+    by (simp add: I.finite_index J.finite_index assms(1) distr_merge)
+  also have "... = emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M))"
+    by (meson A emeasure_distr measurable_merge)
+  also have "... = ?rhs"
     apply (subst J.emeasure_pair_measure_alt[OF merge])
-    apply (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure)
-    done
+    by (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure)
+  finally show "?lhs = ?rhs" .
 
   show ?B
     using IJ.measurable_emeasure_Pair1[OF merge]
@@ -1201,7 +1194,6 @@
 lemma sets_Collect_single:
   "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^sub>M I M). x i \<in> A } \<in> sets (Pi\<^sub>M I M)"
   by simp
-
 lemma pair_measure_eq_distr_PiM:
   fixes M1 :: "'a measure" and M2 :: "'a measure"
   assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
@@ -1282,7 +1274,7 @@
         {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}"
 proof
   have "{(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}} \<subseteq> sets (Pi\<^sub>M I M)"
-  proof (auto)
+  proof clarify
     fix X assume H: "\<forall>i. X i \<in> sets (M i)" "finite {i. X i \<noteq> space (M i)}"
     then have *: "X i \<in> sets (M i)" for i by simp
     define J where "J = {i \<in> I. X i \<noteq> space (M i)}"
@@ -1315,9 +1307,7 @@
   qed
   show "sets (Pi\<^sub>M I M) \<subseteq> sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}"
     unfolding sets_PiM_single
-    apply (rule sigma_sets_mono')
-    apply (auto simp add: PiE_iff *)
-    done
+    by (intro sigma_sets_mono') (auto simp add: PiE_iff *)
 qed
 
 lemma sets_PiM_subset_borel:
@@ -1361,7 +1351,7 @@
     ultimately show ?thesis unfolding \<open>U = (\<Union>B)\<close> by auto
   qed
   have "sigma_sets UNIV (Collect open) \<subseteq> sets (Pi\<^sub>M UNIV (\<lambda>i::'a. (borel::('b measure))))"
-    apply (rule sets.sigma_sets_subset') using ** by auto
+    by (metis "**" mem_Collect_eq open_UNIV sets.sigma_sets_subset' subsetI)
   then show "sets (borel::('a \<Rightarrow> 'b) measure) \<subseteq> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))"
     unfolding borel_def by auto
 qed (simp add: sets_PiM_subset_borel)
--- a/src/HOL/Analysis/Further_Topology.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Further_Topology.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -30,9 +30,9 @@
   proof
     have "\<And>u. \<lbrakk>u \<in> S; norm u *\<^sub>R f (u /\<^sub>R norm u) \<notin> T\<rbrakk> \<Longrightarrow> u = 0"
       by (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \<open>subspace T\<close>] fim image_subset_iff inf_le2 singletonD)
-    then have "g ` (S - {0}) \<subseteq> T"
+    then have "g \<in> (S - {0}) \<rightarrow> T"
       using g_def by blast
-    moreover have "g ` (S - {0}) \<subseteq> UNIV - {0}"
+    moreover have "g \<in> (S - {0}) \<rightarrow> UNIV - {0}"
     proof (clarsimp simp: g_def)
       fix y
       assume "y \<in> S" and f0: "f (y /\<^sub>R norm y) = 0"
@@ -159,30 +159,24 @@
   assumes ST: "subspace S" "subspace T" "dim S < dim T"
       and "S \<subseteq> T"
       and contf: "continuous_on (sphere 0 1 \<inter> S) f"
-      and fim: "f ` (sphere 0 1 \<inter> S) \<subseteq> sphere 0 1 \<inter> T"
+      and fim: "f \<in> (sphere 0 1 \<inter> S) \<rightarrow> sphere 0 1 \<inter> T"
     shows "\<exists>c. homotopic_with_canon (\<lambda>x. True) (sphere 0 1 \<inter> S) (sphere 0 1 \<inter> T) f (\<lambda>x. c)"
 proof -
   have [simp]: "\<And>x. \<lbrakk>norm x = 1; x \<in> S\<rbrakk> \<Longrightarrow> norm (f x) = 1"
-    using fim by (simp add: image_subset_iff)
+    using fim by auto
   have "compact (sphere 0 1 \<inter> S)"
     by (simp add: \<open>subspace S\<close> closed_subspace compact_Int_closed)
-  then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \<inter> S) \<subseteq> T"
+  then obtain g where pfg: "polynomial_function g" and gim: "g \<in> (sphere 0 1 \<inter> S) \<rightarrow> T"
                 and g12: "\<And>x. x \<in> sphere 0 1 \<inter> S \<Longrightarrow> norm(f x - g x) < 1/2"
-    using Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \<open>subspace T\<close>, of "1/2"] fim by auto
+    using Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \<open>subspace T\<close>, of "1/2"] fim
+    by (auto simp: image_subset_iff_funcset)
   have gnz: "g x \<noteq> 0" if "x \<in> sphere 0 1 \<inter> S" for x
-  proof -
-    have "norm (f x) = 1"
-      using fim that by (simp add: image_subset_iff)
-    then show ?thesis
-      using g12 [OF that] by auto
-  qed
+    using g12 that by fastforce
   have diffg: "g differentiable_on sphere 0 1 \<inter> S"
     by (metis pfg differentiable_on_polynomial_function)
   define h where "h \<equiv> \<lambda>x. inverse(norm(g x)) *\<^sub>R g x"
   have h: "x \<in> sphere 0 1 \<inter> S \<Longrightarrow> h x \<in> sphere 0 1 \<inter> T" for x
-    unfolding h_def
-    using gnz [of x]
-    by (auto simp: subspace_mul [OF \<open>subspace T\<close>] subsetD [OF gim])
+    unfolding h_def using \<open>subspace T\<close> gim gnz subspace_mul by fastforce
   have diffh: "h differentiable_on sphere 0 1 \<inter> S"
     unfolding h_def using gnz
     by (fastforce intro: derivative_intros diffg differentiable_on_compose [OF diffg])
@@ -205,7 +199,7 @@
       have "convex T"
         by (simp add: \<open>subspace T\<close> subspace_imp_convex)
       then have "convex hull {f x, g x} \<subseteq> T"
-        by (metis IntD2 closed_segment_subset fim gim image_subset_iff segment_convex_hull that)
+        by (metis IntD2 PiE closed_segment_subset fim gim segment_convex_hull that)
       then show ?thesis
         using that non0fg segment_convex_hull by fastforce
     qed
@@ -279,14 +273,9 @@
       then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \<inter> T)"
         using \<open>aff_dim T = aff_dim S\<close> by simp
       have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \<inter> T)"
-      proof (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \<open>convex S\<close> \<open>bounded S\<close>])
-        show "convex (ball 0 1 \<inter> T)"
-          by (simp add: \<open>subspace T\<close> convex_Int subspace_imp_convex)
-        show "bounded (ball 0 1 \<inter> T)"
-          by (simp add: bounded_Int)
-        show "aff_dim S = aff_dim (ball 0 1 \<inter> T)"
-          by (rule affS_eq)
-      qed
+        using homeomorphic_rel_frontiers_convex_bounded_sets [OF \<open>convex S\<close> \<open>bounded S\<close>]
+        by (simp add: \<open>subspace T\<close> affS_eq assms bounded_Int convex_Int 
+            homeomorphic_rel_frontiers_convex_bounded_sets subspace_imp_convex)
       also have "... = frontier (ball 0 1) \<inter> T"
       proof (rule convex_affine_rel_frontier_Int [OF convex_ball])
         show "affine T"
@@ -318,10 +307,8 @@
   then show ?thesis
   proof (cases "T = {}")
     case True
-    then have "rel_frontier S = {}" "rel_frontier T = {}"
-      using fim by fastforce+
     then show ?thesis
-      using that by (simp add: homotopic_on_emptyI)
+      by (smt (verit, best) False affST aff_dim_negative_iff)
   next
     case False
     obtain T':: "'a set"
@@ -383,16 +370,16 @@
   assumes fin: "finite \<F>"
       and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
       and "\<And>S T. \<lbrakk>S \<in> \<F>; T \<in> \<F>; S \<noteq> T\<rbrakk> \<Longrightarrow> S \<inter> T \<subseteq> K"
-      and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g ` S \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
-   shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g ` (\<Union>\<F>) \<subseteq> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x)"
+      and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g \<in> S \<rightarrow> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
+   shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g \<in> (\<Union>\<F>) \<rightarrow> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x)"
 using assms
 proof (induction \<F>)
   case empty show ?case by simp
 next
   case (insert S \<F>)
-  then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \<subseteq> T" and feq: "\<forall>x \<in> S \<inter> K. f x = h x"
-    by (meson insertI1)
-  obtain g where contg: "continuous_on (\<Union>\<F>) g" and gim: "g ` \<Union>\<F> \<subseteq> T" and geq: "\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x"
+  then obtain f where contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> T" and feq: "\<forall>x \<in> S \<inter> K. f x = h x"
+    by (metis funcset_image insert_iff)
+  obtain g where contg: "continuous_on (\<Union>\<F>) g" and gim: "g \<in> (\<Union>\<F>) \<rightarrow> T" and geq: "\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x"
     using insert by auto
   have fg: "f x = g x" if "x \<in> T" "T \<in> \<F>" "x \<in> S" for x T
   proof -
@@ -403,38 +390,38 @@
   qed
   moreover have "continuous_on (S \<union> \<Union> \<F>) (\<lambda>x. if x \<in> S then f x else g x)"
     by (auto simp: insert closed_Union contf contg  intro: fg continuous_on_cases)
+  moreover have "S \<union> \<Union> \<F> = \<Union> (insert S \<F>)"
+    by auto
   ultimately show ?case
-    by (smt (verit, del_insts) Int_iff UnE complete_lattice_class.Sup_insert feq fim geq gim image_subset_iff)
+    by (smt (verit) Int_iff Pi_iff UnE feq fim geq gim)
 qed
 
 lemma extending_maps_Union:
   assumes fin: "finite \<F>"
-      and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g ` S \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
-      and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
-      and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; \<not> X \<subseteq> Y; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
-    shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g ` (\<Union>\<F>) \<subseteq> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x)"
-apply (simp flip: Union_maximal_sets [OF fin])
-apply (rule extending_maps_Union_aux)
-apply (simp_all add: Union_maximal_sets [OF fin] assms)
-by (metis K psubsetI)
-
+    and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g \<in> S \<rightarrow> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
+    and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
+    and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; \<not> X \<subseteq> Y; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
+  shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g \<in> (\<Union>\<F>) \<rightarrow> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x)"
+proof -
+  have "\<And>S T. \<lbrakk>S \<in> \<F>; \<forall>U\<in>\<F>. \<not> S \<subset> U; T \<in> \<F>; \<forall>U\<in>\<F>. \<not> T \<subset> U; S \<noteq> T\<rbrakk> \<Longrightarrow> S \<inter> T \<subseteq> K"
+    by (metis K psubsetI)
+  then show ?thesis
+    apply (simp flip: Union_maximal_sets [OF fin])
+    apply (rule extending_maps_Union_aux, simp_all add: Union_maximal_sets [OF fin] assms)
+    done
+qed
 
 lemma extend_map_lemma:
   assumes "finite \<F>" "\<G> \<subseteq> \<F>" "convex T" "bounded T"
       and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
       and aff: "\<And>X. X \<in> \<F> - \<G> \<Longrightarrow> aff_dim X < aff_dim T"
       and face: "\<And>S T. \<lbrakk>S \<in> \<F>; T \<in> \<F>\<rbrakk> \<Longrightarrow> (S \<inter> T) face_of S"
-      and contf: "continuous_on (\<Union>\<G>) f" and fim: "f ` (\<Union>\<G>) \<subseteq> rel_frontier T"
-  obtains g where "continuous_on (\<Union>\<F>) g" "g ` (\<Union>\<F>) \<subseteq> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"
+      and contf: "continuous_on (\<Union>\<G>) f" and fim: "f \<in> (\<Union>\<G>) \<rightarrow> rel_frontier T"
+  obtains g where "continuous_on (\<Union>\<F>) g" "g \<in> (\<Union>\<F>) \<rightarrow> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"
 proof (cases "\<F> - \<G> = {}")
   case True
   show ?thesis
-  proof
-    show "continuous_on (\<Union> \<F>) f"
-      using True \<open>\<G> \<subseteq> \<F>\<close> contf by auto
-    show "f ` \<Union> \<F> \<subseteq> rel_frontier T"
-      using True fim by auto
-  qed auto
+    using True assms(2) contf fim that by force
 next
   case False
   then have "0 \<le> aff_dim T"
@@ -467,14 +454,14 @@
       using Suc.IH [OF ple] by auto
     let ?Faces = "{D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D \<le> p}"
     have extendh: "\<exists>g. continuous_on D g \<and>
-                       g ` D \<subseteq> rel_frontier T \<and>
+                       g \<in> D \<rightarrow> rel_frontier T \<and>
                        (\<forall>x \<in> D \<inter> \<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}). g x = h x)"
       if D: "D \<in> \<G> \<union> ?Faces" for D
     proof (cases "D \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p})")
       case True
       have "continuous_on D h"
         using True conth continuous_on_subset by blast
-      moreover have "h ` D \<subseteq> rel_frontier T"
+      moreover have "h \<in> D \<rightarrow> rel_frontier T"
         using True him by blast
       ultimately show ?thesis
         by blast
@@ -574,7 +561,7 @@
      show ?thesis
        using that
         apply clarsimp
-        by (smt (verit, ccfv_SIG) IntI face_of_aff_dim_lt face_of_imp_convex [of X] face_of_imp_convex [of Y] face_of_trans ff)
+       by (smt (verit) IntI face_of_aff_dim_lt face_of_imp_convex face_of_trans ff inf_commute)
     qed
     obtain g where "continuous_on (\<Union>(\<G> \<union> ?Faces)) g"
                    "g ` \<Union>(\<G> \<union> ?Faces) \<subseteq> rel_frontier T"
@@ -589,23 +576,20 @@
     show "\<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < int i}) \<subseteq> \<Union>\<F>"
       using \<open>\<G> \<subseteq> \<F>\<close> face_of_imp_subset by fastforce
     show "\<Union>\<F> \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < i})"
-    proof (rule Union_mono)
-      show "\<F> \<subseteq> \<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < int i}"
-        using face by (fastforce simp: aff i)
-    qed
+        using face by (intro Union_mono) (fastforce simp: aff i)
   qed
   have "int i \<le> aff_dim T" by (simp add: i)
   then show ?thesis
-    using extendf [of i] unfolding eq by (metis that)
+    using extendf [of i] that unfolding eq by fastforce
 qed
 
 lemma extend_map_lemma_cofinite0:
   assumes "finite \<F>"
       and "pairwise (\<lambda>S T. S \<inter> T \<subseteq> K) \<F>"
-      and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g ` (S - {a}) \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
+      and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g \<in> (S - {a}) \<rightarrow> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
       and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
     shows "\<exists>C g. finite C \<and> disjnt C U \<and> card C \<le> card \<F> \<and>
-                 continuous_on (\<Union>\<F> - C) g \<and> g ` (\<Union>\<F> - C) \<subseteq> T
+                 continuous_on (\<Union>\<F> - C) g \<and> g \<in> (\<Union>\<F> - C) \<rightarrow> T
                   \<and> (\<forall>x \<in> (\<Union>\<F> - C) \<inter> K. g x = h x)"
   using assms
 proof induction
@@ -614,18 +598,18 @@
 next
   case (insert X \<F>)
   then have "closed X" and clo: "\<And>X. X \<in> \<F> \<Longrightarrow> closed X"
-        and \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g ` (S - {a}) \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
+        and \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g \<in> (S - {a}) \<rightarrow> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
         and pwX: "\<And>Y. Y \<in> \<F> \<and> Y \<noteq> X \<longrightarrow> X \<inter> Y \<subseteq> K \<and> Y \<inter> X \<subseteq> K"
         and pwF: "pairwise (\<lambda> S T. S \<inter> T \<subseteq> K) \<F>"
     by (simp_all add: pairwise_insert)
   obtain C g where C: "finite C" "disjnt C U" "card C \<le> card \<F>"
                and contg: "continuous_on (\<Union>\<F> - C) g"
-               and gim: "g ` (\<Union>\<F> - C) \<subseteq> T"
+               and gim: "g \<in> (\<Union>\<F> - C) \<rightarrow> T"
                and gh:  "\<And>x. x \<in> (\<Union>\<F> - C) \<inter> K \<Longrightarrow> g x = h x"
     using insert.IH [OF pwF \<F> clo] by auto
   obtain a f where "a \<notin> U"
                and contf: "continuous_on (X - {a}) f"
-               and fim: "f ` (X - {a}) \<subseteq> T"
+               and fim: "f \<in> (X - {a}) \<rightarrow> T"
                and fh: "(\<forall>x \<in> X \<inter> K. f x = h x)"
     using insert.prems by (meson insertI1)
   show ?case
@@ -652,19 +636,19 @@
       by (blast intro: continuous_on_subset)
     show "\<forall>x\<in>(\<Union>(insert X \<F>) - insert a C) \<inter> K. (if x \<in> X then f x else g x) = h x"
       using gh by (auto simp: fh)
-    show "(\<lambda>a. if a \<in> X then f a else g a) ` (\<Union>(insert X \<F>) - insert a C) \<subseteq> T"
-      using fim gim by auto force
+    show "(\<lambda>a. if a \<in> X then f a else g a) \<in> (\<Union>(insert X \<F>) - insert a C) \<rightarrow> T"
+      using fim gim Pi_iff by fastforce
   qed
 qed
 
 
 lemma extend_map_lemma_cofinite1:
 assumes "finite \<F>"
-    and \<F>: "\<And>X. X \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (X - {a}) g \<and> g ` (X - {a}) \<subseteq> T \<and> (\<forall>x \<in> X \<inter> K. g x = h x)"
+    and \<F>: "\<And>X. X \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (X - {a}) g \<and> g \<in> (X - {a}) \<rightarrow> T \<and> (\<forall>x \<in> X \<inter> K. g x = h x)"
     and clo: "\<And>X. X \<in> \<F> \<Longrightarrow> closed X"
     and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; \<not> X \<subseteq> Y; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
   obtains C g where "finite C" "disjnt C U" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"
-                    "g ` (\<Union>\<F> - C) \<subseteq> T"
+                    "g \<in> (\<Union>\<F> - C) \<rightarrow> T"
                     "\<And>x. x \<in> (\<Union>\<F> - C) \<inter> K \<Longrightarrow> g x = h x"
 proof -
   let ?\<F> = "{X \<in> \<F>. \<forall>Y\<in>\<F>. \<not> X \<subset> Y}"
@@ -678,7 +662,7 @@
     by (simp add: \<open>finite \<F>\<close> card_mono)
   moreover
   obtain C g where "finite C \<and> disjnt C U \<and> card C \<le> card ?\<F> \<and>
-                 continuous_on (\<Union>?\<F> - C) g \<and> g ` (\<Union>?\<F> - C) \<subseteq> T
+                 continuous_on (\<Union>?\<F> - C) g \<and> g \<in> (\<Union>?\<F> - C) \<rightarrow> T
                   \<and> (\<forall>x \<in> (\<Union>?\<F> - C) \<inter> K. g x = h x)"
     using extend_map_lemma_cofinite0 [OF fin pw, of U T h] by (fastforce intro!: clo \<F>)
   ultimately show ?thesis
@@ -689,12 +673,12 @@
 lemma extend_map_lemma_cofinite:
   assumes "finite \<F>" "\<G> \<subseteq> \<F>" and T: "convex T" "bounded T"
       and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
-      and contf: "continuous_on (\<Union>\<G>) f" and fim: "f ` (\<Union>\<G>) \<subseteq> rel_frontier T"
+      and contf: "continuous_on (\<Union>\<G>) f" and fim: "f \<in> (\<Union>\<G>) \<rightarrow> rel_frontier T"
       and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X"
       and aff: "\<And>X. X \<in> \<F> - \<G> \<Longrightarrow> aff_dim X \<le> aff_dim T"
   obtains C g where
      "finite C" "disjnt C (\<Union>\<G>)" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"
-     "g ` (\<Union> \<F> - C) \<subseteq> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"
+     "g \<in> (\<Union> \<F> - C) \<rightarrow> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"
 proof -
   define \<H> where "\<H> \<equiv> \<G> \<union> {D. \<exists>C \<in> \<F> - \<G>. D face_of C \<and> aff_dim D < aff_dim T}"
   have "finite \<G>"
@@ -707,7 +691,7 @@
     by (metis face inf_commute)
   have *: "\<And>X Y. \<lbrakk>X \<in> \<H>; Y \<in> \<H>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X"
     by (simp add: \<H>_def) (smt (verit) \<open>\<G> \<subseteq> \<F>\<close> DiffE face' face_of_Int_subface in_mono inf.idem)
-  obtain h where conth: "continuous_on (\<Union>\<H>) h" and him: "h ` (\<Union>\<H>) \<subseteq> rel_frontier T"
+  obtain h where conth: "continuous_on (\<Union>\<H>) h" and him: "h \<in> (\<Union>\<H>) \<rightarrow> rel_frontier T"
              and hf: "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> h x = f x" 
   proof (rule extend_map_lemma [OF \<open>finite \<H>\<close> [unfolded \<H>_def] Un_upper1 T])
     show "\<And>X. \<lbrakk>X \<in> \<G> \<union> {D. \<exists>C\<in>\<F> - \<G>. D face_of C \<and> aff_dim D < aff_dim T}\<rbrakk> \<Longrightarrow> polytope X"
@@ -720,11 +704,11 @@
   then obtain a where a: "a \<notin> \<Union>\<G>"
     by blast
   have \<F>: "\<exists>a g. a \<notin> \<Union>\<G> \<and> continuous_on (D - {a}) g \<and>
-                  g ` (D - {a}) \<subseteq> rel_frontier T \<and> (\<forall>x \<in> D \<inter> \<Union>\<H>. g x = h x)"
+                  g \<in> (D - {a}) \<rightarrow> rel_frontier T \<and> (\<forall>x \<in> D \<inter> \<Union>\<H>. g x = h x)"
        if "D \<in> \<F>" for D
   proof (cases "D \<subseteq> \<Union>\<H>")
     case True
-    then have "h ` (D - {a}) \<subseteq> rel_frontier T" "continuous_on (D - {a}) h"
+    then have "h \<in> (D - {a}) \<rightarrow> rel_frontier T" "continuous_on (D - {a}) h"
       using him by (blast intro!: \<open>a \<notin> \<Union>\<G>\<close> continuous_on_subset [OF conth])+
     then show ?thesis
       using a by blast
@@ -755,7 +739,7 @@
           by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD)
         then obtain r where relfD: "rel_frontier D \<subseteq> affine hull D - {b}"
                         and contr: "continuous_on (affine hull D - {b}) r"
-                        and rim: "r ` (affine hull D - {b}) \<subseteq> rel_frontier D"
+                        and rim: "r \<in> (affine hull D - {b}) \<rightarrow> rel_frontier D"
                         and rid: "\<And>x. x \<in> rel_frontier D \<Longrightarrow> r x = x"
           by (auto simp: retract_of_def retraction_def)
         show ?thesis
@@ -774,7 +758,7 @@
           have "r ` (D - {b}) \<subseteq> r ` (affine hull D - {b})"
             by (simp add: Diff_mono hull_subset image_mono)
           also have "... \<subseteq> rel_frontier D"
-            by (rule rim)
+            using rim by auto
           also have "... \<subseteq> \<Union>{E. E face_of D \<and> aff_dim E < aff_dim T}"
             using affD
             by (force simp: rel_frontier_of_polyhedron [OF \<open>polyhedron D\<close>] facet_of_def)
@@ -788,7 +772,7 @@
             show "continuous_on (r ` (D - {b})) h"
               by (simp add: Diff_mono hull_subset continuous_on_subset [OF conth rsub])
           qed
-          show "(h \<circ> r) ` (D - {b}) \<subseteq> rel_frontier T"
+          show "(h \<circ> r) \<in> (D - {b}) \<rightarrow> rel_frontier T"
             using brelD him rsub by fastforce
           show "(h \<circ> r) x = h x" if x: "x \<in> D \<inter> \<Union>\<H>" for x
           proof -
@@ -825,7 +809,7 @@
   have clo: "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
     by (simp add: poly polytope_imp_closed)
   obtain C g where "finite C" "disjnt C (\<Union>\<G>)" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"
-                   "g ` (\<Union>\<F> - C) \<subseteq> rel_frontier T"
+                   "g \<in> (\<Union>\<F> - C) \<rightarrow> rel_frontier T"
                and gh: "\<And>x. x \<in> (\<Union>\<F> - C) \<inter> \<Union>\<H> \<Longrightarrow> g x = h x"
   proof (rule extend_map_lemma_cofinite1 [OF \<open>finite \<F>\<close> \<F> clo])
     show "X \<inter> Y \<subseteq> \<Union>\<H>" if XY: "X \<in> \<F>" "Y \<in> \<F>" and "\<not> X \<subseteq> Y" "\<not> Y \<subseteq> X" for X Y
@@ -856,7 +840,7 @@
       and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X"
       and contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> rel_frontier T"
   obtains g where "continuous_on (\<Union>\<F>) g"
-     "g ` (\<Union>\<F>) \<subseteq> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+     "g \<in> (\<Union>\<F>) \<rightarrow> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
 proof -
   obtain V g where "S \<subseteq> V" "open V" "continuous_on V g" and gim: "g \<in> V \<rightarrow> rel_frontier T" and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
     using neighbourhood_extension_into_ANR [OF contf fim _ \<open>closed S\<close>] ANR_rel_frontier_convex T by blast
@@ -877,12 +861,12 @@
   proof (rule extend_map_lemma [of \<G> "\<G> \<inter> Pow V" T g])
     show "continuous_on (\<Union>(\<G> \<inter> Pow V)) g"
       by (metis Union_Int_subset Union_Pow_eq \<open>continuous_on V g\<close> continuous_on_subset le_inf_iff)
-  qed (use \<open>finite \<G>\<close> T polyG affG faceG gim in fastforce)+
+  qed (use \<open>finite \<G>\<close> T polyG affG faceG gim image_subset_iff_funcset in auto)
   show ?thesis
   proof
     show "continuous_on (\<Union>\<F>) h"
       using \<open>\<Union>\<G> = \<Union>\<F>\<close> conth by auto
-    show "h ` \<Union>\<F> \<subseteq> rel_frontier T"
+    show "h \<in> \<Union>\<F> \<rightarrow> rel_frontier T"
       using \<open>\<Union>\<G> = \<Union>\<F>\<close> him by auto
     show "h x = f x" if "x \<in> S" for x
     proof -
@@ -910,7 +894,7 @@
       and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X"
       and contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> rel_frontier T"
   obtains C g where "finite C" "disjnt C S" "continuous_on (\<Union>\<F> - C) g"
-     "g ` (\<Union>\<F> - C) \<subseteq> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+     "g \<in> (\<Union>\<F> - C) \<rightarrow> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
 proof -
   obtain V g where "S \<subseteq> V" "open V" "continuous_on V g" and gim: "g \<in> V \<rightarrow> rel_frontier T" and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
     using neighbourhood_extension_into_ANR [OF contf fim _ \<open>closed S\<close>] ANR_rel_frontier_convex T by blast
@@ -926,15 +910,15 @@
     by (rule cell_complex_subdivision_exists [OF \<open>d>0\<close> \<open>finite \<F>\<close> poly aff face]) auto
   obtain C h where "finite C" and dis: "disjnt C (\<Union>(\<G> \<inter> Pow V))"
                and card: "card C \<le> card \<G>" and conth: "continuous_on (\<Union>\<G> - C) h"
-               and him: "h ` (\<Union>\<G> - C) \<subseteq> rel_frontier T"
+               and him: "h \<in> (\<Union>\<G> - C) \<rightarrow> rel_frontier T"
                and hg: "\<And>x. x \<in> \<Union>(\<G> \<inter> Pow V) \<Longrightarrow> h x = g x"
   proof (rule extend_map_lemma_cofinite [of \<G> "\<G> \<inter> Pow V" T g])
     show "continuous_on (\<Union>(\<G> \<inter> Pow V)) g"
       by (metis Union_Int_subset Union_Pow_eq \<open>continuous_on V g\<close> continuous_on_subset le_inf_iff)
-    show "g ` \<Union>(\<G> \<inter> Pow V) \<subseteq> rel_frontier T"
+    show "g \<in> \<Union>(\<G> \<inter> Pow V) \<rightarrow> rel_frontier T"
       using gim by force
   qed (auto intro: \<open>finite \<G>\<close> T polyG affG dest: faceG)
-  have Ssub: "S \<subseteq> \<Union>(\<G> \<inter> Pow V)"
+  have "S \<subseteq> \<Union>(\<G> \<inter> Pow V)"
   proof
     fix x
     assume "x \<in> S"
@@ -948,23 +932,8 @@
     then show "x \<in> \<Union>(\<G> \<inter> Pow V)"
       using \<open>X \<in> \<G>\<close> \<open>x \<in> X\<close> by blast
   qed
-  show ?thesis
-  proof
-    show "continuous_on (\<Union>\<F>-C) h"
-      using \<open>\<Union>\<G> = \<Union>\<F>\<close> conth by auto
-    show "h ` (\<Union>\<F> - C) \<subseteq> rel_frontier T"
-      using \<open>\<Union>\<G> = \<Union>\<F>\<close> him by auto
-    show "h x = f x" if "x \<in> S" for x
-    proof -
-      have "h x = g x"
-        using Ssub hg that by blast
-      also have "... = f x"
-        by (simp add: gf that)
-      finally show "h x = f x" .
-    qed
-    show "disjnt C S"
-      using dis Ssub  by (meson disjnt_iff subset_eq)
-  qed (intro \<open>finite C\<close>)
+  then show ?thesis
+    by (metis PowI Union_Pow_eq \<open>\<Union> \<G> = \<Union> \<F>\<close> \<open>finite C\<close> conth dis disjnt_Union2 gf hg him subsetD that)
 qed
 
 
@@ -1021,7 +990,7 @@
       by (metis aff aff_dim_subset inf_commute inf_le1 order_trans)
     obtain K g where K: "finite K" "disjnt K S"
                  and contg: "continuous_on (\<Union>{bbox \<inter> T} - K) g"
-                 and gim: "g ` (\<Union>{bbox \<inter> T} - K) \<subseteq> rel_frontier U"
+                 and gim: "g \<in> (\<Union>{bbox \<inter> T} - K) \<rightarrow> rel_frontier U"
                  and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
     proof (rule extend_map_cell_complex_to_sphere_cofinite
               [OF _ Ssub _ \<open>convex U\<close> \<open>bounded U\<close> _ _ _ contf fim])
@@ -1090,7 +1059,7 @@
       have "(g \<circ> closest_point (cbox (- c) c \<inter> T)) ` (T - K) \<subseteq> g ` (\<Union>{bbox \<inter> T} - K)"
         by (metis image_comp image_mono cpt_subset)
       also have "... \<subseteq> rel_frontier U"
-        by (rule gim)
+        using gim by blast
       finally show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) \<in> (T - K) \<rightarrow> rel_frontier U"
         by blast
       show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) x = f x" if "x \<in> S" for x
@@ -1112,7 +1081,7 @@
               order_trans [OF \<open>S \<subseteq> T\<close> hull_subset [of T affine]])
   then obtain K g where "finite K" "disjnt K S"
                and contg: "continuous_on (T - K) g"
-               and gim:  "g ` (T - K) \<subseteq> rel_frontier U"
+               and gim:  "g \<in> (T - K) \<rightarrow> rel_frontier U"
                and gf:   "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
     by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset)
   then show ?thesis
@@ -1195,7 +1164,7 @@
     ultimately have "frontier (cball a d) \<inter> U retract_of (U - {a})"
       by metis
     then obtain r where contr: "continuous_on (U - {a}) r"
-                    and rim: "r ` (U - {a}) \<subseteq> sphere a d"  "r ` (U - {a}) \<subseteq> U"
+                    and rim: "r \<in> (U - {a}) \<rightarrow> sphere a d"  "r \<in> (U - {a}) \<rightarrow> U"
                     and req: "\<And>x. x \<in> sphere a d \<inter> U \<Longrightarrow> r x = x"
       using \<open>affine U\<close> by (force simp: retract_of_def retraction_def hull_same)
     define j where "j \<equiv> \<lambda>x. if x \<in> ball a d then r x else x"
@@ -1212,8 +1181,7 @@
         using rim by auto
       then show "j y \<in> S \<union> C - ball a d"
         unfolding j_def
-        using \<open>r y \<in> sphere a d\<close> \<open>y \<in> U - {a}\<close> \<open>y \<in> S \<union> (C - {a})\<close> d rim
-        by (metis Diff_iff Int_iff Un_iff subsetD cball_diff_eq_sphere image_subset_iff)
+        using \<open>y \<in> S \<union> (C - {a})\<close> \<open>y \<in> U - {a}\<close> d rim(2) by auto
     qed
     have contj: "continuous_on (U - {a}) j"
       unfolding j_def Uaeq
@@ -1437,7 +1405,7 @@
                and gim: "g \<in> (T - K) \<rightarrow> rel_frontier U"
                and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
      using assms extend_map_affine_to_sphere_cofinite_simple by metis
-  have "(\<exists>y C. C \<in> components (T - S) \<and> x \<in> C \<and> y \<in> C \<and> y \<in> L)" if "x \<in> K" for x
+  have "\<exists>y C. C \<in> components (T - S) \<and> x \<in> C \<and> y \<in> C \<and> y \<in> L" if "x \<in> K" for x
   proof -
     have "x \<in> T-S"
       using \<open>K \<subseteq> T\<close> \<open>disjnt K S\<close> disjnt_def that by fastforce
@@ -1446,7 +1414,7 @@
     with ovlap [of C] show ?thesis
       by blast
   qed
-  then obtain \<xi> where \<xi>: "\<And>x. x \<in> K \<Longrightarrow> \<exists>C. C \<in> components (T - S) \<and> x \<in> C \<and> \<xi> x \<in> C \<and> \<xi> x \<in> L"
+  then obtain \<xi> where \<xi>: "\<And>x. x \<in> K \<Longrightarrow> \<exists>C. C \<in> components (T - S) \<and> x \<in> C \<and> \<xi> x \<in> C \<and> \<xi> x \<in> L" 
     by metis
   obtain h where conth: "continuous_on (T - \<xi> ` K) h"
              and him: "h \<in> (T - \<xi> ` K) \<rightarrow> rel_frontier U"
@@ -1525,7 +1493,7 @@
   define LU where "LU \<equiv> L \<union> (\<Union> {C \<in> components (T - S). \<not>bounded C} - cbox (-(b+One)) (b+One))"
   obtain K g where "finite K" "K \<subseteq> LU" "K \<subseteq> T" "disjnt K S"
                and contg: "continuous_on (T - K) g"
-               and gim: "g ` (T - K) \<subseteq> rel_frontier U"
+               and gim: "g \<in> (T - K) \<rightarrow> rel_frontier U"
                and gf:  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
   proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim])
     show "C \<inter> LU \<noteq> {}" if "C \<in> components (T - S)" for C
@@ -1577,7 +1545,7 @@
     proof (cases "x \<in> cbox (- c) c")
       case True
       with \<open>x \<in> T\<close> show ?thesis
-        using cbsub(3) Knot  by (force simp: closest_point_self)
+        using cbsub(3) Knot by (force simp: closest_point_self)
     next
       case False
       have clo_in_rf: "closest_point (cbox (- c) c \<inter> T) x \<in> rel_frontier (cbox (- c) c \<inter> T)"
@@ -1617,7 +1585,7 @@
       by (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]; force)
     have "g (closest_point (cbox (- c) c \<inter> T) x) \<in> rel_frontier U"
          if "x \<in> T" "x \<in> K \<longrightarrow> x \<notin> cbox (- b - One) (b + One)" for x
-      using gim [THEN subsetD] that cloTK by blast
+      using cloTK gim that by auto
     then show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) \<in> (T - K \<inter> cbox (- (b + One)) (b + One))
                \<rightarrow> rel_frontier U"
       by force
@@ -1663,8 +1631,7 @@
       and "\<And>C. \<lbrakk>C \<in> components(- S); bounded C\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
   obtains K g where "finite K" "K \<subseteq> L" "disjnt K S" "continuous_on (- K) g"
                     "g \<in> (- K) \<rightarrow> sphere a r" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-  using extend_map_affine_to_sphere_cofinite
-        [OF \<open>compact S\<close> affine_UNIV subset_UNIV] assms
+  using assms extend_map_affine_to_sphere_cofinite [OF \<open>compact S\<close> affine_UNIV subset_UNIV] 
   by (metis Compl_eq_Diff_UNIV aff_dim_UNIV of_nat_le_iff)
 
 corollary extend_map_UNIV_to_sphere_no_bounded_component:
@@ -1789,14 +1756,7 @@
     by auto
   moreover
   have "connected (- S) = connected (- sphere a r)"
-  proof (rule homotopy_eqv_separation)
-    show "S homotopy_eqv sphere a r"
-      using hom homeomorphic_imp_homotopy_eqv by blast
-    show "compact (sphere a r)"
-      by simp
-    then show " compact S"
-      using hom homeomorphic_compactness by blast
-  qed
+    by (meson hom compact_sphere homeomorphic_compactness homeomorphic_imp_homotopy_eqv homotopy_eqv_separation)
   ultimately show ?thesis
     using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \<open>0 < r\<close>)
 qed
@@ -1831,7 +1791,7 @@
       using S by (auto simp: homeomorphic_def)
     show "connected (- T)" if "closed T" "T \<subset> S" for T
     proof -
-      have "f ` T \<subseteq> sphere a r"
+      have "f \<in> T \<rightarrow> sphere a r"
         using \<open>T \<subset> S\<close> hom homeomorphism_image1 by blast
       moreover have "f ` T \<noteq> sphere a r"
         using \<open>T \<subset> S\<close> hom
@@ -1844,7 +1804,7 @@
       moreover then have "compact (f ` T)"
         by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \<open>T \<subset> S\<close>)
       moreover have "T homotopy_eqv f ` T"
-        by (meson \<open>f ` T \<subseteq> sphere a r\<close> dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \<open>T \<subset> S\<close>)
+        by (meson hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets order.refl psubsetE that(2))
       ultimately show ?thesis
         using homotopy_eqv_separation [of T "f`T"] by blast
     qed
@@ -1986,7 +1946,7 @@
 
 lemma inv_of_domain_ss0:
   fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
-  assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S"
+  assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f \<in> U \<rightarrow> S"
       and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)"
       and ope: "openin (top_of_set S) U"
     shows "openin (top_of_set S) (f ` U)"
@@ -2019,19 +1979,20 @@
     show "open (k ` U)"
       by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope])
     show "inj_on (k \<circ> f \<circ> h) (k ` U)"
-      by (smt (verit) comp_apply inj_on_def \<open>U \<subseteq> S\<close> fim homeomorphism_apply2 homhk image_iff injf subsetD)
+      unfolding inj_on_def 
+      by (smt (verit, ccfv_threshold) PiE \<open>U \<subseteq> S\<close> assms(3) comp_apply homeomorphism_def homhk imageE inj_on_def injf subset_eq)
   qed
   moreover
   have eq: "f ` U = h ` (k \<circ> f \<circ> h \<circ> k) ` U"
     unfolding image_comp [symmetric] using \<open>U \<subseteq> S\<close> fim
-    by (metis homeomorphism_image2 homeomorphism_of_subsets homkh subset_image_iff)
+    by (metis homeomorphism_image2 homeomorphism_of_subsets homhk homkh image_subset_iff_funcset top_greatest)
   ultimately show ?thesis
     by (metis (no_types, opaque_lifting) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV)
 qed
 
 lemma inv_of_domain_ss1:
   fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
-  assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S"
+  assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f \<in> U \<rightarrow> S"
       and "subspace S"
       and ope: "openin (top_of_set S) U"
     shows "openin (top_of_set S) (f ` U)"
@@ -2045,7 +2006,7 @@
     show "continuous_on (U \<times> S') g"
       unfolding  g_def
       by (auto intro!: continuous_intros continuous_on_compose2 [OF contf continuous_on_fst])
-    show "g ` (U \<times> S') \<subseteq> S \<times> S'"
+    show "g \<in> (U \<times> S') \<rightarrow> S \<times> S'"
       using fim  by (auto simp: g_def)
     show "inj_on g (U \<times> S')"
       using injf by (auto simp: g_def inj_on_def)
@@ -2073,7 +2034,7 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes ope: "openin (top_of_set U) S"
       and "subspace U" "subspace V" and VU: "dim V \<le> dim U"
-      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
+      and contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> V"
       and injf: "inj_on f S"
     shows "openin (top_of_set V) (f ` S)"
 proof -
@@ -2087,7 +2048,7 @@
   have eq: "f ` S = k ` (h \<circ> f) ` S"
   proof -
     have "k ` h ` f ` S = f ` S"
-      by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl)
+      by (meson equalityD2 fim funcset_image homeomorphism_image2 homeomorphism_of_subsets homhk)
     then show ?thesis
       by (simp add: image_comp)
   qed
@@ -2101,11 +2062,11 @@
     moreover have "openin (top_of_set U) ((h \<circ> f) ` S)"
     proof (rule inv_of_domain_ss1)
       show "continuous_on S (h \<circ> f)"
-        by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk)
+        by (meson contf continuous_on_compose continuous_on_subset fim funcset_image homeomorphism_cont2 homkh)
       show "inj_on (h \<circ> f) S"
-        by (smt (verit, ccfv_SIG) comp_apply fim inj_on_def homeomorphism_apply2 [OF homkh] image_subset_iff injf)
-      show "(h \<circ> f) ` S \<subseteq> U"
-        using \<open>V' \<subseteq> U\<close> hfV' by auto
+        by (smt (verit, ccfv_SIG) Pi_iff comp_apply fim homeomorphism_apply2 homkh inj_on_def injf)
+      show "h \<circ> f \<in> S \<rightarrow> U"
+        using \<open>V' \<subseteq> U\<close> hfV' by blast
       qed (auto simp: assms)
     ultimately show "openin (top_of_set V') ((h \<circ> f) ` S)"
       using openin_subset_trans \<open>V' \<subseteq> U\<close> by force
@@ -2116,7 +2077,7 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes ope: "openin (top_of_set U) S"
       and "subspace U" "subspace V"
-      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
+      and contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> V"
       and injf: "inj_on f S" and "S \<noteq> {}"
     shows "dim U \<le> dim V"
 proof -
@@ -2130,11 +2091,12 @@
     then obtain h k where homhk: "homeomorphism V T h k"
       using homeomorphic_def  by blast
     have "continuous_on S (h \<circ> f)"
-      by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk)
+      by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_def homhk image_subset_iff_funcset)
     moreover have "(h \<circ> f) ` S \<subseteq> U"
       using \<open>T \<subseteq> U\<close> fim homeomorphism_image1 homhk by fastforce
     moreover have "inj_on (h \<circ> f) S"
-      by (smt (verit, best) comp_apply inj_on_def fim homeomorphism_apply1 homhk image_subset_iff injf)
+      unfolding inj_on_def
+      by (metis Pi_iff comp_apply fim homeomorphism_def homhk inj_on_def injf)
     ultimately have ope_hf: "openin (top_of_set U) ((h \<circ> f) ` S)"
       using invariance_of_domain_subspaces [OF ope \<open>subspace U\<close> \<open>subspace U\<close>] by blast
     have "(h \<circ> f) ` S \<subseteq> T"
@@ -2155,7 +2117,7 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes ope: "openin (top_of_set U) S"
       and aff: "affine U" "affine V" "aff_dim V \<le> aff_dim U"
-      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
+      and contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> V"
       and injf: "inj_on f S"
     shows "openin (top_of_set V) (f ` S)"
 proof (cases "S = {}")
@@ -2177,7 +2139,7 @@
       by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff)
     show "continuous_on ((+) (- a) ` S) ((+) (- b) \<circ> f \<circ> (+) a)"
       by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois)
-    show "((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S \<subseteq> (+) (- b) ` V"
+    show "(+) (- b) \<circ> f \<circ> (+) a \<in> (+) (- a) ` S \<rightarrow> (+) (- b) ` V"
       using fim by auto
     show "inj_on ((+) (- b) \<circ> f \<circ> (+) a) ((+) (- a) ` S)"
       by (auto simp: inj_on_def) (meson inj_onD injf)
@@ -2190,7 +2152,7 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes ope: "openin (top_of_set U) S"
       and aff: "affine U" "affine V"
-      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
+      and contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> V"
       and injf: "inj_on f S" and "S \<noteq> {}"
     shows "aff_dim U \<le> aff_dim V"
 proof -
@@ -2206,7 +2168,7 @@
       by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace_subtract \<open>affine V\<close> cong: image_cong_simp)
     show "continuous_on ((+) (- a) ` S) ((+) (- b) \<circ> f \<circ> (+) a)"
       by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois)
-    show "((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S \<subseteq> (+) (- b) ` V"
+    show "(+) (- b) \<circ> f \<circ> (+) a \<in> (+) (- a) ` S \<rightarrow> (+) (- b) ` V"
       using fim by auto
     show "inj_on ((+) (- b) \<circ> f \<circ> (+) a) ((+) (- a) ` S)"
       by (auto simp: inj_on_def) (meson inj_onD injf)
@@ -2227,7 +2189,7 @@
 corollary continuous_injective_image_subspace_dim_le:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "subspace S" "subspace T"
-      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> T"
+      and contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> T"
       and injf: "inj_on f S"
     shows "dim S \<le> dim T"
   using invariance_of_dimension_subspaces [of S S _ f] assms by (auto simp: subspace_affine)
@@ -2235,7 +2197,7 @@
 lemma invariance_of_dimension_convex_domain:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "convex S"
-      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> affine hull T"
+      and contf: "continuous_on S f" and fim: "f \<in> S \<rightarrow> affine hull T"
       and injf: "inj_on f S"
     shows "aff_dim S \<le> aff_dim T"
 proof (cases "S = {}")
@@ -2249,7 +2211,7 @@
       by (simp add: openin_rel_interior)
     show "continuous_on (rel_interior S) f"
       using contf continuous_on_subset rel_interior_subset by blast
-    show "f ` rel_interior S \<subseteq> affine hull T"
+    show "f \<in> rel_interior S \<rightarrow> affine hull T"
       using fim rel_interior_subset by blast
     show "inj_on f (rel_interior S)"
       using inj_on_subset injf rel_interior_subset by blast
@@ -2271,8 +2233,8 @@
   proof (rule invariance_of_dimension_convex_domain [OF \<open>convex S\<close>])
     show "continuous_on S h"
       using homeomorphism_def homhk by blast
-    show "h ` S \<subseteq> affine hull T"
-      by (metis homeomorphism_def homhk hull_subset)
+    show "h \<in> S \<rightarrow> affine hull T"
+      using homeomorphism_image1 homhk hull_subset by fastforce
     show "inj_on h S"
       by (meson homeomorphism_apply1 homhk inj_on_inverseI)
   qed
@@ -2342,13 +2304,13 @@
 lemma continuous_image_subset_interior:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
-  shows "f ` (interior S) \<subseteq> interior(f ` S)"
+  shows "f \<in> (interior S) \<rightarrow> interior(f ` S)"
 proof -
   have "open (f ` interior S)"
     using assms
     by (intro invariance_of_domain_gen) (auto simp: subset_inj_on interior_subset continuous_on_subset)
   then show ?thesis
-    by (simp add: image_mono interior_maximal interior_subset)
+    by (simp add: image_mono interiorI interior_subset)
 qed
 
 lemma homeomorphic_interiors_same_dimension:
@@ -2363,9 +2325,9 @@
      and contf: "continuous_on S f" and contg: "continuous_on T g"
   then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
     by (auto simp: inj_on_def intro: rev_image_eqI) metis+
-  have fim: "f ` interior S \<subseteq> interior T"
+  have fim: "f \<in> interior S \<rightarrow> interior T"
     using continuous_image_subset_interior [OF contf \<open>inj_on f S\<close>] dimeq fST by simp
-  have gim: "g ` interior T \<subseteq> interior S"
+  have gim: "g \<in> interior T \<rightarrow> interior S"
     using continuous_image_subset_interior [OF contg \<open>inj_on g T\<close>] dimeq gTS by simp
   show "homeomorphism (interior S) (interior T) f g"
     unfolding homeomorphism_def
@@ -2437,16 +2399,14 @@
      and contf: "continuous_on S f" and contg: "continuous_on T g"
   then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
     by (auto simp: inj_on_def intro: rev_image_eqI) metis+
-  have "g ` interior T \<subseteq> interior S"
+  have "g \<in> interior T \<rightarrow> interior S"
     using continuous_image_subset_interior [OF contg \<open>inj_on g T\<close>] dimeq gTS by simp
   then have fim: "f ` frontier S \<subseteq> frontier T"
-    unfolding frontier_def
-    using continuous_image_subset_interior assms(2) assms(3) S by auto
-  have "f ` interior S \<subseteq> interior T"
+    unfolding frontier_def using Pi_mem S assms by fastforce
+  have "f \<in> interior S \<rightarrow> interior T"
     using continuous_image_subset_interior [OF contf \<open>inj_on f S\<close>] dimeq fST by simp
   then have gim: "g ` frontier T \<subseteq> frontier S"
-    unfolding frontier_def
-    using continuous_image_subset_interior T assms(2) assms(3) by auto
+    unfolding frontier_def using Pi_mem T assms by fastforce
   show "homeomorphism (frontier S) (frontier T) f g"
     unfolding homeomorphism_def
   proof (intro conjI ballI)
@@ -2500,9 +2460,10 @@
 
 lemma continuous_image_subset_rel_interior:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
+  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f \<in> S \<rightarrow> T"
       and TS: "aff_dim T \<le> aff_dim S"
-  shows "f ` (rel_interior S) \<subseteq> rel_interior(f ` S)"
+    shows "f \<in> (rel_interior S) \<rightarrow> rel_interior(f ` S)"
+unfolding image_subset_iff_funcset [symmetric]
 proof (rule rel_interior_maximal)
   show "f ` rel_interior S \<subseteq> f ` S"
     by(simp add: image_mono rel_interior_subset)
@@ -2511,9 +2472,9 @@
     show "openin (top_of_set (affine hull S)) (rel_interior S)"
       by (simp add: openin_rel_interior)
     show "aff_dim (affine hull f ` S) \<le> aff_dim (affine hull S)"
-      by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans)
-    show "f ` rel_interior S \<subseteq> affine hull f ` S"
-      by (meson \<open>f ` rel_interior S \<subseteq> f ` S\<close> hull_subset order_trans)
+      by (metis TS aff_dim_affine_hull aff_dim_subset fim image_subset_iff_funcset order_trans)
+    show "f \<in> rel_interior S \<rightarrow> affine hull f ` S"
+      using \<open>f ` rel_interior S \<subseteq> f ` S\<close> hull_subset by fastforce
     show "continuous_on (rel_interior S) f"
       using contf continuous_on_subset rel_interior_subset by blast
     show "inj_on f (rel_interior S)"
@@ -2533,10 +2494,10 @@
      and contf: "continuous_on S f" and contg: "continuous_on T g"
   then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
     by (auto simp: inj_on_def intro: rev_image_eqI) metis+
-  have fim: "f ` rel_interior S \<subseteq> rel_interior T"
-    by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
-  have gim: "g ` rel_interior T \<subseteq> rel_interior S"
-    by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior gTS order_refl)
+  have fim: "f \<in> rel_interior S \<rightarrow> rel_interior T"
+    by (smt (verit, best) PiE Pi_I S \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST)
+  have gim: "g \<in> rel_interior T \<rightarrow> rel_interior S"
+    by (metis T \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior dual_order.refl funcsetI gTS)
   show "homeomorphism (rel_interior S) (rel_interior T) f g"
     unfolding homeomorphism_def
   proof (intro conjI ballI)
@@ -2553,7 +2514,7 @@
         by (metis fg \<open>x \<in> rel_interior T\<close> imageI)
     qed
     moreover have "f ` rel_interior S \<subseteq> rel_interior T"
-      by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
+      using fim by blast
     ultimately show "f ` rel_interior S = rel_interior T"
       by blast
     show "continuous_on (rel_interior S) f"
@@ -2587,8 +2548,8 @@
   proof (rule invariance_of_dimension_affine_sets)
     show "continuous_on (rel_interior S) f"
       using contf continuous_on_subset rel_interior_subset by blast
-    show "f ` rel_interior S \<subseteq> affine hull T"
-      by (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD)
+    show "f \<in> rel_interior S \<rightarrow> affine hull T"
+      by (simp add: S hull_inc mem_rel_interior_ball)
     show "inj_on f (rel_interior S)"
       by (metis S inj_on_inverseI inj_on_subset rel_interior_subset)
   qed (simp_all add: openin_rel_interior assms)
@@ -2625,10 +2586,10 @@
      and contf: "continuous_on S f" and contg: "continuous_on T g"
   then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
     by (auto simp: inj_on_def intro: rev_image_eqI) metis+
-  have fim: "f ` rel_interior S \<subseteq> rel_interior T"
-    by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
-  have gim: "g ` rel_interior T \<subseteq> rel_interior S"
-    by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior gTS order_refl)
+  have fim: "f \<in> rel_interior S \<rightarrow> rel_interior T"
+    by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior dual_order.refl fST image_subset_iff_funcset)
+  have gim: "g \<in> rel_interior T \<rightarrow> rel_interior S"
+    by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior dual_order.refl gTS image_subset_iff_funcset)
   show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g"
     unfolding homeomorphism_def
   proof (intro conjI ballI)
@@ -2637,11 +2598,11 @@
     show fg: "\<And>y. y \<in> T - rel_interior T \<Longrightarrow> f (g y) = y"
       using T mem_rel_interior_ball by blast
     show "f ` (S - rel_interior S) = T - rel_interior T"
-      using S fST fim gim by auto
+      using S fST fim gim image_subset_iff_funcset by fastforce
     show "continuous_on (S - rel_interior S) f"
       using contf continuous_on_subset rel_interior_subset by blast
     show "g ` (T - rel_interior T) = S - rel_interior S"
-      using T gTS gim fim by auto
+      using T gTS gim fim image_subset_iff_funcset by fastforce
     show "continuous_on (T - rel_interior T) g"
       using contg continuous_on_subset rel_interior_subset by blast
   qed
@@ -2744,7 +2705,7 @@
                           (q \<circ> (\<lambda>z. (Arg2pi z / (2 * pi))))"
 proof -
   obtain h where conth: "continuous_on ({0..1::real} \<times> {0..1}) h"
-             and him: "h ` ({0..1} \<times> {0..1}) \<subseteq> S"
+             and him: "h \<in> ({0..1} \<times> {0..1}) \<rightarrow> S"
              and h0: "(\<forall>x. h (0, x) = p x)"
              and h1: "(\<forall>x. h (1, x) = q x)"
              and h01: "(\<forall>t\<in>{0..1}. h (t, 1) = h (t, 0)) "
@@ -2795,13 +2756,13 @@
 lemma simply_connected_eq_homotopic_circlemaps1:
   fixes f :: "complex \<Rightarrow> 'a::topological_space" and g :: "complex \<Rightarrow> 'a"
   assumes S: "simply_connected S"
-      and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \<subseteq> S"
-      and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \<subseteq> S"
+      and contf: "continuous_on (sphere 0 1) f" and fim: "f \<in> (sphere 0 1) \<rightarrow> S"
+      and contg: "continuous_on (sphere 0 1) g" and gim: "g \<in> (sphere 0 1) \<rightarrow> S"
     shows "homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g"
 proof -
   let ?h = "(\<lambda>t. complex_of_real (2 * pi * t) * \<i>)"
   have "homotopic_loops S (f \<circ> exp \<circ> ?h) (f \<circ> exp \<circ> ?h) \<and> homotopic_loops S (g \<circ> exp \<circ> ?h) (g \<circ> exp \<circ> ?h)"
-    by (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim)
+    by (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim image_subset_iff_funcset)
   then have "homotopic_loops S (f \<circ> exp \<circ> ?h) (g \<circ> exp \<circ> ?h)"
     using S simply_connected_homotopic_loops by blast
   then show ?thesis
@@ -2812,19 +2773,19 @@
 
 lemma simply_connected_eq_homotopic_circlemaps2a:
   fixes h :: "complex \<Rightarrow> 'a::topological_space"
-  assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \<subseteq> S"
+  assumes conth: "continuous_on (sphere 0 1) h" and him: "h \<in> sphere 0 1 \<rightarrow> S"
       and hom: "\<And>f g::complex \<Rightarrow> 'a.
-                \<lbrakk>continuous_on (sphere 0 1) f; f ` (sphere 0 1) \<subseteq> S;
-                continuous_on (sphere 0 1) g; g ` (sphere 0 1) \<subseteq> S\<rbrakk>
+                \<lbrakk>continuous_on (sphere 0 1) f; f \<in> (sphere 0 1) \<rightarrow> S;
+                continuous_on (sphere 0 1) g; g \<in> (sphere 0 1) \<rightarrow> S\<rbrakk>
                 \<Longrightarrow> homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g"
             shows "\<exists>a. homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S h (\<lambda>x. a)"
-  by (metis conth continuous_on_const him hom image_subset_iff)
+  by (metis conth continuous_on_const him hom image_subset_iff image_subset_iff_funcset)
 
 lemma simply_connected_eq_homotopic_circlemaps2b:
   fixes S :: "'a::real_normed_vector set"
   assumes "\<And>f g::complex \<Rightarrow> 'a.
-                \<lbrakk>continuous_on (sphere 0 1) f; f ` (sphere 0 1) \<subseteq> S;
-                continuous_on (sphere 0 1) g; g ` (sphere 0 1) \<subseteq> S\<rbrakk>
+                \<lbrakk>continuous_on (sphere 0 1) f; f \<in> (sphere 0 1) \<rightarrow> S;
+                continuous_on (sphere 0 1) g; g \<in> (sphere 0 1) \<rightarrow> S\<rbrakk>
                 \<Longrightarrow> homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g"
   shows "path_connected S"
 proof (clarsimp simp add: path_connected_eq_homotopic_points)
@@ -2869,10 +2830,10 @@
   fixes S :: "'a::real_normed_vector set"
   shows "simply_connected S \<longleftrightarrow>
          (\<forall>f g::complex \<Rightarrow> 'a.
-              continuous_on (sphere 0 1) f \<and> f ` (sphere 0 1) \<subseteq> S \<and>
-              continuous_on (sphere 0 1) g \<and> g ` (sphere 0 1) \<subseteq> S
+              continuous_on (sphere 0 1) f \<and> f \<in> (sphere 0 1) \<rightarrow> S \<and>
+              continuous_on (sphere 0 1) g \<and> g \<in> (sphere 0 1) \<rightarrow> S
               \<longrightarrow> homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g)"
-  by (metis simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a 
+  by (metis image_subset_iff_funcset simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a
       simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3)
 
 proposition simply_connected_eq_contractible_circlemap:
@@ -2882,7 +2843,7 @@
          (\<forall>f::complex \<Rightarrow> 'a.
               continuous_on (sphere 0 1) f \<and> f `(sphere 0 1) \<subseteq> S
               \<longrightarrow> (\<exists>a. homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f (\<lambda>x. a)))"
-  by (metis simply_connected_eq_homotopic_circlemaps simply_connected_eq_homotopic_circlemaps2a 
+  by (metis image_subset_iff_funcset simply_connected_eq_homotopic_circlemaps simply_connected_eq_homotopic_circlemaps2a 
       simply_connected_eq_homotopic_circlemaps3 simply_connected_imp_path_connected)
 
 corollary homotopy_eqv_simple_connectedness:
@@ -2966,12 +2927,7 @@
   then obtain f g where hom: "homeomorphism S T f g"
     using homeomorphic_def by blast
   show "dim S = dim T"
-  proof (rule order_antisym)
-    show "dim S \<le> dim T"
-      by (metis assms dual_order.refl inj_onI homeomorphism_cont1 [OF hom] homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] continuous_injective_image_subspace_dim_le)
-    show "dim T \<le> dim S"
-      by (metis assms dual_order.refl inj_onI homeomorphism_cont2 [OF hom] homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] continuous_injective_image_subspace_dim_le)
-  qed
+    by (metis \<open>S homeomorphic T\<close> aff_dim_subspace assms homeomorphic_convex_sets of_nat_eq_iff subspace_imp_convex)
 next
   assume "dim S = dim T"
   then show "S homeomorphic T"
@@ -3034,11 +2990,11 @@
   qed
 qed
 
-subsection\<open>more invariance of domain\<close>(*FIX ME title? *)
+subsection\<open>more invariance of domain\<close>  (*FIX ME title? *)
 
 proposition invariance_of_domain_sphere_affine_set_gen:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
+  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f \<in> S \<rightarrow> T"
       and U: "bounded U" "convex U"
       and "affine T" and affTU: "aff_dim T < aff_dim U"
       and ope: "openin (top_of_set (rel_frontier U)) S"
@@ -3085,8 +3041,8 @@
     qed (use contf continuous_on_subset hgsub in blast)
     show "inj_on (f \<circ> h) (g ` (S - {b}))"
       by (smt (verit, del_insts) SU homeomorphism_def inj_on_def injf gh Diff_iff comp_apply imageE subset_iff)
-    show "(f \<circ> h) ` g ` (S - {b}) \<subseteq> T"
-      by (metis fim image_comp image_mono hgsub subset_trans)
+    show "f \<circ> h \<in> g ` (S - {b}) \<rightarrow> T"
+      using fim hgsub by fastforce
   qed (auto simp: assms)
   moreover
   have "openin (top_of_set T) ((f \<circ> k) ` j ` (S - {c}))"
@@ -3100,20 +3056,15 @@
     qed (use contf continuous_on_subset kjsub in blast)
     show "inj_on (f \<circ> k) (j ` (S - {c}))"
       by (smt (verit, del_insts) SU homeomorphism_def inj_on_def injf jk Diff_iff comp_apply imageE subset_iff)
-    show "(f \<circ> k) ` j ` (S - {c}) \<subseteq> T"
-      by (metis fim image_comp image_mono kjsub subset_trans)
+    show "f \<circ> k \<in> j ` (S - {c}) \<rightarrow> T"
+      using fim kjsub by fastforce
   qed (auto simp: assms)
   ultimately have "openin (top_of_set T) ((f \<circ> h) ` g ` (S - {b}) \<union> ((f \<circ> k) ` j ` (S - {c})))"
     by (rule openin_Un)
   moreover have "(f \<circ> h) ` g ` (S - {b}) = f ` (S - {b})"
   proof -
     have "h ` g ` (S - {b}) = (S - {b})"
-    proof
-      show "h ` g ` (S - {b}) \<subseteq> S - {b}"
-        using homeomorphism_apply1 [OF gh] SU by (fastforce simp add: image_iff image_subset_iff)
-      show "S - {b} \<subseteq> h ` g ` (S - {b})"
-        by (metis Diff_mono SU gh homeomorphism_image2 homeomorphism_of_subsets set_eq_subset)
-    qed
+      by (meson Diff_mono Diff_subset SU gh homeomorphism_def homeomorphism_of_subsets subset_singleton_iff)
     then show ?thesis
       by (metis image_comp)
   qed
@@ -3134,7 +3085,7 @@
 
 lemma invariance_of_domain_sphere_affine_set:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
+  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f \<in> S \<rightarrow> T"
       and "r \<noteq> 0" "affine T" and affTU: "aff_dim T < DIM('a)"
       and ope: "openin (top_of_set (sphere a r)) S"
    shows "openin (top_of_set T) (f ` S)"
@@ -3166,9 +3117,10 @@
     then have "\<not> open (f ` sphere a r)"
       using compact_open
       by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty)
-    then show False
-      using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \<open>r > 0\<close>
-      by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that)
+    then have "r=0"
+      by (metis Pi_I UNIV_I aff_dim_UNIV affine_UNIV contf injf invariance_of_domain_sphere_affine_set
+               of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that)
+    with \<open>r>0\<close> show False by auto
   qed
   then show ?thesis
     using not_less by blast
@@ -3712,10 +3664,10 @@
   assume L: ?lhs
   then obtain g where contg: "continuous_on S g" and g: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
     using inessential_imp_continuous_logarithm_circle by blast
-  have "f ` S \<subseteq> sphere 0 1"
-    by (metis L homotopic_with_imp_subset1)
+  have "f \<in> S \<rightarrow> sphere 0 1"
+    by (metis L image_subset_iff_funcset homotopic_with_imp_subset1)
   then have "\<And>x. x \<in> S \<Longrightarrow> Re (g x) = 0"
-    using g by auto
+    using g by (simp add: Pi_iff)
   then show ?rhs
     by (rule_tac x="Im \<circ> g" in exI) (auto simp: Euler g intro: contg continuous_intros)
 next
@@ -3747,7 +3699,7 @@
     shows "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x * h x) (\<lambda>x. g x * h x)"
 proof -
   obtain k where contk: "continuous_on ({0..1::real} \<times> S) k"
-             and kim: "k ` ({0..1} \<times> S) \<subseteq> sphere 0 1"
+             and kim: "k \<in> ({0..1} \<times> S) \<rightarrow> sphere 0 1"
              and k0:  "\<And>x. k(0, x) = f x"
              and k1: "\<And>x. k(1, x) = g x"
     using hom by (auto simp: homotopic_with_def)
@@ -3760,8 +3712,8 @@
 proposition homotopic_circlemaps_divide:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
     shows "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) f g \<longleftrightarrow>
-           continuous_on S f \<and> f ` S \<subseteq> sphere 0 1 \<and>
-           continuous_on S g \<and> g ` S \<subseteq> sphere 0 1 \<and>
+           continuous_on S f \<and> f \<in> S \<rightarrow> sphere 0 1 \<and>
+           continuous_on S g \<and> g \<in> S \<rightarrow> sphere 0 1 \<and>
            (\<exists>c. homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c))"
 proof -
   have "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"
@@ -3780,8 +3732,8 @@
                 homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"
     by auto
   have "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) f g \<longleftrightarrow>
-           continuous_on S f \<and> f ` S \<subseteq> sphere 0 1 \<and>
-           continuous_on S g \<and> g ` S \<subseteq> sphere 0 1 \<and>
+           continuous_on S f \<and> f \<in> S \<rightarrow> sphere 0 1 \<and>
+           continuous_on S g \<and> g \<in> S \<rightarrow> sphere 0 1 \<and>
            homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"
         (is "?lhs \<longleftrightarrow> ?rhs")
   proof
@@ -3802,7 +3754,7 @@
       using homotopic_with_eq [OF homotopic_with_sphere_times [OF L cont]]
       by (auto simp: divide_inverse norm_inverse)
     with L show ?rhs
-      by (simp add: homotopic_with_imp_continuous homotopic_with_imp_subset1 homotopic_with_imp_subset2)
+      by (simp add: homotopic_with_imp_continuous homotopic_with_imp_funspace1)
   next
     assume ?rhs then show ?lhs
       by (elim conjE homotopic_with_eq [OF homotopic_with_sphere_times]; force)
@@ -3882,7 +3834,7 @@
 qed
 
 lemma open_map_iff_lower_hemicontinuous_preimage:
-  assumes "f ` S \<subseteq> T"
+  assumes "f \<in> S \<rightarrow> T"
     shows "((\<forall>U. openin (top_of_set S) U
                  \<longrightarrow> openin (top_of_set T) (f ` U)) \<longleftrightarrow>
             (\<forall>U. closedin (top_of_set S) U
@@ -3913,7 +3865,7 @@
 qed
 
 lemma closed_map_iff_upper_hemicontinuous_preimage:
-  assumes "f ` S \<subseteq> T"
+  assumes "f \<in> S \<rightarrow> T"
     shows "((\<forall>U. closedin (top_of_set S) U
                  \<longrightarrow> closedin (top_of_set T) (f ` U)) \<longleftrightarrow>
             (\<forall>U. openin (top_of_set S) U
@@ -4519,11 +4471,11 @@
   assume contf: "continuous_on (S \<union> T) f" and 0: "\<forall>i\<in>S \<union> T. f i \<noteq> 0"
   then have contfS: "continuous_on S f" and contfT: "continuous_on T f"
     using continuous_on_subset by auto
-  have "\<lbrakk>continuous_on S f; f ` S \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))"
+  have "\<lbrakk>continuous_on S f; f \<in> S \<rightarrow> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))"
     using BS by (auto simp: Borsukian_continuous_logarithm)
   then obtain g where contg: "continuous_on S g" and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
     using "0" contfS by force
-  have "\<lbrakk>continuous_on T f; f ` T \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on T g \<and> (\<forall>x \<in> T. f x = exp(g x))"
+  have "\<lbrakk>continuous_on T f; f \<in> T \<rightarrow> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on T g \<and> (\<forall>x \<in> T. f x = exp(g x))"
     using BT by (auto simp: Borsukian_continuous_logarithm)
   then obtain h where conth: "continuous_on T h" and fh: "\<And>x. x \<in> T \<Longrightarrow> f x = exp(h x)"
     using "0" contfT by force
@@ -4718,12 +4670,12 @@
     proof (rule upper_lower_hemicontinuous_explicit [of T "\<lambda>y. {z \<in> S. f z = y}" S])
       show "\<And>U. openin (top_of_set S) U
                  \<Longrightarrow> openin (top_of_set T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
-        using closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]]
-        by (simp add: Abstract_Topology_2.continuous_imp_closed_map \<open>compact S\<close> contf fim)
+        using closed_map_iff_upper_hemicontinuous_preimage [of f S T] fim contf \<open>compact S\<close>
+        using Abstract_Topology_2.continuous_imp_closed_map by blast
       show "\<And>U. closedin (top_of_set S) U \<Longrightarrow>
                  closedin (top_of_set T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
-        using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]]
-        by meson
+        using ope open_map_iff_lower_hemicontinuous_preimage[of f S T] fim [THEN equalityD1]
+        by blast
       show "bounded {z \<in> S. f z = y}"
         by (metis (no_types, lifting) compact_imp_bounded [OF \<open>compact S\<close>] bounded_subset mem_Collect_eq subsetI)
     qed (use \<open>y \<in> T\<close> \<open>0 < d\<close> fk kTS in \<open>force+\<close>)
@@ -4918,7 +4870,7 @@
     assume "connected U" "connected V" and T: "T = U \<union> V"
       and cloU: "closedin (top_of_set T) U"
       and cloV: "closedin (top_of_set T) V"
-    have "f ` (g ` U \<inter> g ` V) \<subseteq> U" "f ` (g ` U \<inter> g ` V) \<subseteq> V"
+    have "f \<in> (g ` U \<inter> g ` V) \<rightarrow> U" "f \<in> (g ` U \<inter> g ` V) \<rightarrow> V"
       using gf fim T by auto (metis UnCI image_iff)+
     moreover have "U \<inter> V \<subseteq> f ` (g ` U \<inter> g ` V)"
       using gf fim by (force simp: image_iff T)
@@ -5393,14 +5345,8 @@
 lemma nonseparation_by_component_eq:
   fixes S :: "'a :: euclidean_space set"
   assumes "open S \<or> closed S"
-  shows "((\<forall>C \<in> components S. connected(-C)) \<longleftrightarrow> connected(- S))" (is "?lhs = ?rhs")
-proof
-  assume ?lhs with assms show ?rhs
-    by (meson separation_by_component_closed separation_by_component_open)
-next
-  assume ?rhs with assms show ?lhs
-    using component_complement_connected by force
-qed
+  shows "((\<forall>C \<in> components S. connected(-C)) \<longleftrightarrow> connected(- S))" 
+  by (metis assms component_complement_connected double_complement separation_by_component_closed separation_by_component_open)
 
 
 text\<open>Another interesting equivalent of an inessential mapping into C-{0}\<close>
@@ -5421,7 +5367,7 @@
     case False
     have anr: "ANR (-{0::complex})"
       by (simp add: ANR_delete open_Compl open_imp_ANR)
-    obtain g where contg: "continuous_on UNIV g" and gim: "g ` UNIV \<subseteq> -{0}"
+    obtain g where contg: "continuous_on UNIV g" and gim: "g \<in> UNIV \<rightarrow> -{0}"
                    and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
     proof (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ homotopic_with_symD [OF a]])
       show "closedin (top_of_set UNIV) S"
@@ -5456,13 +5402,7 @@
         by (auto intro: continuous_on_subset [OF contj] simp flip: homeomorphism_image2 [OF hk])
     qed
     show "f x = exp ((j \<circ> k) x)" if "x \<in> S" for x
-    proof -
-      have "f x = (g \<circ> h) (k x)"
-        by (simp add: gf that)
-      also have "... = exp (j (k x))"
-        by (metis rangeI homeomorphism_image2 [OF hk] j)
-      finally show ?thesis by simp
-    qed
+      by (metis UNIV_I comp_apply gf hk homeomorphism_def image_eqI j that)
   qed
   then show ?lhs
     by (simp add: inessential_eq_continuous_logarithm)
@@ -5487,7 +5427,7 @@
     and ope: "openin (top_of_set (\<Union>\<F>)) C"
     and "homotopic_with_canon (\<lambda>x. True) C T f (\<lambda>x. a)"
     using assms by blast
-  with \<open>C \<noteq> {}\<close> have "f ` C \<subseteq> T" "a \<in> T"
+  with \<open>C \<noteq> {}\<close> have "f \<in> C \<rightarrow> T" "a \<in> T"
     using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+
   have "homotopic_with_canon (\<lambda>x. True) (\<Union>\<F>) T f (\<lambda>x. a)"
   proof (rule homotopic_on_clopen_Union)
@@ -5514,15 +5454,9 @@
 
 proposition Janiszewski_dual:
   fixes S :: "complex set"
-  assumes
-   "compact S" "compact T" "connected S" "connected T" "connected(- (S \<union> T))"
+  assumes "compact S" "compact T" "connected S" "connected T" "connected(- (S \<union> T))"
  shows "connected(S \<inter> T)"
-proof -
-  have ST: "compact (S \<union> T)"
-    by (simp add: assms compact_Un)
-  with Borsukian_imp_unicoherent [of "S \<union> T"] ST assms
-  show ?thesis
-    by (simp add: Borsukian_separation_compact closed_subset compact_imp_closed unicoherentD)
-qed
+  by (meson Borsukian_imp_unicoherent Borsukian_separation_compact assms closed_subset compact_Un 
+      compact_imp_closed sup_ge1 sup_ge2 unicoherentD)
 
 end
--- a/src/HOL/Analysis/Homotopy.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Homotopy.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -2950,11 +2950,11 @@
   obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B"
              and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
              and "independent B" "finite B" "card B = dim S" "span B = S"
-    by (metis orthonormal_basis_subspace [OF S] independent_finite)
+    by (metis orthonormal_basis_subspace [OF S] independent_imp_finite)
   obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C"
              and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1"
              and "independent C" "finite C" "card C = dim T" "span C = T"
-    by (metis orthonormal_basis_subspace [OF T] independent_finite)
+    by (metis orthonormal_basis_subspace [OF T] independent_imp_finite)
   obtain fb where "fb ` B \<subseteq> C" "inj_on fb B"
     by (metis \<open>card B = dim S\<close> \<open>card C = dim T\<close> \<open>finite B\<close> \<open>finite C\<close> card_le_inj d)
   then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B"
@@ -3003,11 +3003,11 @@
   obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B"
              and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
              and "independent B" "finite B" "card B = dim S" "span B = S"
-    by (metis orthonormal_basis_subspace [OF S] independent_finite)
+    by (metis orthonormal_basis_subspace [OF S] independent_imp_finite)
   obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C"
              and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1"
              and "independent C" "finite C" "card C = dim T" "span C = T"
-    by (metis orthonormal_basis_subspace [OF T] independent_finite)
+    by (metis orthonormal_basis_subspace [OF T] independent_imp_finite)
   obtain fb where "bij_betw fb B C"
     by (metis \<open>finite B\<close> \<open>finite C\<close> bij_betw_iff_card \<open>card B = dim S\<close> \<open>card C = dim T\<close> d)
   then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B"
--- a/src/HOL/Analysis/Path_Connected.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Analysis/Path_Connected.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -50,16 +50,6 @@
   unfolding path_def path_image_def
   using continuous_on_compose by blast
 
-lemma continuous_on_translation_eq:
-  fixes g :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_vector"
-  shows "continuous_on A ((+) a \<circ> g) = continuous_on A g"
-proof -
-  have g: "g = (\<lambda>x. -a + x) \<circ> ((\<lambda>x. a + x) \<circ> g)"
-    by (rule ext) simp
-  show ?thesis
-    by (metis (no_types, opaque_lifting) g continuous_on_compose homeomorphism_def homeomorphism_translation)
-qed
-
 lemma path_translation_eq:
   fixes g :: "real \<Rightarrow> 'a :: real_normed_vector"
   shows "path((\<lambda>x. a + x) \<circ> g) = path g"
@@ -251,23 +241,11 @@
   by auto
 
 lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g"
-proof -
-  have *: "\<And>g. path_image (reversepath g) \<subseteq> path_image g"
-    unfolding path_image_def subset_eq reversepath_def Ball_def image_iff
-    by force
-  show ?thesis
-    using *[of g] *[of "reversepath g"]
-    unfolding reversepath_reversepath
-    by auto
-qed
+  by (metis cancel_comm_monoid_add_class.diff_cancel diff_zero image_comp 
+      image_diff_atLeastAtMost path_image_def reversepath_o)
 
 lemma path_reversepath [simp]: "path (reversepath g) \<longleftrightarrow> path g"
-proof -
-  have *: "\<And>g. path g \<Longrightarrow> path (reversepath g)"
-    by (metis cancel_comm_monoid_add_class.diff_cancel continuous_on_compose 
-        continuous_on_op_minus diff_zero image_diff_atLeastAtMost path_def reversepath_o)
-  then show ?thesis by force
-qed
+  by (metis continuous_on_compose continuous_on_op_minus image_comp image_ident path_def path_image_def path_image_reversepath reversepath_o reversepath_reversepath)
 
 lemma arc_reversepath:
   assumes "arc g" shows "arc(reversepath g)"
@@ -375,8 +353,8 @@
   by auto
 
 lemma subset_path_image_join:
-  assumes "path_image g1 \<subseteq> s" and "path_image g2 \<subseteq> s"
-  shows "path_image (g1 +++ g2) \<subseteq> s"
+  assumes "path_image g1 \<subseteq> S" and "path_image g2 \<subseteq> S"
+  shows "path_image (g1 +++ g2) \<subseteq> S"
   using path_image_join_subset[of g1 g2] and assms
   by auto
 
@@ -495,16 +473,10 @@
           "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
     shows "arc(g1 +++ g2)"
 proof -
-  have injg1: "inj_on g1 {0..1}"
-    using assms
-    by (simp add: arc_def)
-  have injg2: "inj_on g2 {0..1}"
+  have injg1: "inj_on g1 {0..1}" and injg2: "inj_on g2 {0..1}" 
+     and g11: "g1 1 = g2 0" and sb: "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g2 0}"
     using assms
-    by (simp add: arc_def)
-  have g11: "g1 1 = g2 0"
-   and sb:  "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g2 0}"
-    using assms
-    by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def)
+    by (auto simp: arc_def pathfinish_def pathstart_def path_image_def)
   { fix x and y::real
     assume xy: "g2 (2 * x - 1) = g1 (2 * y)" "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1"
     then have "g1 (2 * y) = g2 0"
@@ -660,9 +632,9 @@
 
 lemma simple_path_join_loop_eq:
   assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2"
-    shows "simple_path(g1 +++ g2) \<longleftrightarrow>
+  shows "simple_path(g1 +++ g2) \<longleftrightarrow>
              arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
-by (metis assms simple_path_joinE simple_path_join_loop)
+  by (metis assms simple_path_joinE simple_path_join_loop)
 
 lemma arc_join_eq:
   assumes "pathfinish g1 = pathstart g2"
@@ -672,8 +644,8 @@
 proof
   assume ?lhs then show ?rhs 
     using reversepath_simps assms
-    by (smt (verit, ccfv_threshold) Int_commute arc_distinct_ends arc_imp_simple_path arc_reversepath 
-            in_mono insertE pathfinish_join reversepath_joinpaths simple_path_joinE subsetI)
+    by (smt (verit, best) Int_commute arc_reversepath arc_simple_path in_mono insertE pathstart_join 
+          reversepath_joinpaths simple_path_joinE subsetI)
 next
   assume ?rhs then show ?lhs
     using assms
@@ -681,18 +653,17 @@
 qed
 
 lemma arc_join_eq_alt:
-        "pathfinish g1 = pathstart g2
-        \<Longrightarrow> (arc(g1 +++ g2) \<longleftrightarrow>
-             arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 = {pathstart g2})"
-using pathfinish_in_path_image by (fastforce simp: arc_join_eq)
+  "pathfinish g1 = pathstart g2
+   \<Longrightarrow> arc(g1 +++ g2) \<longleftrightarrow> arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 = {pathstart g2}"
+  using pathfinish_in_path_image by (fastforce simp: arc_join_eq)
 
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>The joining of paths is associative\<close>
 
 lemma path_assoc:
-    "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart r\<rbrakk>
+  "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart r\<rbrakk>
      \<Longrightarrow> path(p +++ (q +++ r)) \<longleftrightarrow> path((p +++ q) +++ r)"
-by simp
+  by simp
 
 lemma simple_path_assoc:
   assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r"
@@ -740,18 +711,18 @@
 subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Symmetry and loops\<close>
 
 lemma path_sym:
-    "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> path(p +++ q) \<longleftrightarrow> path(q +++ p)"
+  "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> path(p +++ q) \<longleftrightarrow> path(q +++ p)"
   by auto
 
 lemma simple_path_sym:
-    "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk>
+  "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk>
      \<Longrightarrow> simple_path(p +++ q) \<longleftrightarrow> simple_path(q +++ p)"
-by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop)
+  by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop)
 
 lemma path_image_sym:
-    "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk>
+  "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk>
      \<Longrightarrow> path_image(p +++ q) = path_image(q +++ p)"
-by (simp add: path_image_join sup_commute)
+  by (simp add: path_image_join sup_commute)
 
 
 subsection\<open>Subpath\<close>
@@ -821,7 +792,7 @@
 
 lemma sum_le_prod1:
   fixes a::real shows "\<lbrakk>a \<le> 1; b \<le> 1\<rbrakk> \<Longrightarrow> a + b \<le> 1 + a * b"
-by (metis add.commute affine_ineq mult.right_neutral)
+  by (metis add.commute affine_ineq mult.right_neutral)
 
 lemma simple_path_subpath_eq:
   "simple_path(subpath u v g) \<longleftrightarrow>
@@ -871,9 +842,8 @@
   assumes "simple_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<noteq> v"
   shows "simple_path(subpath u v g)"
   using assms
-  apply (simp add: simple_path_subpath_eq simple_path_imp_path)
-  apply (simp add: simple_path_def loop_free_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce)
-  done
+  unfolding simple_path_subpath_eq
+  by (force simp:  simple_path_def loop_free_def closed_segment_real_eq image_affinity_atLeastAtMost)
 
 lemma arc_simple_path_subpath:
     "\<lbrakk>simple_path g; u \<in> {0..1}; v \<in> {0..1}; g u \<noteq> g v\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
@@ -1049,7 +1019,7 @@
   shows "pathfinish (shiftpath a g) = g a"
     and "pathstart (shiftpath a g) = g a"
   using assms
-  by (auto intro!: pathfinish_shiftpath pathstart_shiftpath)
+  by (simp_all add: pathstart_shiftpath pathfinish_shiftpath)
 
 lemma closed_shiftpath:
   assumes "pathfinish g = pathstart g"
@@ -1082,7 +1052,7 @@
     proof (rule continuous_on_eq)
       show "continuous_on {1-a..1} (g \<circ> (+) (a - 1))"
         by (intro continuous_intros continuous_on_subset [OF contg]) (use \<open>a \<in> {0..1}\<close> in auto)
-    qed (auto simp:  "**" add.commute add_diff_eq)
+    qed (auto simp: "**" add.commute add_diff_eq)
   qed auto
 qed
 
@@ -1699,19 +1669,8 @@
 
 lemma path_connected_path_component [simp]:
   "path_connected (path_component_set S x)"
-proof (clarsimp simp: path_connected_def)
-  fix y z
-  assume pa: "path_component S x y" "path_component S x z"
-  then have pae: "path_component_set S x = path_component_set S y"
-    using path_component_eq by auto
-  obtain g where g: "path g \<and> path_image g \<subseteq> S \<and> pathstart g = y \<and> pathfinish g = z"
-    using pa path_component_sym path_component_trans path_component_def by metis
-  then have "path_image g \<subseteq> path_component_set S x"
-    using pae path_component_maximal path_connected_path_image by blast
-  then show "\<exists>g. path g \<and> path_image g \<subseteq> path_component_set S x \<and>
-                 pathstart g = y \<and> pathfinish g = z"
-    using g by blast
-qed
+  by (smt (verit) mem_Collect_eq path_component_def path_component_eq path_component_maximal 
+      path_connected_component path_connected_path_image pathstart_in_path_image)
 
 lemma path_component: 
   "path_component S x y \<longleftrightarrow> (\<exists>t. path_connected t \<and> t \<subseteq> S \<and> x \<in> t \<and> y \<in> t)"
@@ -1726,13 +1685,7 @@
 
 lemma path_component_path_component [simp]:
    "path_component_set (path_component_set S x) x = path_component_set S x"
-proof (cases "x \<in> S")
-  case True show ?thesis
-    by (metis True mem_Collect_eq path_component_refl path_connected_component_set path_connected_path_component)
-next
-  case False then show ?thesis
-    by (metis False empty_iff path_component_eq_empty)
-qed
+  by (metis (full_types) mem_Collect_eq path_component_eq_empty path_component_refl path_connected_component_set path_connected_path_component)
 
 lemma path_component_subset_connected_component:
    "(path_component_set S x) \<subseteq> (connected_component_set S x)"
@@ -1750,8 +1703,8 @@
 lemma path_connected_linear_image:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "path_connected S" "bounded_linear f"
-    shows "path_connected(f ` S)"
-by (auto simp: linear_continuous_on assms path_connected_continuous_image)
+  shows "path_connected(f ` S)"
+  by (auto simp: linear_continuous_on assms path_connected_continuous_image)
 
 lemma is_interval_path_connected: "is_interval S \<Longrightarrow> path_connected S"
   by (simp add: convex_imp_path_connected is_interval_convex)
@@ -1814,8 +1767,8 @@
 lemma linear_homeomorphic_image:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "linear f" "inj f"
-    shows "S homeomorphic f ` S"
-by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms])
+  shows "S homeomorphic f ` S"
+  by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms])
 
 lemma path_connected_Times:
   assumes "path_connected s" "path_connected t"
@@ -1851,7 +1804,7 @@
 lemma is_interval_path_connected_1:
   fixes s :: "real set"
   shows "is_interval s \<longleftrightarrow> path_connected s"
-using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast
+  using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast
 
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Path components\<close>
@@ -1885,20 +1838,15 @@
 qed
 
 lemma path_component_unique:
-  assumes "x \<in> c" "c \<subseteq> S" "path_connected c"
-          "\<And>c'. \<lbrakk>x \<in> c'; c' \<subseteq> S; path_connected c'\<rbrakk> \<Longrightarrow> c' \<subseteq> c"
-   shows "path_component_set S x = c"
-    (is "?lhs = ?rhs")
-proof 
-  show "?lhs \<subseteq> ?rhs"
-    using assms
-    by (metis mem_Collect_eq path_component_refl path_component_subset path_connected_path_component subsetD)
-qed (simp add: assms path_component_maximal)
+  assumes "x \<in> C" "C \<subseteq> S" "path_connected C"
+          "\<And>C'. \<lbrakk>x \<in> C'; C' \<subseteq> S; path_connected C'\<rbrakk> \<Longrightarrow> C' \<subseteq> C"
+        shows "path_component_set S x = C"
+  by (smt (verit, best) Collect_cong assms path_component path_component_of_subset path_connected_component_set)
 
 lemma path_component_intermediate_subset:
-   "path_component_set u a \<subseteq> t \<and> t \<subseteq> u
-        \<Longrightarrow> path_component_set t a = path_component_set u a"
-by (metis (no_types) path_component_mono path_component_path_component subset_antisym)
+  "path_component_set U a \<subseteq> T \<and> T \<subseteq> U
+        \<Longrightarrow> path_component_set T a = path_component_set U a"
+  by (metis (no_types) path_component_mono path_component_path_component subset_antisym)
 
 lemma complement_path_component_Union:
   fixes x :: "'a :: topological_space"
@@ -2078,9 +2026,8 @@
 proof
   assume ?lhs
   then show ?rhs
-    apply (simp add: fun_eq_iff path_component_in_topspace)
-    apply (meson path_component_of_sym path_component_of_trans)
-    done
+    unfolding fun_eq_iff path_component_in_topspace
+    by (metis path_component_in_topspace path_component_of_sym path_component_of_trans)
 qed (simp add: path_component_of_refl)
 
 lemma path_component_of_disjoint:
@@ -2105,7 +2052,7 @@
   have "topspace (subtopology X (path_component_of_set X x)) = path_component_of_set X x"
     by (meson path_component_of_subset_topspace topspace_subtopology_subset)
   then have "path_connected_space (subtopology X (path_component_of_set X x))"
-    by (metis (full_types) path_component_of_aux mem_Collect_eq path_component_of_equiv path_connected_space_iff_path_component)
+    by (metis mem_Collect_eq path_component_of_aux path_component_of_equiv path_connected_space_iff_path_component)
   then show ?thesis
     by (simp add: path_component_of_subset_topspace path_connectedin_def)
 qed
@@ -2152,13 +2099,14 @@
   by (metis imageI nonempty_path_components_of path_component_of_eq_empty path_components_of_def)
 
 lemma path_connectedin_Union:
-  assumes \<A>: "\<And>S. S \<in> \<A> \<Longrightarrow> path_connectedin X S" "\<Inter>\<A> \<noteq> {}"
+  assumes \<A>: "\<And>S. S \<in> \<A> \<Longrightarrow> path_connectedin X S" and "\<Inter>\<A> \<noteq> {}"
   shows "path_connectedin X (\<Union>\<A>)"
 proof -
   obtain a where "\<And>S. S \<in> \<A> \<Longrightarrow> a \<in> S"
     using assms by blast
   then have "\<And>x. x \<in> topspace (subtopology X (\<Union>\<A>)) \<Longrightarrow> path_component_of (subtopology X (\<Union>\<A>)) a x"
-    by simp (meson Union_upper \<A> path_component_of path_connectedin_subtopology)
+    unfolding topspace_subtopology path_component_of
+    by (metis (full_types) IntD2 Union_iff Union_upper \<A> path_connectedin_subtopology)
   then show ?thesis
     using \<A> unfolding path_connectedin_def
     by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component)
@@ -2462,8 +2410,8 @@
     {PiE I B |B. \<forall>i \<in> I. B i \<in> path_components_of(X i)}"  (is "?lhs=?rhs")
 proof
   show "?lhs \<subseteq> ?rhs"
-    apply (simp add: path_components_of_def image_subset_iff)
-    by (smt (verit, best) PiE_iff image_eqI path_component_of_product_topology)
+    unfolding path_components_of_def image_subset_iff
+    by (smt (verit) image_iff mem_Collect_eq path_component_of_product_topology topspace_product_topology_alt)
 next
   show "?rhs \<subseteq> ?lhs"
   proof
@@ -2542,14 +2490,6 @@
   assumes "2 \<le> DIM('a)"
   shows "path_connected(sphere a r)"
 proof (cases r "0::real" rule: linorder_cases)
-  case less
-  then show ?thesis
-    by simp
-next
-  case equal
-  then show ?thesis
-    by simp
-next
   case greater
   then have eq: "(sphere (0::'a) r) = (\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a})"
     by (force simp: image_iff split: if_split_asm)
@@ -2557,13 +2497,11 @@
     by (intro continuous_intros) auto
   then have "path_connected ((\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))"
     by (intro path_connected_continuous_image path_connected_punctured_universe assms)
-  with eq have "path_connected (sphere (0::'a) r)"
-    by auto
-  then have "path_connected((+) a ` (sphere (0::'a) r))"
+  with eq have "path_connected((+) a ` (sphere (0::'a) r))"
     by (simp add: path_connected_translation)
   then show ?thesis
     by (metis add.right_neutral sphere_translation)
-qed
+qed auto
 
 lemma connected_sphere:
     fixes a :: "'a :: euclidean_space"
@@ -3000,14 +2938,13 @@
 proof -
   obtain i::'a where i: "i \<in> Basis"
     using nonempty_Basis by blast
-  obtain B where B: "B>0" "-S \<subseteq> ball 0 B"
+  obtain B where "B>0" and B: "-S \<subseteq> ball 0 B"
     using bounded_subset_ballD [OF bs, of 0] by auto
   then have *: "\<And>x. B \<le> norm x \<Longrightarrow> x \<in> S"
     by (force simp: ball_def dist_norm)
-  obtain x' where x': "connected_component S x x'" "norm x' > B"
-    using B(1) bo(1) bounded_pos by force
-  obtain y' where y': "connected_component S y y'" "norm y' > B"
-    using B(1) bo(2) bounded_pos by force
+  obtain x' y' where x': "connected_component S x x'" "norm x' > B"
+  and  y': "connected_component S y y'" "norm y' > B"
+    using \<open>B>0\<close> bo bounded_pos by (metis linorder_not_le mem_Collect_eq) 
   have x'y': "connected_component S x' y'"
     unfolding connected_component_def
   proof (intro exI conjI)
@@ -3197,8 +3134,8 @@
   fixes S :: "'a::euclidean_space set"
   assumes S: "bounded S"  "2 \<le> DIM('a)"
   shows "- (outside S) = {x. \<forall>B. \<exists>y. B \<le> norm(y) \<and> \<not> connected_component (- S) x y}"
-  apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms])
-  by (meson gt_ex less_le_trans)
+  unfolding not_outside_connected_component_lt [OF assms]
+  by (metis (no_types, opaque_lifting) dual_order.strict_trans1 gt_ex pinf(8))
 
 lemma inside_connected_component_lt:
     fixes S :: "'a::euclidean_space set"
@@ -3327,10 +3264,10 @@
 qed
 
 lemma connected_component_UNIV [simp]:
-    fixes x :: "'a::real_normed_vector"
-    shows "connected_component_set UNIV x = UNIV"
-using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV
-by auto
+  fixes x :: "'a::real_normed_vector"
+  shows "connected_component_set UNIV x = UNIV"
+  using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV
+  by auto
 
 lemma connected_component_eq_UNIV:
     fixes x :: "'a::real_normed_vector"
@@ -3484,10 +3421,8 @@
 lemma inside_frontier_eq_interior:
      fixes S :: "'a :: {real_normed_vector, perfect_space} set"
      shows "\<lbrakk>bounded S; convex S\<rbrakk> \<Longrightarrow> inside(frontier S) = interior S"
-  apply (simp add: inside_outside outside_frontier_eq_complement_closure)
-  using closure_subset interior_subset
-  apply (auto simp: frontier_def)
-  done
+  unfolding inside_outside outside_frontier_eq_complement_closure
+  using closure_subset interior_subset by (auto simp: frontier_def)
 
 lemma open_inside:
     fixes S :: "'a::real_normed_vector set"
@@ -3527,10 +3462,10 @@
 qed
 
 lemma closure_inside_subset:
-    fixes S :: "'a::real_normed_vector set"
-    assumes "closed S"
-      shows "closure(inside S) \<subseteq> S \<union> inside S"
-by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside)
+  fixes S :: "'a::real_normed_vector set"
+  assumes "closed S"
+  shows "closure(inside S) \<subseteq> S \<union> inside S"
+  by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside)
 
 lemma frontier_inside_subset:
     fixes S :: "'a::real_normed_vector set"
@@ -3594,7 +3529,7 @@
 lemma inside_outside_intersect_connected:
       "\<lbrakk>connected T; inside S \<inter> T \<noteq> {}; outside S \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow> S \<inter> T \<noteq> {}"
   apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify)
-  by (metis (no_types, opaque_lifting) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl)
+  by (metis compl_le_swap1 connected_componentI connected_component_eq mem_Collect_eq)
 
 lemma outside_bounded_nonempty:
   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
@@ -4045,7 +3980,8 @@
     have "embedding_map (top_of_set S) euclideanreal f"
       using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto
     then show ?thesis
-      by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that)
+      unfolding embedding_map_def topspace_euclidean_subtopology
+      by (metis f homeomorphic_map_closedness_eq homeomorphism_injective_closed_map that)
   qed
 qed (metis homeomorphism_def inj_onI)
 
--- a/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -316,15 +316,13 @@
                  ((\<lambda>u. f u / (u - w)) has_contour_integral (\<lambda>x. 2 * of_real pi * \<i> * f x) w) (circlepath z r)"
     by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute)
   show ?thes1
-    apply (simp add: power2_eq_square)
-    apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified])
-    apply (blast intro: int)
-    done
+    unfolding power2_eq_square
+    using int Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1]
+    by fastforce
   have "((\<lambda>x. 2 * of_real pi * \<i> * f x) has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2)) (at w)"
-    apply (simp add: power2_eq_square)
-    apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\<lambda>x. 2 * of_real pi * \<i> * f x", simplified])
-    apply (blast intro: int)
-    done
+    unfolding power2_eq_square
+    using int Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\<lambda>x. 2 * of_real pi * \<i> * f x"]
+    by fastforce
   then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2) / (2 * of_real pi * \<i>)) (at w)"
     by (rule DERIV_cdivide [where f = "\<lambda>x. 2 * of_real pi * \<i> * f x" and c = "2 * of_real pi * \<i>", simplified])
   show ?thes2
@@ -396,8 +394,7 @@
 lemma has_field_derivative_higher_deriv:
      "\<lbrakk>f holomorphic_on S; open S; x \<in> S\<rbrakk>
       \<Longrightarrow> ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)"
-by (metis (no_types, opaque_lifting) DERIV_deriv_iff_field_differentiable at_within_open comp_apply
-         funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def)
+  using holomorphic_derivI holomorphic_higher_deriv by fastforce
   
 lemma higher_deriv_cmult:
   assumes "f holomorphic_on A" "x \<in> A" "open A"
@@ -644,7 +641,7 @@
     by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T)
   have holo3: "(\<lambda>z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S"
     by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros)
-  have "(*) u holomorphic_on S" "f holomorphic_on (*) u ` S"
+  have u: "(*) u holomorphic_on S" "f holomorphic_on (*) u ` S"
     by (rule holo0 holomorphic_intros)+
   then have holo1: "(\<lambda>w. f (u * w)) holomorphic_on S"
     by (rule holomorphic_on_compose [where g=f, unfolded o_def])
@@ -658,12 +655,7 @@
     have "(deriv ^^ n) f analytic_on T"
       by (simp add: analytic_on_open f holomorphic_higher_deriv T)
     then have "(\<lambda>w. (deriv ^^ n) f (u * w)) analytic_on S"
-    proof -
-      have "(deriv ^^ n) f \<circ> (*) u holomorphic_on S"
-        by (simp add: holo2 holomorphic_on_compose)
-      then show ?thesis
-        by (simp add: S analytic_on_open o_def)
-    qed
+      by (meson S u analytic_on_open holo2 holomorphic_on_compose holomorphic_transform o_def)
     then show ?thesis
       by (intro deriv_cmult analytic_on_imp_differentiable_at [OF _ Suc.prems])
   qed
@@ -1298,15 +1290,15 @@
 lemma series_and_derivative_comparison_complex:
   fixes S :: "complex set"
   assumes S: "open S"
-      and hfd: "\<And>n x. x \<in> S \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
-      and to_g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>d h. 0 < d \<and> summable h \<and> range h \<subseteq> \<real>\<^sub>\<ge>\<^sub>0 \<and> (\<forall>\<^sub>F n in sequentially. \<forall>y\<in>ball x d \<inter> S. cmod(f n y) \<le> cmod (h n))"
+    and hfd: "\<And>n x. x \<in> S \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
+    and to_g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>d h. 0 < d \<and> summable h \<and> range h \<subseteq> \<real>\<^sub>\<ge>\<^sub>0 \<and> (\<forall>\<^sub>F n in sequentially. \<forall>y\<in>ball x d \<inter> S. cmod(f n y) \<le> cmod (h n))"
   shows "\<exists>g g'. \<forall>x \<in> S. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
-apply (rule series_and_derivative_comparison_local [OF S hfd], assumption)
-apply (rule ex_forward [OF to_g], assumption)
-apply (erule exE)
-apply (rule_tac x="Re \<circ> h" in exI)
-apply (force simp: summable_Re o_def nonneg_Reals_cmod_eq_Re image_subset_iff)
-done
+  apply (rule series_and_derivative_comparison_local [OF S hfd], assumption)
+  apply (rule ex_forward [OF to_g], assumption)
+  apply (erule exE)
+  apply (rule_tac x="Re \<circ> h" in exI)
+  apply (force simp: summable_Re o_def nonneg_Reals_cmod_eq_Re image_subset_iff)
+  done
 
 text\<open>Sometimes convenient to compare with a complex series of positive reals. (?)\<close>
 lemma series_differentiable_comparison_complex:
@@ -1410,9 +1402,8 @@
 corollary holomorphic_iff_power_series:
      "f holomorphic_on ball z r \<longleftrightarrow>
       (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
-  apply (intro iffI ballI holomorphic_power_series, assumption+)
-  apply (force intro: power_series_holomorphic [where a = "\<lambda>n. (deriv ^^ n) f z / (fact n)"])
-  done
+  using power_series_holomorphic [where a = "\<lambda>n. (deriv ^^ n) f z / (fact n)"] holomorphic_power_series
+  by blast
 
 lemma power_series_analytic:
      "(\<And>w. w \<in> ball z r \<Longrightarrow> (\<lambda>n. a n*(w - z)^n) sums f w) \<Longrightarrow> f analytic_on ball z r"
@@ -1670,12 +1661,8 @@
     obtain \<delta> where "\<delta>>0" and \<delta>: "cball w \<delta> \<subseteq> U" using open_contains_cball \<open>open U\<close> \<open>w \<in> U\<close> by force
     let ?TZ = "cball w \<delta>  \<times> closed_segment a b"
     have "uniformly_continuous_on ?TZ (\<lambda>(x,y). F x y)"
-    proof (rule compact_uniformly_continuous)
-      show "continuous_on ?TZ (\<lambda>(x,y). F x y)"
-        by (rule continuous_on_subset[OF cond_uu]) (use SigmaE \<delta> abu in blast)
-      show "compact ?TZ"
-        by (simp add: compact_Times)
-    qed
+      by (metis Sigma_mono \<delta> abu compact_Times compact_cball compact_segment compact_uniformly_continuous 
+          cond_uu continuous_on_subset)
     then obtain \<eta> where "\<eta>>0"
         and \<eta>: "\<And>x x'. \<lbrakk>x\<in>?TZ; x'\<in>?TZ; dist x' x < \<eta>\<rbrakk> \<Longrightarrow>
                          dist ((\<lambda>(x,y). F x y) x') ((\<lambda>(x,y). F x y) x) < \<epsilon>/norm(b - a)"
@@ -1992,16 +1979,10 @@
           by (auto intro: continuous_on_swap_args cond_uu)
       qed
       have cont_cint_d\<gamma>: "continuous_on {0..1} ((\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) \<circ> \<gamma>)"
-      proof (rule continuous_on_compose)
-        show "continuous_on {0..1} \<gamma>"
-          using \<open>path \<gamma>\<close> path_def by blast
-        show "continuous_on (\<gamma> ` {0..1}) (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
-          using pasz unfolding path_image_def
-          by (auto intro!: continuous_on_subset [OF cont_cint_d])
-      qed
+        by (metis Diff_subset \<open>path \<gamma>\<close> cont_cint_d continuous_on_compose continuous_on_subset pasz path_def path_image_def)
       have "continuous_on {0..1} (\<lambda>x. vector_derivative \<gamma> (at x))"
         using pf\<gamma>' by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \<gamma>'])
-      then      have cint_cint: "(\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) contour_integrable_on \<gamma>"
+      then have cint_cint: "(\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) contour_integrable_on \<gamma>"
         apply (simp add: contour_integrable_on)
         apply (rule integrable_continuous_real)
         by (rule continuous_on_mult [OF cont_cint_d\<gamma> [unfolded o_def]])
@@ -2632,8 +2613,8 @@
   have g_nz: "g \<noteq> 0"
   proof -
     define z :: complex where "z = (if r = \<infinity> then 1 else of_real (real_of_ereal r / 2))"
-    from \<open>r > 0\<close> have "z \<in> eball 0 r"
-      by (cases r) (auto simp: z_def eball_def)
+    have "z \<in> eball 0 r"
+      using \<open>r > 0\<close> ereal_less_real_iff z_def by fastforce
     moreover have "z \<noteq> 0" using \<open>r > 0\<close> 
       by (cases r) (auto simp: z_def)
     ultimately have "eval_fps g z \<noteq> 0" by (rule assms(6))
--- a/src/HOL/Complex_Analysis/Cauchy_Integral_Theorem.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Complex_Analysis/Cauchy_Integral_Theorem.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -842,7 +842,7 @@
     by (simp add: has_field_derivative_def has_derivative_at2 bounded_linear_mult_right)
 qed
 
-(** Existence of a primitive.*)
+text \<open>Existence of a primitive\<close>
 lemma holomorphic_starlike_primitive:
   fixes f :: "complex \<Rightarrow> complex"
   assumes contf: "continuous_on S f"
@@ -1017,8 +1017,7 @@
   assumes gpd: "g piecewise_differentiable_on {a..b}"
       and dh: "\<And>x. x \<in> S \<Longrightarrow> (f has_field_derivative f' x) (at x within S)"
       and gs: "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> S"
-  shows 
-    "(\<lambda>x. f' (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b}"
+    shows "(\<lambda>x. f' (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b}"
 proof (cases "cbox a b = {}")
   case False
   then show ?thesis
--- a/src/HOL/Complex_Analysis/Complex_Singularities.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Complex_Analysis/Complex_Singularities.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -74,10 +74,14 @@
   assumes "open A" "x \<in> A" "f holomorphic_on A"
   shows   "\<not>is_pole f x"
 proof -
-  have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact
-  with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at)
-  hence "f \<midarrow>x\<rightarrow> f x" by (simp add: isCont_def)
-  thus "\<not>is_pole f x" unfolding is_pole_def
+  have "continuous_on A f" 
+    by (intro holomorphic_on_imp_continuous_on) fact
+  with assms have "isCont f x" 
+    by (simp add: continuous_on_eq_continuous_at)
+  hence "f \<midarrow>x\<rightarrow> f x" 
+    by (simp add: isCont_def)
+  thus "\<not>is_pole f x" 
+    unfolding is_pole_def
     using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto
 qed
 
@@ -87,17 +91,19 @@
      (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros)
 
 lemma is_pole_cmult_iff [simp]:
-  "c \<noteq> 0 \<Longrightarrow> is_pole (\<lambda>z. c * f z :: 'a :: real_normed_field) z \<longleftrightarrow> is_pole f z"
+  assumes "c \<noteq> 0"
+  shows "is_pole (\<lambda>z. c * f z :: 'a :: real_normed_field) z \<longleftrightarrow> is_pole f z"
 proof
-  assume *: "c \<noteq> 0" "is_pole (\<lambda>z. c * f z) z"
-  have "is_pole (\<lambda>z. inverse c * (c * f z)) z" unfolding is_pole_def
-    by (rule tendsto_mult_filterlim_at_infinity tendsto_const)+ (use * in \<open>auto simp: is_pole_def\<close>)
+  assume "is_pole (\<lambda>z. c * f z) z"
+  with \<open>c\<noteq>0\<close> have "is_pole (\<lambda>z. inverse c * (c * f z)) z" 
+    unfolding is_pole_def
+    by (force intro: tendsto_mult_filterlim_at_infinity)
   thus "is_pole f z"
-    using *(1) by (simp add: field_simps)
+    using \<open>c\<noteq>0\<close> by (simp add: field_simps)
 next
-  assume *: "c \<noteq> 0" "is_pole f z"
-  show "is_pole (\<lambda>z. c * f z) z" unfolding is_pole_def 
-    by (rule tendsto_mult_filterlim_at_infinity tendsto_const)+ (use * in \<open>auto simp: is_pole_def\<close>)
+  assume "is_pole f z"
+  with \<open>c\<noteq>0\<close> show "is_pole (\<lambda>z. c * f z) z"  
+    by (auto intro!: tendsto_mult_filterlim_at_infinity simp: is_pole_def)
 qed
 
 lemma is_pole_uminus_iff [simp]: "is_pole (\<lambda>z. -f z :: 'a :: real_normed_field) z \<longleftrightarrow> is_pole f z"
@@ -184,7 +190,13 @@
 lemma is_pole_mult_analytic_nonzero2:
   assumes "is_pole f x" "g analytic_on {x}" "g x \<noteq> 0"
   shows   "is_pole (\<lambda>x. f x * g x) x"
-  by (subst mult.commute, rule is_pole_mult_analytic_nonzero1) (use assms in auto)
+proof -
+  have g: "g analytic_on {x}"
+    using assms by auto
+  show ?thesis
+    using is_pole_mult_analytic_nonzero1 [OF \<open>is_pole f x\<close> g] \<open>g x \<noteq> 0\<close>
+    by (simp add: mult.commute)
+qed
 
 lemma is_pole_mult_analytic_nonzero1_iff:
   assumes "f analytic_on {x}" "f x \<noteq> 0"
@@ -433,7 +445,8 @@
         \<and> f w = g2 w * (w - z) powi n2  \<and> g2 w\<noteq>0"
       using \<open>fac n1 g1 r1\<close> \<open>fac n2 g2 r2\<close>   unfolding fac_def r_def
       by fastforce
-    ultimately show "n1=n2" using g1_holo g2_holo \<open>g1 z\<noteq>0\<close> \<open>g2 z\<noteq>0\<close>
+    ultimately show "n1=n2" 
+      using g1_holo g2_holo \<open>g1 z\<noteq>0\<close> \<open>g2 z\<noteq>0\<close>
       apply (elim holomorphic_factor_unique)
       by (auto simp add:r_def)
   qed
@@ -561,24 +574,7 @@
   assumes "isolated_singularity_at g z"
   assumes "\<forall>\<^sub>F w in (at z). g w = f w"
   shows "isolated_singularity_at f z"
-proof -
-  obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}"
-    using assms(1) unfolding isolated_singularity_at_def by auto
-  obtain r2 where "r2>0" and r2:" \<forall>x. x \<noteq> z \<and> dist x z < r2 \<longrightarrow> g x = f x"
-    using assms(2) unfolding eventually_at by auto
-  define r3 where "r3=min r1 r2"
-  have "r3>0" unfolding r3_def using \<open>r1>0\<close> \<open>r2>0\<close> by auto
-  moreover have "f analytic_on ball z r3 - {z}"
-  proof -
-    have "g holomorphic_on ball z r3 - {z}"
-      using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto)
-    then have "f holomorphic_on ball z r3 - {z}"
-      using r2 unfolding r3_def
-      by (auto simp add:dist_commute elim!:holomorphic_transform)
-    then show ?thesis by (subst analytic_on_open,auto)
-  qed
-  ultimately show ?thesis unfolding isolated_singularity_at_def by auto
-qed
+  using assms isolated_singularity_at_cong by blast
 
 lemma not_essential_powr[singularity_intros]:
   assumes "LIM w (at z). f w :> (at x)"
@@ -704,8 +700,7 @@
   have ?thesis when "\<not> ((\<exists>\<^sub>Fw in (at z). f w\<noteq>0) \<and> (\<exists>\<^sub>Fw in (at z). g w\<noteq>0))"
   proof -
     have "\<forall>\<^sub>Fw in (at z). fg w=0"
-      using that[unfolded frequently_def, simplified] unfolding fg_def
-      by (auto elim: eventually_rev_mp)
+      using fg_def frequently_elim1 not_eventually that by fastforce
     from tendsto_cong[OF this] have "fg \<midarrow>z\<rightarrow>0" by auto
     then show ?thesis unfolding not_essential_def fg_def by auto
   qed
@@ -787,9 +782,8 @@
   proof -
     have "\<forall>\<^sub>Fw in (at z). f w=0"
       using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp)
-    then have "\<forall>\<^sub>Fw in (at z). vf w=0"
-      unfolding vf_def by auto
-    from tendsto_cong[OF this] have "vf \<midarrow>z\<rightarrow>0" unfolding vf_def by auto
+    then have "vf \<midarrow>z\<rightarrow>0" 
+      unfolding vf_def by (simp add: tendsto_eventually)
     then show ?thesis unfolding not_essential_def vf_def by auto
   qed
   moreover have ?thesis when "is_pole f z"
@@ -802,22 +796,18 @@
   proof -
     from that obtain fz where fz:"f\<midarrow>z\<rightarrow>fz" by auto
     have ?thesis when "fz=0"
+
     proof -
       have "(\<lambda>w. inverse (vf w)) \<midarrow>z\<rightarrow>0"
         using fz that unfolding vf_def by auto
       moreover have "\<forall>\<^sub>F w in at z. inverse (vf w) \<noteq> 0"
         using non_zero_neighbour[OF f_iso f_ness f_nconst]
         unfolding vf_def by auto
-      ultimately have "is_pole vf z"
-        using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto
-      then show ?thesis unfolding not_essential_def vf_def by auto
+      ultimately show ?thesis unfolding not_essential_def vf_def
+         using filterlim_atI filterlim_inverse_at_iff is_pole_def by blast
     qed
     moreover have ?thesis when "fz\<noteq>0"
-    proof -
-      have "vf \<midarrow>z\<rightarrow>inverse fz"
-        using fz that unfolding vf_def by (auto intro:tendsto_eq_intros)
-      then show ?thesis unfolding not_essential_def vf_def by auto
-    qed
+      using fz not_essential_def tendsto_inverse that by blast
     ultimately show ?thesis by auto
   qed
   ultimately show ?thesis using f_ness unfolding not_essential_def by auto
@@ -856,7 +846,7 @@
     moreover
     have "f analytic_on ball z d3 - {z}"
       by (smt (verit, best) Diff_iff analytic_on_analytic_at d2 d3_def mem_ball)
-      then have "vf analytic_on ball z d3 - {z}"
+    then have "vf analytic_on ball z d3 - {z}"
       unfolding vf_def
       by (intro analytic_on_inverse; simp add: d1(2) d3_def dist_commute)
     ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto
@@ -878,8 +868,8 @@
   assumes f_iso:"isolated_singularity_at f z"
       and g_iso:"isolated_singularity_at g z"
     shows isolated_singularity_at_times[singularity_intros]:
-              "isolated_singularity_at (\<lambda>w. f w * g w) z" and
-          isolated_singularity_at_add[singularity_intros]:
+              "isolated_singularity_at (\<lambda>w. f w * g w) z"
+      and isolated_singularity_at_add[singularity_intros]:
               "isolated_singularity_at (\<lambda>w. f w + g w) z"
 proof -
   obtain d1 d2 where "d1>0" "d2>0"
@@ -912,19 +902,18 @@
   unfolding isolated_singularity_at_def by (simp add: gt_ex)
 
 lemma isolated_singularity_at_minus[singularity_intros]:
-  assumes f_iso:"isolated_singularity_at f z"
-      and g_iso:"isolated_singularity_at g z"
-    shows "isolated_singularity_at (\<lambda>w. f w - g w) z"
-  using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\<lambda>w. - g w"]
-        ,OF g_iso] by simp
+  assumes "isolated_singularity_at f z" and "isolated_singularity_at g z"
+  shows "isolated_singularity_at (\<lambda>w. f w - g w) z"
+  unfolding diff_conv_add_uminus
+  using assms isolated_singularity_at_add isolated_singularity_at_uminus by blast
 
 lemma isolated_singularity_at_divide[singularity_intros]:
-  assumes f_iso:"isolated_singularity_at f z"
-      and g_iso:"isolated_singularity_at g z"
-      and g_ness:"not_essential g z"
+  assumes "isolated_singularity_at f z"
+      and "isolated_singularity_at g z"
+      and "not_essential g z"
     shows "isolated_singularity_at (\<lambda>w. f w / g w) z"
-  using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso,
-          of "\<lambda>w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps)
+  unfolding divide_inverse
+  by (simp add: assms isolated_singularity_at_inverse isolated_singularity_at_times)
 
 lemma isolated_singularity_at_const[singularity_intros]:
     "isolated_singularity_at (\<lambda>w. c) z"
@@ -1013,14 +1002,7 @@
 lemma not_essential_holomorphic:
   assumes "f holomorphic_on A" "x \<in> A" "open A"
   shows   "not_essential f x"
-proof -
-  have "continuous_on A f"
-    using assms holomorphic_on_imp_continuous_on by blast
-  hence "f \<midarrow>x\<rightarrow> f x"
-    using assms continuous_on_eq_continuous_at isContD by blast
-  thus ?thesis
-    by (auto simp: not_essential_def)
-qed
+  by (metis assms at_within_open continuous_on holomorphic_on_imp_continuous_on not_essential_def)
 
 lemma not_essential_analytic:
   assumes "f analytic_on {z}"
@@ -1045,11 +1027,7 @@
   then have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
     by (intro eventually_at_in_open) auto
   thus "eventually (\<lambda>w. \<not>is_pole f w) (at z)"
-  proof eventually_elim
-    case (elim w)
-    with r show ?case
-      using analytic_imp_holomorphic not_is_pole_holomorphic open_delete by blast
-  qed
+    by (metis (no_types, lifting) analytic_at analytic_on_analytic_at eventually_mono not_is_pole_holomorphic r)
 qed
 
 lemma not_islimpt_poles:
@@ -1077,13 +1055,7 @@
 lemma isolated_singularity_at_analytic:
   assumes "f analytic_on {z}"
   shows   "isolated_singularity_at f z"
-proof -
-  from assms obtain r where r: "r > 0" "f holomorphic_on ball z r"
-    by (auto simp: analytic_on_def)
-  show ?thesis
-    by (rule isolated_singularity_at_holomorphic[of f "ball z r"])
-       (use \<open>r > 0\<close> in \<open>auto intro!: holomorphic_on_subset[OF r(2)]\<close>)
-qed
+  by (meson Diff_subset analytic_at assms holomorphic_on_subset isolated_singularity_at_holomorphic)
 
 subsection \<open>The order of non-essential singularities (i.e. removable singularities or poles)\<close>
 
@@ -1100,7 +1072,7 @@
 
 lemma zorder_exist:
   fixes f::"complex \<Rightarrow> complex" and z::complex
-  defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
+  defines "n \<equiv> zorder f z" and "g \<equiv> zor_poly f z"
   assumes f_iso:"isolated_singularity_at f z"
       and f_ness:"not_essential f z"
       and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
@@ -1109,7 +1081,7 @@
 proof -
   define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
           \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi n \<and> g w\<noteq>0))"
-  have "\<exists>!n. \<exists>g r. P n g r"
+  have "\<exists>!k. \<exists>g r. P k g r"
     using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto
   then have "\<exists>g r. P n g r"
     unfolding n_def P_def zorder_def
@@ -1168,8 +1140,8 @@
         and fr_nz: "inverse (fp w) \<noteq> 0"
     when "w\<in>ball z fr - {z}" for w
   proof -
-    have "f w = fp w * (w - z) powi fn" "fp w\<noteq>0"
-      using fr(2)[rule_format,of w] that by auto
+    have "f w = fp w * (w - z) powi fn" "fp w \<noteq> 0"
+      using fr(2) that by auto
     then show "vf w = (inverse (fp w)) * (w - z) powi (-fn)" "inverse (fp w)\<noteq>0"
       by (simp_all add: power_int_minus vf_def)
   qed
@@ -1256,13 +1228,7 @@
     define n where "n \<equiv> zorder f z"
 
     have "f w = zor_poly f z w * (w - z) powi n"
-    proof -
-      have "w\<in>cball z r1 - {z}"
-        using r_def that by auto
-      from rball1[rule_format, OF this]
-      show ?thesis unfolding n_def by auto
-    qed
-
+      using n_def r_def rball1 that by auto
     moreover have "f w = zor_poly ff 0 (w - z) * (w - z) powi n"
     proof -
       have "w-z\<in>cball 0 r2 - {0}"
@@ -1286,23 +1252,16 @@
   then have "\<forall>\<^sub>F w in at z. zor_poly f z w
                   = zor_poly ff 0 (w - z)"
     unfolding eventually_at
-    apply (rule_tac x=r in exI)
-    using \<open>r>0\<close> by (auto simp:dist_commute)
+    by (metis DiffI \<open>0 < r\<close> dist_commute mem_ball singletonD)
   moreover have "isCont (zor_poly f z) z"
     using holo1[THEN holomorphic_on_imp_continuous_on]
-    apply (elim continuous_on_interior)
-    using \<open>r1>0\<close> by auto
-  moreover have "isCont (\<lambda>w. zor_poly ff 0 (w - z)) z"
-  proof -
-    have "isCont (zor_poly ff 0) 0"
-      using holo2[THEN holomorphic_on_imp_continuous_on]
-      apply (elim continuous_on_interior)
-      using \<open>r2>0\<close> by auto
-    then show ?thesis
+    by (simp add: \<open>0 < r1\<close> continuous_on_interior)
+  moreover 
+  have "isCont (zor_poly ff 0) 0"
+    using \<open>0 < r2\<close> centre_in_ball continuous_on_interior holo2 holomorphic_on_imp_continuous_on interior_cball by blast  
+  then have "isCont (\<lambda>w. zor_poly ff 0 (w - z)) z"
       unfolding isCont_iff by simp
-  qed
-  ultimately show "\<forall>\<^sub>F w in nhds z. zor_poly f z w
-                  = zor_poly ff 0 (w - z)"
+  ultimately show "\<forall>\<^sub>F w in nhds z. zor_poly f z w = zor_poly ff 0 (w - z)"
     by (elim at_within_isCont_imp_nhds;auto)
 qed
 
@@ -1335,27 +1294,26 @@
   have fg_times:"fg w = (fp w * gp w) * (w - z) powi (fn+gn)" and fgp_nz:"fp w*gp w\<noteq>0"
     when "w\<in>ball z r1 - {z}" for w
   proof -
-    have "f w = fp w * (w - z) powi fn" "fp w\<noteq>0"
+    have "f w = fp w * (w - z) powi fn" "fp w \<noteq> 0"
       using fr(2)[rule_format,of w] that unfolding r1_def by auto
     moreover have "g w = gp w * (w - z) powi gn" "gp w \<noteq> 0"
-      using gr(2)[rule_format, of w] that unfolding r1_def by auto
+      using gr(2) that unfolding r1_def by auto
     ultimately show "fg w = (fp w * gp w) * (w - z) powi (fn+gn)" "fp w*gp w\<noteq>0"
-      using that
-      unfolding fg_def by (auto simp add:power_int_add)
+      using that unfolding fg_def by (auto simp add:power_int_add)
   qed
 
   obtain fgr where [simp]:"fgp z \<noteq> 0" and "fgr > 0"
           and fgr: "fgp holomorphic_on cball z fgr"
                   "\<forall>w\<in>cball z fgr - {z}. fg w = fgp w * (w - z) powi fgn \<and> fgp w \<noteq> 0"
   proof -
-    have "fgp z \<noteq> 0 \<and> (\<exists>r>0. fgp holomorphic_on cball z r
-            \<and> (\<forall>w\<in>cball z r - {z}. fg w = fgp w * (w - z) powi fgn \<and> fgp w \<noteq> 0))"
-      apply (rule zorder_exist[of fg z, folded fgn_def fgp_def])
-      subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] .
-      subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] .
-      subgoal unfolding fg_def using fg_nconst .
-      done
-    then show ?thesis using that by blast
+    have "isolated_singularity_at fg z"
+      unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] .
+    moreover have "not_essential fg z"
+      by (simp add: f_iso f_ness fg_def g_iso g_ness not_essential_times)
+    moreover have "\<exists>\<^sub>F w in at z. fg w \<noteq> 0"
+      using fg_def fg_nconst by blast
+    ultimately show ?thesis 
+      using that zorder_exist[of fg z] fgn_def fgp_def by fastforce
   qed
   define r2 where "r2 = min fgr r1"
   have "r2>0" using \<open>r1>0\<close> \<open>fgr>0\<close> unfolding r2_def by simp
@@ -1368,9 +1326,9 @@
     proof (rule ballI)
       fix w assume "w \<in> ball z r2 - {z}"
       then have "w \<in> ball z r1 - {z}" "w \<in> cball z fgr - {z}"  unfolding r2_def by auto
-      from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)]
-      show "fg w = fgp w * (w - z) powi fgn \<and> fgp w \<noteq> 0
-              \<and> fg w = fp w * gp w * (w - z) powi (fn + gn) \<and> fp w * gp w \<noteq> 0" by auto
+      then show "fg w = fgp w * (w - z) powi fgn \<and> fgp w \<noteq> 0
+              \<and> fg w = fp w * gp w * (w - z) powi (fn + gn) \<and> fp w * gp w \<noteq> 0"
+        using fg_times fgp_nz fgr(2) by blast
     qed
     subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros)
     subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros)
@@ -1398,22 +1356,20 @@
   have f_nconst:"\<exists>\<^sub>Fw in (at z). f w \<noteq> 0" and g_nconst:"\<exists>\<^sub>Fw in (at z).g w\<noteq> 0"
     using fg_nconst by (auto elim!:frequently_elim1)
   define vg where "vg=(\<lambda>w. inverse (g w))"
-  have "zorder (\<lambda>w. f w * vg w) z = zorder f z + zorder vg z"
-    apply (rule zorder_times[OF f_iso _ f_ness,of vg])
-    subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] .
-    subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] .
-    subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1)
-    done
+  have 1: "isolated_singularity_at vg z"
+    by (simp add: g_iso g_ness isolated_singularity_at_inverse vg_def)
+  moreover have 2: "not_essential vg z"
+    by (simp add: g_iso g_ness not_essential_inverse vg_def)
+  moreover have 3: "\<exists>\<^sub>F w in at z. f w * vg w \<noteq> 0"
+    using fg_nconst vg_def by auto
+  ultimately  have "zorder (\<lambda>w. f w * vg w) z = zorder f z + zorder vg z"
+    using zorder_times[OF f_iso _ f_ness] by blast
   then show "zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z"
     using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
     by (auto simp add:field_simps)
 
   have "\<forall>\<^sub>F w in at z. zor_poly (\<lambda>w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w"
-    apply (rule zor_poly_times[OF f_iso _ f_ness,of vg])
-    subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] .
-    subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] .
-    subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1)
-    done
+    using zor_poly_times[OF f_iso _ f_ness,of vg] 1 2 3 by blast
   then show "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w / g w) z w = zor_poly f z w  / zor_poly g z w"
     using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
     by eventually_elim (auto simp add:field_simps)
@@ -1447,15 +1403,14 @@
           by (rule non_zero_neighbour_alt[OF holo \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close>])
       qed
       then show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
-        apply (elim eventually_frequentlyE)
-        by auto
+        by (auto elim: eventually_frequentlyE)
     qed
     then obtain r1 where "g z \<noteq> 0" "r1>0" and r1:"g holomorphic_on cball z r1"
             "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w - z) powi n \<and> g w \<noteq> 0)"
       by auto
     obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> s"
       using assms(4,6) open_contains_cball_eq by blast
-    define r3 where "r3=min r1 r2"
+    define r3 where "r3 \<equiv> min r1 r2"
     have "r3>0" "cball z r3 \<subseteq> s" using \<open>r1>0\<close> r2 unfolding r3_def by auto
     moreover have "g holomorphic_on cball z r3"
       using r1(1) unfolding r3_def by auto
@@ -1517,7 +1472,7 @@
       have "(\<lambda>x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \<midarrow>z\<rightarrow> 0" 
         using tendsto_mult by fastforce
       then have "(\<lambda>x. 1::complex) \<midarrow>z\<rightarrow> 0"
-        by (elim Lim_transform_within_open[where s=UNIV],auto)
+        using Lim_transform_within_open by fastforce
       then show False using LIM_const_eq by fastforce
     qed
     ultimately show ?thesis by fastforce
@@ -1555,13 +1510,11 @@
 lemma zorder_exist_pole:
   fixes f::"complex \<Rightarrow> complex" and z::complex
   defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
-  assumes  holo: "f holomorphic_on s-{z}" and
-          "open s" "z\<in>s"
-      and "is_pole f z"
-  shows "n < 0 \<and> g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> cball z r \<subseteq> s \<and> g holomorphic_on cball z r
+  assumes  holo: "f holomorphic_on S-{z}" and "open S" "z\<in>S" and "is_pole f z"
+  shows "n < 0 \<and> g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> cball z r \<subseteq> S \<and> g holomorphic_on cball z r
     \<and> (\<forall>w\<in>cball z r - {z}. f w  = g w / (w-z) ^ nat (- n) \<and> g w \<noteq>0))"
 proof -
-  obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> s" "g holomorphic_on cball z r"
+  obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> S" "g holomorphic_on cball z r"
             "(\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powi n \<and> g w \<noteq> 0)"
   proof -
     have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r
@@ -1574,16 +1527,15 @@
         using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on
         by fastforce
       from non_zero_neighbour_pole[OF \<open>is_pole f z\<close>] show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
-        apply (elim eventually_frequentlyE)
-        by auto
+        by (auto elim: eventually_frequentlyE)
     qed
     then obtain r1 where "g z \<noteq> 0" "r1>0" and r1:"g holomorphic_on cball z r1"
             "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w - z) powi n \<and> g w \<noteq> 0)"
       by auto
-    obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> s"
+    obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> S"
       using assms(4,5) open_contains_cball_eq by metis
     define r3 where "r3=min r1 r2"
-    have "r3>0" "cball z r3 \<subseteq> s" using \<open>r1>0\<close> r2 unfolding r3_def by auto
+    have "r3>0" "cball z r3 \<subseteq> S" using \<open>r1>0\<close> r2 unfolding r3_def by auto
     moreover have "g holomorphic_on cball z r3"
       using r1(1) unfolding r3_def by auto
     moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w - z) powi n \<and> g w \<noteq> 0)"
@@ -1624,16 +1576,16 @@
 qed
 
 lemma zorder_eqI:
-  assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0"
-  assumes fg_eq:"\<And>w. \<lbrakk>w \<in> s;w\<noteq>z\<rbrakk> \<Longrightarrow> f w = g w * (w - z) powi n"
+  assumes "open S" "z \<in> S" "g holomorphic_on S" "g z \<noteq> 0"
+  assumes fg_eq:"\<And>w. \<lbrakk>w \<in> S;w\<noteq>z\<rbrakk> \<Longrightarrow> f w = g w * (w - z) powi n"
   shows   "zorder f z = n"
 proof -
-  have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact
+  have "continuous_on S g" by (rule holomorphic_on_imp_continuous_on) fact
   moreover have "open (-{0::complex})" by auto
-  ultimately have "open ((g -` (-{0})) \<inter> s)"
-    unfolding continuous_on_open_vimage[OF \<open>open s\<close>] by blast
-  moreover from assms have "z \<in> (g -` (-{0})) \<inter> s" by auto
-  ultimately obtain r where r: "r > 0" "cball z r \<subseteq>  s \<inter> (g -` (-{0}))"
+  ultimately have "open ((g -` (-{0})) \<inter> S)"
+    unfolding continuous_on_open_vimage[OF \<open>open S\<close>] by blast
+  moreover from assms have "z \<in> (g -` (-{0})) \<inter> S" by auto
+  ultimately obtain r where r: "r > 0" "cball z r \<subseteq>  S \<inter> (g -` (-{0}))"
     unfolding open_contains_cball by blast
 
   let ?gg= "(\<lambda>w. g w * (w - z) powi n)"
@@ -1644,18 +1596,18 @@
   then have "\<exists>g r. P n g r" by auto
   moreover have unique: "\<exists>!n. \<exists>g r. P n g r" unfolding P_def
   proof (rule holomorphic_factor_puncture)
-    have "ball z r-{z} \<subseteq> s" using r using ball_subset_cball by blast
+    have "ball z r-{z} \<subseteq> S" using r using ball_subset_cball by blast
     then have "?gg holomorphic_on ball z r-{z}"
-      using \<open>g holomorphic_on s\<close> r by (auto intro!: holomorphic_intros)
+      using \<open>g holomorphic_on S\<close> r by (auto intro!: holomorphic_intros)
     then have "f holomorphic_on ball z r - {z}"
-      by (smt (verit, best) DiffD2 \<open>ball z r-{z} \<subseteq> s\<close> fg_eq holomorphic_cong singleton_iff subset_iff)
+      by (smt (verit, best) DiffD2 \<open>ball z r-{z} \<subseteq> S\<close> fg_eq holomorphic_cong singleton_iff subset_iff)
     then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
       using analytic_on_open open_delete r(1) by blast
   next
     have "not_essential ?gg z"
     proof (intro singularity_intros)
       show "not_essential g z"
-        by (meson \<open>continuous_on s g\<close> assms continuous_on_eq_continuous_at
+        by (meson \<open>continuous_on S g\<close> assms continuous_on_eq_continuous_at
             isCont_def not_essential_def)
       show " \<forall>\<^sub>F w in at z. w - z \<noteq> 0" by (simp add: eventually_at_filter)
       then show "LIM w at z. w - z :> at 0"
@@ -1679,7 +1631,7 @@
       proof (subst fg_eq[OF _ \<open>z'\<noteq>z\<close>])
         have "z' \<in> cball z r"
           unfolding z'_def using \<open>r>0\<close> \<open>d>0\<close> by (auto simp add:dist_norm)
-        then show " z' \<in> s" using r(2) by blast
+        then show " z' \<in> S" using r(2) by blast
         show "g z' * (z' - z) powi n \<noteq> 0"
           using P_def \<open>P n g r\<close> \<open>z' \<in> cball z r\<close> \<open>z' \<noteq> z\<close> by auto
       qed
@@ -1692,8 +1644,8 @@
 qed
 
 lemma simple_zeroI:
-  assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0"
-  assumes "\<And>w. w \<in> s \<Longrightarrow> f w = g w * (w - z)"
+  assumes "open S" "z \<in> S" "g holomorphic_on S" "g z \<noteq> 0"
+  assumes "\<And>w. w \<in> S \<Longrightarrow> f w = g w * (w - z)"
   shows   "zorder f z = 1"
   using assms zorder_eqI by force
 
@@ -1726,13 +1678,13 @@
 qed
 
 lemma zorder_zero_eqI:
-  assumes  f_holo:"f holomorphic_on s" and "open s" "z \<in> s"
+  assumes  f_holo:"f holomorphic_on S" and "open S" "z \<in> S"
   assumes zero: "\<And>i. i < nat n \<Longrightarrow> (deriv ^^ i) f z = 0"
   assumes nz: "(deriv ^^ nat n) f z \<noteq> 0" and "n\<ge>0"
   shows   "zorder f z = n"
 proof -
-  obtain r where [simp]:"r>0" and "ball z r \<subseteq> s"
-    using \<open>open s\<close> \<open>z\<in>s\<close> openE by blast
+  obtain r where [simp]:"r>0" and "ball z r \<subseteq> S"
+    using \<open>open S\<close> \<open>z\<in>S\<close> openE by blast
   have nz':"\<exists>w\<in>ball z r. f w \<noteq> 0"
   proof (rule ccontr)
     assume "\<not> (\<exists>w\<in>ball z r. f w \<noteq> 0)"
@@ -1746,17 +1698,17 @@
   qed
 
   define zn g where "zn = zorder f z" and "g = zor_poly f z"
-  obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and
-            [simp]:"e>0" and "cball z e \<subseteq> ball z r" and
-            g_holo:"g holomorphic_on cball z e" and
-            e_fac:"(\<forall>w\<in>cball z e. f w = g w * (w - z) ^ nat zn \<and> g w \<noteq> 0)"
+  obtain e where e_if: "if f z = 0 then 0 < zn else zn = 0" and
+            [simp]: "e>0" and "cball z e \<subseteq> ball z r" and
+            g_holo: "g holomorphic_on cball z e" and
+            e_fac: "(\<forall>w\<in>cball z e. f w = g w * (w - z) ^ nat zn \<and> g w \<noteq> 0)"
   proof -
     have "f holomorphic_on ball z r"
-      using f_holo \<open>ball z r \<subseteq> s\<close> by auto
+      using f_holo \<open>ball z r \<subseteq> S\<close> by auto
     from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def]
-    show ?thesis by blast
+    show thesis by blast
   qed
-  then obtain "zn \<ge> 0" "g z\<noteq>0"
+  then obtain "zn \<ge> 0" "g z \<noteq> 0"
     by (metis centre_in_cball less_le_not_le order_refl)
 
   define A where "A \<equiv> (\<lambda>i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)"
@@ -1850,10 +1802,7 @@
   assumes "isolated_singularity_at f z" "not_essential f z"
   assumes "g analytic_on {z}" "frequently (\<lambda>z. f z * g z \<noteq> 0) (at z)"
   shows   "zorder (\<lambda>x. f x * g x) z = zorder f z + zorder g z"
-proof (rule zorder_times)
-  show "isolated_singularity_at g z" "not_essential g z"
-    by (intro isolated_singularity_at_analytic not_essential_analytic assms)+
-qed (use assms in auto)
+  using assms isolated_singularity_at_analytic not_essential_analytic zorder_times by blast
 
 lemma zorder_cmult:
   assumes "c \<noteq> 0"
@@ -1874,7 +1823,7 @@
 qed
 
 lemma zorder_nonzero_div_power:
-  assumes sz: "open s" "z \<in> s" "f holomorphic_on s" "f z \<noteq> 0" and "n > 0"
+  assumes sz: "open S" "z \<in> S" "f holomorphic_on S" "f z \<noteq> 0" and "n > 0"
   shows  "zorder (\<lambda>w. f w / (w - z) ^ n) z = - n"
   by (intro zorder_eqI [OF sz]) (simp add: inverse_eq_divide power_int_minus)
 
@@ -1893,7 +1842,7 @@
 qed
 
 lemma zor_poly_zero_eq:
-  assumes "f holomorphic_on s" "open s" "connected s" "z \<in> s" "\<exists>w\<in>s. f w \<noteq> 0"
+  assumes "f holomorphic_on S" "open S" "connected S" "z \<in> S" "\<exists>w\<in>S. f w \<noteq> 0"
   shows "eventually (\<lambda>w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)"
 proof -
   obtain r where r:"r>0"
@@ -2104,7 +2053,7 @@
   have *: "x \<in> ball x r" "open (ball x r)" "open (ball x r - {x})"
     using \<open>r > 0\<close> by auto
   show "is_pole (deriv f) x" "zorder (deriv f) x = zorder f x - 1"
-    by (rule is_pole_deriv' zorder_deriv', (rule assms * holomorphic_derivI holo | assumption)+)+
+    by (meson "*" assms(1) holo holomorphic_derivI is_pole_deriv' zorder_deriv')+
 qed
 
 lemma removable_singularity_deriv':
@@ -2271,11 +2220,7 @@
   have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
     using r by (intro eventually_at_in_open) auto
   thus "eventually (\<lambda>w. f w = 0) (at z)"
-  proof eventually_elim
-    case (elim w)
-    thus ?case using g_eq_0[of w]
-      by (auto simp: g_def)
-  qed
+    by (metis freq non_zero_neighbour not_eventually not_frequently sing)
 qed
 
 lemma pole_imp_not_constant:
@@ -2337,8 +2282,7 @@
   qed
   moreover have "\<forall>\<^sub>F w in at z. f w =  P w * (w - z) powi n"
     unfolding eventually_at_le
-    apply (rule exI[where x=r])
-    using w_Pn \<open>r>0\<close> by (simp add: dist_commute)
+    using w_Pn \<open>r>0\<close> by (force simp add: dist_commute)
   ultimately show ?thesis using is_pole_cong by fast
 qed
 
@@ -2387,34 +2331,20 @@
     using is_pole_deriv[OF \<open>is_pole f z\<close> f_iso,THEN non_zero_neighbour_pole]
     .
   ultimately have "\<forall>\<^sub>F w in at z. False"
-    apply eventually_elim
-    by auto
+    by eventually_elim auto
   then show False by auto
 qed
 
 lemma isolated_pole_imp_neg_zorder:
-  assumes f_iso:"isolated_singularity_at f z"
-    and "is_pole f z"
-  shows "zorder f z<0"
-proof -
-  obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}"
-    using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast
-  show ?thesis
-    using zorder_exist_pole[OF f_holo,simplified,OF \<open>is_pole f z\<close>]
-    by auto
-qed
+  assumes "isolated_singularity_at f z" and "is_pole f z"
+  shows "zorder f z < 0"
+  using analytic_imp_holomorphic assms centre_in_ball isolated_singularity_at_def zorder_exist_pole by blast
+
 
 lemma isolated_singularity_at_deriv[singularity_intros]:
   assumes "isolated_singularity_at f x"
   shows "isolated_singularity_at (deriv f) x"
-proof -
-  obtain r where "r>0" "f analytic_on ball x r - {x}"
-    using assms unfolding isolated_singularity_at_def by auto
-  from analytic_deriv[OF this(2)]
-  have "deriv f analytic_on ball x r - {x}" .
-  with \<open>r>0\<close> show ?thesis
-    unfolding isolated_singularity_at_def by auto
-qed
+  by (meson analytic_deriv assms isolated_singularity_at_def)
 
 lemma zorder_deriv_minus_1:
   fixes f g::"complex \<Rightarrow> complex" and z::complex
@@ -2501,13 +2431,9 @@
 
   have "zorder ff z =  zorder (deriv f) z - zorder f z"
     unfolding ff_def using f_iso f_ness fg_nconst
-    apply (rule_tac zorder_divide)
-    by (auto intro:singularity_intros)
+    using isolated_singularity_at_deriv not_essential_deriv zorder_divide by blast
   moreover have "zorder (deriv f) z = zorder f z - 1"
-  proof (rule zorder_deriv_minus_1)
-    show " \<exists>\<^sub>F w in at z. f w \<noteq> 0"
-      using fg_nconst frequently_elim1 by fastforce
-  qed (use f_iso f_ness f_ord in auto)
+    using f_iso f_ness f_ord fg_nconst frequently_elim1 zorder_deriv_minus_1 by fastforce
   ultimately show "zorder ff z < 0" by auto
     
   show "\<exists>\<^sub>F w in at z. ff w \<noteq> 0" 
@@ -2523,8 +2449,7 @@
   show "not_essential f z" 
     using \<open>is_pole f z\<close> unfolding not_essential_def by auto
   show "\<exists>\<^sub>F w in at z. deriv f w * f w \<noteq> 0"
-    apply (rule isolated_pole_imp_nzero_times)
-    using assms by auto
+    using assms f_iso isolated_pole_imp_nzero_times by blast
   show "zorder f z \<noteq> 0"
     using isolated_pole_imp_neg_zorder assms by fastforce
 qed
@@ -2541,7 +2466,7 @@
   assumes "isolated_zero f x" "isolated_zero g x"
   shows   "isolated_zero (\<lambda>x. f x * g x) x"
 proof -
-  have "eventually (\<lambda>x. f x \<noteq> 0) (at x)" "eventually (\<lambda>x. g x \<noteq> 0) (at x)"
+  have "\<forall>\<^sub>F x in at x. f x \<noteq> 0" "\<forall>\<^sub>F x in at x. g x \<noteq> 0"
     using assms unfolding isolated_zero_def by auto
   hence "eventually (\<lambda>x. f x * g x \<noteq> 0) (at x)"
     by eventually_elim auto
@@ -2576,10 +2501,7 @@
 lemma non_isolated_zero':
   assumes "isolated_singularity_at f z" "not_essential f z" "f z = 0" "\<not>isolated_zero f z"
   shows   "eventually (\<lambda>z. f z = 0) (at z)"
-proof (rule not_essential_frequently_0_imp_eventually_0)
-  from assms show "frequently (\<lambda>z. f z = 0) (at z)"
-    by (auto simp: frequently_def isolated_zero_def)
-qed fact+
+  by (metis assms isolated_zero_def non_zero_neighbour not_eventually)
 
 lemma non_isolated_zero:
   assumes "\<not>isolated_zero f z" "f analytic_on {z}" "f z = 0"
--- a/src/HOL/Complex_Analysis/Conformal_Mappings.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Complex_Analysis/Conformal_Mappings.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -30,8 +30,7 @@
     by (metis \<open>m \<noteq> 0\<close> dist_norm mem_ball norm_minus_commute not_gr_zero)
   have "0 < min r s"  by (simp add: \<open>0 < r\<close> \<open>0 < s\<close>)
   then show thesis
-    apply (rule that)
-    using r s by auto
+  proof qed (use r s in auto)
 qed
 
 proposition analytic_continuation:
@@ -166,7 +165,7 @@
     then have "0 < norm (f \<xi>)"
       by (simp add: \<open>0 < r\<close>)
     have fnz': "\<And>w. w \<in> cball \<xi> r \<Longrightarrow> f w \<noteq> 0"
-      by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero)
+      using dist_complex_def fnz norm_less order_le_less by fastforce
     have "frontier(cball \<xi> r) \<noteq> {}"
       using \<open>0 < r\<close> by simp
     define g where [abs_def]: "g z = inverse (f z)" for z
@@ -204,7 +203,7 @@
       by (metis (no_types) dist_norm frontier_cball mem_sphere w)
     ultimately obtain wr: "norm (\<xi> - w) = r" and nfw: "norm (f w) \<le> norm (f \<xi>)"
       unfolding g_def
-        by (metis (no_types) \<open>0 < cmod (f \<xi>)\<close> less_imp_inverse_less norm_inverse not_le now order_trans v)
+      by (smt (verit, del_insts) \<open>0 < cmod (f \<xi>)\<close> inverse_le_imp_le norm_inverse now v)
     with fw have False
       using norm_less by force
   }
@@ -339,13 +338,10 @@
 qed
 
 corollary\<^marker>\<open>tag unimportant\<close> open_mapping_thm3:
-  assumes holf: "f holomorphic_on S"
-      and "open S" and injf: "inj_on f S"
+  assumes "f holomorphic_on S"
+      and "open S" and  "inj_on f S"
     shows  "open (f ` S)"
-proof (rule open_mapping_thm2 [OF holf])
-  show "\<And>X. \<lbrakk>open X; X \<subseteq> S; X \<noteq> {}\<rbrakk> \<Longrightarrow> \<not> f constant_on X"
-    using inj_on_subset injective_not_constant injf by blast
-qed (use assms in auto)
+  by (meson assms inj_on_subset injective_not_constant open_mapping_thm2 order.refl)
 
 subsection\<open>Maximum modulus principle\<close>
 
@@ -423,7 +419,7 @@
     then obtain w where w: "w \<in> frontier(connected_component_set (interior S) z)"
        by auto
     then have "norm (f z) = norm (f w)"  by (simp add: "2" c cc frontier_def)
-    also have "... \<le> B"
+    also have "\<dots> \<le> B"
       using w frontier_interior_subset frontier_of_connected_component_subset 
       by (blast intro: leB)
     finally show ?thesis .
@@ -434,14 +430,14 @@
 
 corollary\<^marker>\<open>tag unimportant\<close> maximum_real_frontier:
   assumes holf: "f holomorphic_on (interior S)"
-      and contf: "continuous_on (closure S) f"
-      and bos: "bounded S"
-      and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> Re(f z) \<le> B"
-      and "\<xi> \<in> S"
-    shows "Re(f \<xi>) \<le> B"
-using maximum_modulus_frontier [of "exp o f" S "exp B"]
-      Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms
-by auto
+    and contf: "continuous_on (closure S) f"
+    and bos: "bounded S"
+    and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> Re(f z) \<le> B"
+    and "\<xi> \<in> S"
+  shows "Re(f \<xi>) \<le> B"
+  using maximum_modulus_frontier [of "exp o f" S "exp B"]
+    Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms
+  by auto
 
 subsection\<^marker>\<open>tag unimportant\<close> \<open>Factoring out a zero according to its order\<close>
 
@@ -541,9 +537,10 @@
     by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open)
   then have con: "continuous_on (ball \<xi> r) (\<lambda>x. exp (h x) / g x)"
     by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne)
+  have gfd: "dist \<xi> x < r \<Longrightarrow> g field_differentiable at x" if "dist \<xi> x < r" for x
+    using holg holomorphic_on_imp_differentiable_at by auto
   have 0: "dist \<xi> x < r \<Longrightarrow> ((\<lambda>x. exp (h x) / g x) has_field_derivative 0) (at x)" for x
-    apply (rule h derivative_eq_intros DERIV_deriv_iff_field_differentiable [THEN iffD2] | simp)+
-    using holg by (auto simp: holomorphic_on_imp_differentiable_at gne h)
+    by (rule gfd h derivative_eq_intros DERIV_deriv_iff_field_differentiable [THEN iffD2] | simp add: gne)+
   obtain c where c: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> exp (h x) / g x = c"
     by (rule DERIV_zero_connected_constant [of "ball \<xi> r" "{}" "\<lambda>x. exp(h x) / g x"]) (auto simp: con 0)
   have hol: "(\<lambda>z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \<xi> r"
@@ -564,8 +561,8 @@
   fixes k :: "'a::wellorder"
   assumes a_def: "a == LEAST x. P x" and P: "P k"
   shows def_LeastI: "P a" and def_Least_le: "a \<le> k"
-unfolding a_def
-by (rule LeastI Least_le; rule P)+
+  unfolding a_def
+  by (rule LeastI Least_le; rule P)+
 
 lemma holomorphic_factor_zero_nonconstant:
   assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
@@ -646,8 +643,8 @@
   then have leg: "\<And>w. w \<in> cball \<xi> d \<Longrightarrow> norm x \<le> norm (g w)"
     by auto
   have "ball \<xi> d \<subseteq> cball \<xi> d" by auto
-  also have "... \<subseteq> ball \<xi> e" using \<open>0 < d\<close> d_def by auto
-  also have "... \<subseteq> S" by (rule e)
+  also have "\<dots> \<subseteq> ball \<xi> e" using \<open>0 < d\<close> d_def by auto
+  also have "\<dots> \<subseteq> S" by (rule e)
   finally have dS: "ball \<xi> d \<subseteq> S" .
   have "x \<noteq> 0" using gnz x \<open>d < r\<close> by auto
   show thesis
@@ -687,20 +684,20 @@
       using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast
     then have "(g \<longlongrightarrow> g \<xi>) (at \<xi>)"
       by (simp add: \<xi>)
+    then have "\<forall>\<^sub>F z in at \<xi>. cmod (f z) \<le> cmod (g \<xi>) + 1"
+      by (rule eventually_mp [OF * tendstoD [where e=1]], auto)
     then show ?thesis
-      apply (rule_tac x="norm(g \<xi>) + 1" in exI)
-      apply (rule eventually_mp [OF * tendstoD [where e=1]], auto)
-      done
+      by blast
   qed
   moreover have "?Q" if "\<forall>\<^sub>F z in at \<xi>. cmod (f z) \<le> B" for B
     by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero)
   moreover have "?P" if "(\<lambda>z. (z - \<xi>) * f z) \<midarrow>\<xi>\<rightarrow> 0"
   proof -
     define h where [abs_def]: "h z = (z - \<xi>)^2 * f z" for z
-    have h0: "(h has_field_derivative 0) (at \<xi>)"
-      apply (simp add: h_def has_field_derivative_iff)
-      apply (auto simp: field_split_simps power2_eq_square Lim_transform_within [OF that, of 1])
-      done
+    have "(\<lambda>y. (y - \<xi>)\<^sup>2 * f y / (y - \<xi>)) \<midarrow>\<xi>\<rightarrow> 0"
+      by (simp add: LIM_cong power2_eq_square that)
+    then have h0: "(h has_field_derivative 0) (at \<xi>)"
+      by (simp add: h_def has_field_derivative_iff)
     have holh: "h holomorphic_on S"
     proof (simp add: holomorphic_on_def, clarify)
       fix z assume "z \<in> S"
@@ -755,22 +752,20 @@
         by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+
       have 2: "0 \<in> interior (ball 0 r)"
         using \<open>0 < r\<close> by simp
-      have "\<exists>B. 0<B \<and> eventually (\<lambda>z. cmod ((inverse \<circ> f \<circ> inverse) z) \<le> B) (at 0)"
-        apply (rule exI [where x=1])
-        using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one]
-        by (simp add: eventually_mono)
-      with holomorphic_on_extend_bounded [OF 1 2]
       obtain g where holg: "g holomorphic_on ball 0 r"
                  and geq: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> g z = (inverse \<circ> f \<circ> inverse) z"
-        by meson
+        using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] holomorphic_on_extend_bounded [OF 1 2]
+        by (smt (verit, del_insts) \<open>l = 0\<close> eventually_mono norm_conv_dist)
       have ifi0: "(inverse \<circ> f \<circ> inverse) \<midarrow>0\<rightarrow> 0"
         using \<open>l = 0\<close> lim lim_at_infinity_0 by blast
       have g2g0: "g \<midarrow>0\<rightarrow> g 0"
         using \<open>0 < r\<close> centre_in_ball continuous_at continuous_on_eq_continuous_at holg
         by (blast intro: holomorphic_on_imp_continuous_on)
       have g2g1: "g \<midarrow>0\<rightarrow> 0"
-        apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]])
-        using \<open>0 < r\<close> by (auto simp: geq)
+      proof (rule Lim_transform_within_open [OF ifi0 open_ball])
+        show "\<And>x. \<lbrakk>x \<in> ball 0 r; x \<noteq> 0\<rbrakk> \<Longrightarrow> (inverse \<circ> f \<circ> inverse) x = g x"
+          by (auto simp: geq)
+      qed (auto simp: \<open>0 < r\<close>)
       have [simp]: "g 0 = 0"
         by (rule tendsto_unique [OF _ g2g0 g2g1]) simp
       have "ball 0 r - {0::complex} \<noteq> {}"
@@ -831,7 +826,7 @@
           using **[of w] fi0 \<open>0 < r\<close>  that by force
         then show ?thesis
           unfolding lim_at_infinity_0
-          using eventually_at \<open>r > 0\<close> by (force simp add: intro: tendsto_eventually)
+          using eventually_at \<open>r > 0\<close> by (force simp: intro: tendsto_eventually)
       qed
       obtain w where "w \<in> ball 0 r - {0}" and "f (inverse w) = 0"
         using False \<open>0 < r\<close> by blast
@@ -916,7 +911,7 @@
     with k m show ?thesis
       by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right)
   qed
-  have \<section>: "((inverse \<circ> f) \<longlongrightarrow> 0) at_infinity"
+  have *: "((inverse \<circ> f) \<longlongrightarrow> 0) at_infinity"
   proof (rule Lim_at_infinityI)
     fix e::real assume "0 < e"
     with compf [of "cball 0 (inverse e)"]
@@ -927,7 +922,7 @@
   qed
   then obtain a n where "\<And>z. f z = (\<Sum>i\<le>n. a i * z^i)"
     using assms pole_at_infinity by blast
-  with \<section> 2 show ?rhs by blast
+  with * 2 show ?rhs by blast
 next
   assume ?rhs
   then obtain c n where "0 < n" "c n \<noteq> 0" "f = (\<lambda>z. \<Sum>i\<le>n. c i * z ^ i)" by blast
@@ -967,11 +962,12 @@
     using dnz by simp
   then obtain g' where g': "linear g'" "g' \<circ> (*) (deriv f \<xi>) = id"
     using linear_injective_left_inverse [of "(*) (deriv f \<xi>)"] by auto
+
+  have fder: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative (*) (deriv f x)) (at x)"
+    using \<open>open S\<close> has_field_derivative_imp_has_derivative holf holomorphic_derivI by blast
   show ?thesis
     apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\<lambda>z h. deriv f z * h" and g' = g'])
-    using g' * 
-    apply (simp_all add: linear_conv_bounded_linear that)
-    using \<open>open S\<close> has_field_derivative_imp_has_derivative holf holomorphic_derivI by blast
+    using g' * by (simp_all add: fder linear_conv_bounded_linear that)
 qed
 
 lemma has_complex_derivative_locally_invertible:
@@ -1175,7 +1171,7 @@
     using continuous_openin_preimage [OF contg gim]
     by (meson \<open>open V\<close> contg continuous_openin_preimage_eq)
   ultimately obtain \<epsilon> where "\<epsilon>>0" and e: "ball z \<epsilon> \<inter> U \<subseteq> g -` V"
-    by (force simp add: openin_contains_ball)
+    by (force simp: openin_contains_ball)
   show "g field_differentiable at z within U"
   proof (rule field_differentiable_transform_within)
     show "(0::real) < \<epsilon>"
@@ -1398,7 +1394,7 @@
     show "convex (S \<inter> {x. d \<bullet> x \<le> k})"
       by (rule convex_Int [OF \<open>convex S\<close> convex_halfspace_le])
   qed
-  also have "... \<subseteq> {z \<in> S. d \<bullet> z < k}"
+  also have "\<dots> \<subseteq> {z \<in> S. d \<bullet> z < k}"
     by (force simp: interior_open [OF \<open>open S\<close>] \<open>d \<noteq> 0\<close>)
   finally have *: "interior (convex hull {a, b, c}) \<subseteq> {z \<in> S. d \<bullet> z < k}" .
   have "continuous_on (convex hull {a,b,c}) f"
@@ -1584,16 +1580,17 @@
       using \<open>p \<in> S\<close> openE S by blast
     then have "continuous_on (ball p e) f"
       using contf continuous_on_subset by blast
-    moreover have "f holomorphic_on {z. dist p z < e \<and> d \<bullet> z < k}"
-      apply (rule holomorphic_on_subset [OF holf1])
+    moreover 
+    have "{z. dist p z < e \<and> d \<bullet> z < k} \<subseteq> S \<inter> {z. d \<bullet> z < k}" 
+         "{z. dist p z < e \<and> k < d \<bullet> z} \<subseteq> S \<inter> {z. k < d \<bullet> z}"
       using e by auto
-    moreover have "f holomorphic_on {z. dist p z < e \<and> k < d \<bullet> z}"
-      apply (rule holomorphic_on_subset [OF holf2])
-      using e by auto
+    then have "f holomorphic_on {z. dist p z < e \<and> d \<bullet> z < k}" 
+              "f holomorphic_on {z. dist p z < e \<and> k < d \<bullet> z}"
+      using holomorphic_on_subset holf1 holf2 by presburger+
     ultimately show ?thesis
       apply (rule_tac x="ball p e" in exI)
-      using \<open>e > 0\<close> e \<open>d \<noteq> 0\<close> hol_pal_lem4 [of "ball p e" _ _ _ d _ k]
-      by (force simp add: subset_hull)
+      using \<open>e > 0\<close> e \<open>d \<noteq> 0\<close> hol_pal_lem4 [of "ball p e" _ _ _ d _ k] 
+      by (force simp: subset_hull)
   qed
   show ?thesis
     by (blast intro: * Morera_local_triangle analytic_imp_holomorphic)
@@ -1730,12 +1727,12 @@
       proof -
         have [simp]: "x * norm z < r"
           using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero)
-        have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \<le> norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C"
-          apply (rule Le1) using r x \<open>0 < r\<close> by simp
-        also have "... \<le> norm (x *\<^sub>R z) / (r - norm z) * C"
-          using r x \<open>0 < r\<close>
-          apply (simp add: field_split_simps)
-          by (simp add: \<open>0 < C\<close> mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono)
+        then have "cmod (x *\<^sub>R z) < r"
+          by (simp add: x)
+        then have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \<le> norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C"
+          by (metis Le1) 
+        also have "\<dots> \<le> norm (x *\<^sub>R z) / (r - norm z) * C"
+          using r x \<open>0 < r\<close> \<open>0 < C\<close> by (simp add: frac_le mult_left_le_one_le)
         finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm (x *\<^sub>R z)  / (r - norm z) * C * norm z"
           by (rule mult_right_mono) simp
         with x show ?thesis by (simp add: algebra_simps)
@@ -1762,9 +1759,8 @@
       then have \<section>: "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \<le> norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)"
         by (simp add: algebra_simps)
       show ?thesis
-        apply (rule le_norm [OF _ int_le])
-        using \<open>norm z < r\<close>
-        by (simp add: power2_eq_square divide_simps C_def norm_mult \<section>)
+        using \<open>norm z < r\<close> 
+        by (force simp add: power2_eq_square divide_simps C_def norm_mult \<section> intro!: le_norm [OF _ int_le])
     qed
     have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2)  < 1"
       by (auto simp:  sqrt2_less_2)
@@ -1788,20 +1784,20 @@
     have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) =
           ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))"
       by simp
-    also have "...  \<subseteq> f ` ball 0 ((1 - sqrt 2 / 2) * r)"
+    also have "\<dots>  \<subseteq> f ` ball 0 ((1 - sqrt 2 / 2) * r)"
     proof -
       have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \<le> norm (f z)"
         if "norm z = (1 - sqrt 2 / 2) * r" for z
-        apply (rule order_trans [OF _ *])
-        using  \<open>0 < r\<close>
-         apply (simp_all add: field_simps power2_eq_square that)
-        apply (simp add: mult.assoc [symmetric])
-        done
+      proof (rule order_trans [OF _ *])
+        show "(3 - 2 * sqrt 2) * r * cmod (deriv f 0)
+           \<le> (cmod z - (cmod z)\<^sup>2 / (r - cmod z)) * cmod (deriv f 0)"
+          by (simp add: le_less algebra_simps divide_simps power2_eq_square that)
+      qed (use \<open>0 < r\<close> that in auto)
       show ?thesis
-        apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball])
-        using \<open>0 < r\<close> sq201 3 C_def \<open>0 < C\<close> sq3 by auto
+        using \<open>0 < r\<close> sq201 3 C_def \<open>0 < C\<close> sq3 
+        by (intro ball_subset_open_map_image [OF 1 2 _ bounded_ball]) auto
     qed
-    also have "...  \<subseteq> f ` ball 0 r"
+    also have "\<dots> \<subseteq> f ` ball 0 r"
     proof -
       have "\<And>x. (1 - sqrt 2 / 2) * r \<le> r"
         using \<open>0 < r\<close> by (auto simp: field_simps)
@@ -1833,7 +1829,7 @@
                 \<subseteq> (\<lambda>z. f (a + z) - f a) ` ball 0 r"
     apply (rule Bloch_lemma_0)
     using \<open>0 < r\<close>
-      apply (simp_all add: \<open>0 < r\<close>)
+      apply (simp_all add: \<open>0 < r\<close> )
     apply (simp add: fz deriv_chain dist_norm le)
     done
   show ?thesis
@@ -1892,7 +1888,7 @@
     using gen_le_dfp [of a] \<open>r > 0\<close> by auto
   have 1: "f holomorphic_on cball p t"
     using cpt \<open>r < 1\<close> order_subst1 subset_ball
-    by (force simp add: intro!: holomorphic_on_subset [OF holf])
+    by (force simp: intro!: holomorphic_on_subset [OF holf])
   have 2: "norm (deriv f z) \<le> 2 * norm (deriv f p)" if "z \<in> ball p t" for z
   proof -
     have z: "z \<in> cball a r"
@@ -1904,7 +1900,7 @@
     with \<open>norm (z - a) < r\<close> \<open>norm (p - a) < r\<close>
     have "norm (deriv f z) \<le> (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)"
       by (simp add: field_simps)
-    also have "... \<le> 2 * norm (deriv f p)"
+    also have "\<dots> \<le> 2 * norm (deriv f p)"
     proof (rule mult_right_mono)
       show "(r - cmod (p - a)) / (r - cmod (z - a)) \<le> 2"
         using that \<open>norm (p - a) < r\<close> \<open>norm(z - a) < r\<close> dist_triangle3 [of z a p] 
@@ -1917,7 +1913,7 @@
   then have sq3: "0 < 3 - 2 * sqrt 2" by simp
   have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r"
     using sq3 sqrt2 by (auto simp: field_simps r_def)
-  also have "... \<le> cmod (deriv f p) * (r - cmod (p - a))"
+  also have "\<dots> \<le> cmod (deriv f p) * (r - cmod (p - a))"
     using \<open>norm (p - a) < r\<close> le_norm_dfp   by (simp add: pos_divide_le_eq)
   finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)"
     using pos_divide_less_eq half_gt_zero_iff sq3 by blast
@@ -1925,13 +1921,8 @@
     using sq3 by (simp add: mult.commute t_def)
   have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball p t"
     by (rule Bloch_lemma [OF 1 \<open>0 < t\<close> 2])
-  also have "... \<subseteq> f ` ball a 1"
-  proof -
-    have "ball a r \<subseteq> ball a 1"
-      using \<open>0 < t\<close> \<open>r < 1\<close> by (simp add: ball_subset_ball_iff dist_norm)
-    then show ?thesis
-      using ball_subset_cball cpt by blast
-  qed
+  also have "\<dots> \<subseteq> f ` ball a 1"
+    by (meson \<open>r < 1\<close> ball_subset_cball cpt dual_order.trans image_mono less_le_not_le subset_ball)
   finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball a 1" .
   with ** show ?thesis
     by (rule that)
@@ -1970,11 +1961,17 @@
       apply (rule derivative_eq_intros * | simp)+
       using \<open>0 < r\<close> by (auto simp: C_def False)
   qed
-  have "deriv (\<lambda>z. f (a + of_real r * z) / (C * of_real r)) 0 = deriv (\<lambda>z. f (a + complex_of_real r * z)) 0 /
-    (C * complex_of_real r)"
-    apply (rule deriv_cdivide_right)
-    by (metis (no_types) DERIV_chain2 add.right_neutral dfa field_differentiable_add_const field_differentiable_def field_differentiable_linear fo mult_zero_right)
-  also have "... = 1"
+  obtain f' where "(f has_field_derivative f') (at a)"
+    using dfa field_differentiable_def by blast
+  then have "\<exists>c. ((\<lambda>c. f (a + complex_of_real r * c)) has_field_derivative c) (at 0)"
+    by (metis (no_types) DERIV_chain2 add_cancel_left_right field_differentiable_add_const 
+        field_differentiable_def field_differentiable_linear mult_eq_0_iff)
+  then have "(\<lambda>w. f (a + complex_of_real r * w)) field_differentiable at 0"
+    by (simp add: field_differentiable_def)
+  then have "deriv (\<lambda>z. f (a + of_real r * z) / (C * of_real r)) 0 
+           = deriv (\<lambda>z. f (a + of_real r * z)) 0 / (C * of_real r)"
+    by (rule deriv_cdivide_right)
+  also have "\<dots> = 1"
     using \<open>0 < r\<close> by (simp add: C_def False fo derivative_intros dfa deriv_chain)
   finally have 2: "deriv (\<lambda>z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" .
   have sb1: "(*) (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1
--- a/src/HOL/Complex_Analysis/Contour_Integration.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Complex_Analysis/Contour_Integration.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -44,14 +44,14 @@
   unfolding contour_integrable_on_def contour_integral_def by blast
 
 lemma contour_integral_unique: "(f has_contour_integral i) g \<Longrightarrow> contour_integral g f = i"
-  apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def)
+  unfolding contour_integral_def has_contour_integral_def contour_integrable_on_def
   using has_integral_unique by blast
 
 lemma has_contour_integral_eqpath:
-     "\<lbrakk>(f has_contour_integral y) p; f contour_integrable_on \<gamma>;
+  "\<lbrakk>(f has_contour_integral y) p; f contour_integrable_on \<gamma>;
        contour_integral p f = contour_integral \<gamma> f\<rbrakk>
       \<Longrightarrow> (f has_contour_integral y) \<gamma>"
-using contour_integrable_on_def contour_integral_unique by auto
+  using contour_integrable_on_def contour_integral_unique by auto
 
 lemma has_contour_integral_integral:
     "f contour_integrable_on i \<Longrightarrow> (f has_contour_integral (contour_integral i f)) i"
@@ -329,12 +329,12 @@
 qed
 
 lemma contour_integrable_join [simp]:
-    "\<lbrakk>valid_path g1; valid_path g2\<rbrakk>
+  "\<lbrakk>valid_path g1; valid_path g2\<rbrakk>
      \<Longrightarrow> f contour_integrable_on (g1 +++ g2) \<longleftrightarrow> f contour_integrable_on g1 \<and> f contour_integrable_on g2"
-using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast
+  using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast
 
 lemma contour_integral_join [simp]:
-    "\<lbrakk>f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\<rbrakk>
+  "\<lbrakk>f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\<rbrakk>
         \<Longrightarrow> contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f"
   by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique)
 
@@ -353,11 +353,7 @@
     using assms by (auto simp: has_contour_integral)
   then have i: "i = integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)) +
                     integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x))"
-    apply (rule has_integral_unique)
-    apply (subst add.commute)
-    apply (subst Henstock_Kurzweil_Integration.integral_combine)
-    using assms * integral_unique by auto
-
+    by (smt (verit, ccfv_threshold) Henstock_Kurzweil_Integration.integral_combine a add.commute atLeastAtMost_iff has_integral_iff)
   have vd1: "vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))"
     if "0 \<le> x" "x + a < 1" "x \<notin> (\<lambda>x. x - a) ` S" for x
     unfolding shiftpath_def
@@ -371,8 +367,7 @@
       then show "(g has_vector_derivative vector_derivative g (at (a + x))) (at (x + a))"
         by (metis add.commute vector_derivative_works)
     qed
-    then
-    show "((\<lambda>x. g (a + x)) has_vector_derivative vector_derivative g (at (x + a))) (at x)"
+    then show "((\<lambda>x. g (a + x)) has_vector_derivative vector_derivative g (at (x + a))) (at x)"
       by (auto simp: field_simps)
     show "0 < dist (1 - a) x"
       using that by auto
@@ -474,8 +469,8 @@
 
 lemma contour_integrable_on_shiftpath_eq:
   assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
-    shows "f contour_integrable_on (shiftpath a g) \<longleftrightarrow> f contour_integrable_on g"
-using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto
+  shows "f contour_integrable_on (shiftpath a g) \<longleftrightarrow> f contour_integrable_on g"
+  using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto
 
 lemma contour_integral_shiftpath:
   assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
@@ -556,26 +551,17 @@
 lemma contour_integrable_subpath:
   assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
     shows "f contour_integrable_on (subpath u v g)"
-proof (cases u v rule: linorder_class.le_cases)
-  case le
-  then show ?thesis
-    by (metis contour_integrable_on_def has_contour_integral_subpath [OF assms])
-next
-  case ge
-  with assms show ?thesis
-    by (metis (no_types, lifting) contour_integrable_on_def contour_integrable_reversepath_eq has_contour_integral_subpath reversepath_subpath valid_path_subpath)
-qed
+  by (smt (verit, ccfv_threshold) assms contour_integrable_on_def contour_integrable_reversepath_eq
+      has_contour_integral_subpath reversepath_subpath valid_path_subpath)
 
 lemma has_integral_contour_integral_subpath:
   assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
-    shows "(((\<lambda>x. f(g x) * vector_derivative g (at x)))
+    shows "((\<lambda>x. f(g x) * vector_derivative g (at x))
             has_integral  contour_integral (subpath u v g) f) {u..v}"
-  using assms
+          (is "(?fg has_integral _)_")
 proof -
-  have "(\<lambda>r. f (g r) * vector_derivative g (at r)) integrable_on {u..v}"
-    by (metis (full_types) assms(1) assms(3) assms(4) atLeastAtMost_iff atLeastatMost_subset_iff contour_integrable_on integrable_on_subinterval)
-  then have "((\<lambda>r. f (g r) * vector_derivative g (at r)) has_integral integral {u..v} (\<lambda>r. f (g r) * vector_derivative g (at r))) {u..v}"
-    by blast
+  have "(?fg has_integral integral {u..v} ?fg) {u..v}"
+    using assms contour_integrable_on integrable_on_subinterval by fastforce
   then show ?thesis
     by (metis (full_types) assms contour_integral_unique has_contour_integral_subpath)
 qed
@@ -653,9 +639,9 @@
 next
   fix x assume "x \<in> {0..1} - ({0, 1} \<union> g -` A \<inter> {0<..<1})"
   hence "g x \<in> path_image g - A" by (auto simp: path_image_def)
-  from assms(4)[OF this] and assms(3)
-    show "f' (g' x) * vector_derivative g' (at x) = f (g x) * vector_derivative g (at x)" by simp
-  qed
+  with assms show "f' (g' x) * vector_derivative g' (at x) = f (g x) * vector_derivative g (at x)" 
+    by simp
+qed
 
 
 text \<open>Contour integral along a segment on the real axis\<close>
@@ -664,7 +650,7 @@
   fixes a b :: complex and f :: "complex \<Rightarrow> complex"
   assumes "a \<in> Reals" "b \<in> Reals" "Re a < Re b"
   shows   "(f has_contour_integral I) (linepath a b) \<longleftrightarrow>
-             ((\<lambda>x. f (of_real x)) has_integral I) {Re a..Re b}"
+           ((\<lambda>x. f (of_real x)) has_integral I) {Re a..Re b}"
 proof -
   from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b"
     by (simp_all add: complex_eq_iff)
@@ -738,9 +724,8 @@
     have "((\<lambda>x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})"
       using diff_chain_within [OF gdiff fdiff]
       by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac)
-  } note * = this
-  show ?thesis
-    using assms cfg *
+  } then show ?thesis
+    using assms cfg 
     by (force simp: at_within_Icc_at intro: fundamental_theorem_of_calculus_interior_strong [OF \<open>finite K\<close>])
 qed
 
@@ -759,7 +744,7 @@
     shows "(f' has_contour_integral 0) g"
   using assms by (metis diff_self contour_integral_primitive)
 
-text\<open>Existence of path integral for continuous function\<close>
+
 lemma contour_integrable_continuous_linepath:
   assumes "continuous_on (closed_segment a b) f"
   shows "f contour_integrable_on (linepath a b)"
@@ -951,10 +936,9 @@
   by fastforce
 
 lemma contour_integrable_sum:
-    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
+  "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
      \<Longrightarrow> (\<lambda>x. sum (\<lambda>a. f a x) s) contour_integrable_on p"
-   unfolding contour_integrable_on_def
-   by (metis has_contour_integral_sum)
+  unfolding contour_integrable_on_def by (metis has_contour_integral_sum)
 
 lemma contour_integrable_neg_iff:
   "(\<lambda>x. -f x) contour_integrable_on g \<longleftrightarrow> f contour_integrable_on g"
@@ -1027,10 +1011,10 @@
       apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc)
       apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"])
       done
-  } note fj = this
-  show ?thesis
+  } 
+  then show ?thesis
     using f k unfolding has_contour_integral_linepath
-    by (simp add: linepath_def has_integral_combine [OF _ _ fi fj])
+    by (simp add: linepath_def has_integral_combine [OF _ _ fi])
 qed
 
 lemma continuous_on_closed_segment_transform:
@@ -1107,10 +1091,10 @@
   then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\<lambda>z. vector_derivative g (at (fst z)))"
        and  hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\<lambda>x. vector_derivative h (at (snd x)))"
     by auto
-  have "continuous_on (cbox (0, 0) (1, 1)) ((\<lambda>(y1, y2). f y1 y2) \<circ> (\<lambda>w. ((g \<circ> fst) w, (h \<circ> snd) w)))"
-    apply (intro gcon hcon continuous_intros | simp)+
-    apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon])
-    done
+  have "continuous_on ((\<lambda>x. (g (fst x), h (snd x))) ` cbox (0,0) (1,1)) (\<lambda>(y1, y2). f y1 y2)"
+    by (auto simp: path_image_def intro: continuous_on_subset [OF fcon])
+  then have "continuous_on (cbox (0, 0) (1, 1)) ((\<lambda>(y1, y2). f y1 y2) \<circ> (\<lambda>w. ((g \<circ> fst) w, (h \<circ> snd) w)))"
+    by (intro gcon hcon continuous_intros | simp)+
   then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\<lambda>x. f (g (fst x)) (h (snd x)))"
     by auto
   have "integral {0..1} (\<lambda>x. contour_integral h (f (g x)) * vector_derivative g (at x)) =
@@ -1186,7 +1170,7 @@
 lemma valid_path_polynomial_function:
   fixes p :: "real \<Rightarrow> 'a::euclidean_space"
   shows "polynomial_function p \<Longrightarrow> valid_path p"
-by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function)
+  by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function)
 
 lemma valid_path_subpath_trivial [simp]:
     fixes g :: "real \<Rightarrow> 'a::euclidean_space"
@@ -1199,15 +1183,15 @@
   where "part_circlepath z r s t \<equiv> \<lambda>x. z + of_real r * exp (\<i> * of_real (linepath s t x))"
 
 lemma pathstart_part_circlepath [simp]:
-     "pathstart(part_circlepath z r s t) = z + r*exp(\<i> * s)"
-by (metis part_circlepath_def pathstart_def pathstart_linepath)
+  "pathstart(part_circlepath z r s t) = z + r*exp(\<i> * s)"
+  by (metis part_circlepath_def pathstart_def pathstart_linepath)
 
 lemma pathfinish_part_circlepath [simp]:
-     "pathfinish(part_circlepath z r s t) = z + r*exp(\<i>*t)"
-by (metis part_circlepath_def pathfinish_def pathfinish_linepath)
+  "pathfinish(part_circlepath z r s t) = z + r*exp(\<i>*t)"
+  by (metis part_circlepath_def pathfinish_def pathfinish_linepath)
 
 lemma reversepath_part_circlepath[simp]:
-    "reversepath (part_circlepath z r s t) = part_circlepath z r t s"
+  "reversepath (part_circlepath z r s t) = part_circlepath z r t s"
   unfolding part_circlepath_def reversepath_def linepath_def
   by (auto simp:algebra_simps)
 
@@ -1284,24 +1268,12 @@
 lemma in_path_image_part_circlepath:
   assumes "w \<in> path_image(part_circlepath z r s t)" "s \<le> t" "0 \<le> r"
     shows "norm(w - z) = r"
-proof -
-  have "w \<in> {c. dist z c = r}"
-    by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms)
-  thus ?thesis
-    by (simp add: dist_norm norm_minus_commute)
-qed
+  by (smt (verit) assms dist_norm mem_Collect_eq norm_minus_commute path_image_part_circlepath_subset sphere_def subsetD)
 
 lemma path_image_part_circlepath_subset':
   assumes "r \<ge> 0"
   shows   "path_image (part_circlepath z r s t) \<subseteq> sphere z r"
-proof (cases "s \<le> t")
-  case True
-  thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp
-next
-  case False
-  thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms
-    by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all
-qed
+  by (smt (verit) assms path_image_part_circlepath_subset reversepath_part_circlepath reversepath_simps(2))
 
 lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x"
   by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps)
@@ -1458,7 +1430,7 @@
     have *: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1; part_circlepath z r s t x \<notin> k\<rbrakk> \<Longrightarrow> cmod (f (part_circlepath z r s t x)) \<le> B"
       by (auto intro!: B [unfolded path_image_def image_def])
     show ?thesis
-      apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified])
+      using has_integral_bound [where 'a=real, simplified, OF _ **]
       using assms le * "2" \<open>r > 0\<close> by (auto simp add: norm_mult vector_derivative_part_circlepath)
   qed
 qed
@@ -1520,13 +1492,16 @@
   qed
   have abs_away: "\<And>P. (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. P \<bar>x - y\<bar>) \<longleftrightarrow> (\<forall>x::real. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> P x)"
     by force
+  have "\<And>x n. \<lbrakk>s \<noteq> t; \<bar>s - t\<bar> \<le> 2 * pi; 0 \<le> x; x < 1;
+            x * (t - s) = 2 * (real_of_int n * pi)\<rbrakk>
+           \<Longrightarrow> x = 0"
+    by (rule ccontr) (auto simp: 2 field_split_simps abs_mult dest: of_int_leD)
+  then
   show ?thesis using False
     apply (simp add: simple_path_def loop_free_def)
     apply (simp add: part_circlepath_def linepath_def exp_eq  * ** abs01 del: Set.insert_iff)
     apply (subst abs_away)
     apply (auto simp: 1)
-    apply (rule ccontr)
-    apply (auto simp: 2 field_split_simps abs_mult dest: of_int_leD)
     done
 qed
 
@@ -1719,16 +1694,18 @@
                and inta: "(\<lambda>t. f a (\<gamma> t) * vector_derivative \<gamma> (at t)) integrable_on {0..1}"
       using eventually_happens [OF eventually_conj]
       by (fastforce simp: contour_integrable_on path_image_def)
-    have Ble: "B * e / (\<bar>B\<bar> + 1) \<le> e"
-      using \<open>0 \<le> B\<close>  \<open>0 < e\<close> by (simp add: field_split_simps)
     have "\<exists>h. (\<forall>x\<in>{0..1}. cmod (l (\<gamma> x) * vector_derivative \<gamma> (at x) - h x) \<le> e) \<and> h integrable_on {0..1}"
     proof (intro exI conjI ballI)
       show "cmod (l (\<gamma> x) * vector_derivative \<gamma> (at x) - f a (\<gamma> x) * vector_derivative \<gamma> (at x)) \<le> e"
         if "x \<in> {0..1}" for x
-        apply (rule order_trans [OF _ Ble])
-        using noleB [OF that] fga [OF that] \<open>0 \<le> B\<close> \<open>0 < e\<close>
-        apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le] simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps)
-        done
+      proof -
+        have "cmod (l (\<gamma> x) * vector_derivative \<gamma> (at x) - f a (\<gamma> x) * vector_derivative \<gamma> (at x)) \<le> B * e / (\<bar>B\<bar> + 1)"
+          using noleB [OF that] fga [OF that] \<open>0 \<le> B\<close> \<open>0 < e\<close>
+          by (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le] simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps)
+        also have "\<dots> \<le> e"
+          using \<open>0 \<le> B\<close>  \<open>0 < e\<close> by (simp add: field_split_simps)
+        finally show ?thesis .
+      qed
     qed (rule inta)
   }
   then show lintg: "l contour_integrable_on \<gamma>"
--- a/src/HOL/Complex_Analysis/Laurent_Convergence.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Complex_Analysis/Laurent_Convergence.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -7,17 +7,17 @@
 (* TODO: Move *)
 text \<open>TODO: Better than @{thm deriv_compose_linear}?\<close>
 lemma deriv_compose_linear':
-  assumes "f field_differentiable at (c * z+a)"
-  shows "deriv (\<lambda>w. f (c * w+a)) z = c * deriv f (c * z+a)"
-  apply (subst deriv_chain[where f="\<lambda>w. c * w+a",unfolded comp_def])
+  assumes "f field_differentiable at (c*z + a)"
+  shows "deriv (\<lambda>w. f (c*w + a)) z = c * deriv f (c*z + a)"
+  apply (subst deriv_chain[where f="\<lambda>w. c*w + a",unfolded comp_def])
   using assms by (auto intro:derivative_intros)
 
 text \<open>TODO: Better than @{thm higher_deriv_compose_linear}?\<close>
 lemma higher_deriv_compose_linear':
   fixes z::complex
   assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \<in> S"
-      and fg: "\<And>w. w \<in> S \<Longrightarrow> u * w+c \<in> T"
-    shows "(deriv ^^ n) (\<lambda>w. f (u * w+c)) z = u^n * (deriv ^^ n) f (u * z+c)"
+      and fg: "\<And>w. w \<in> S \<Longrightarrow> u*w + c \<in> T"
+    shows "(deriv ^^ n) (\<lambda>w. f (u*w + c)) z = u^n * (deriv ^^ n) f (u*z + c)"
 using z
 proof (induction n arbitrary: z)
   case 0 then show ?case by simp
@@ -116,14 +116,7 @@
 instance fls :: (semiring_char_0) semiring_char_0
 proof
   show "inj (of_nat :: nat \<Rightarrow> 'a fls)"
-  proof
-    fix m n :: nat
-    assume "of_nat m = (of_nat n :: 'a fls)"
-    hence "fls_nth (of_nat m) 0 = (fls_nth (of_nat n) 0 :: 'a)"
-      by (simp only: )
-    thus "m = n"
-      by (simp add: fls_of_nat_nth)
-  qed
+    by (metis fls_regpart_of_nat injI of_nat_eq_iff)
 qed
 
 lemma fls_const_eq_0_iff [simp]: "fls_const c = 0 \<longleftrightarrow> c = 0"
@@ -547,20 +540,14 @@
   assumes "g holomorphic_on A"
   assumes "g ` A \<subseteq> eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
   shows   "(\<lambda>x. eval_fls f (g x)) holomorphic_on A"
-proof -
-  have "eval_fls f \<circ> g holomorphic_on A"
-    by (intro holomorphic_on_compose[OF assms(1) holomorphic_on_eval_fls]) (use assms in auto)
-  thus ?thesis
-    by (simp add: o_def)
-qed
+  by (meson assms holomorphic_on_compose holomorphic_on_eval_fls holomorphic_transform o_def)
 
 lemma continuous_on_eval_fls:
   fixes f
   defines "n \<equiv> fls_subdegree f"
   assumes "A \<subseteq> eball 0 (fls_conv_radius f) - (if n \<ge> 0 then {} else {0})"
   shows   "continuous_on A (eval_fls f)"
-  by (intro holomorphic_on_imp_continuous_on holomorphic_on_eval_fls)
-     (use assms in auto)
+  using assms holomorphic_on_eval_fls holomorphic_on_imp_continuous_on by blast
 
 lemma continuous_on_eval_fls' [continuous_intros]:
   fixes f
@@ -568,9 +555,7 @@
   assumes "g ` A \<subseteq> eball 0 (fls_conv_radius f) - (if n \<ge> 0 then {} else {0})"
   assumes "continuous_on A g"
   shows   "continuous_on A (\<lambda>x. eval_fls f (g x))"
-  using assms(3)
-  by (intro continuous_on_compose2[OF continuous_on_eval_fls _ assms(2)])
-     (auto simp: n_def)
+  by (metis assms continuous_on_compose2 continuous_on_eval_fls order.refl)
 
 lemmas has_field_derivative_eval_fps' [derivative_intros] =
   DERIV_chain2[OF has_field_derivative_eval_fps]
@@ -621,7 +606,7 @@
 lemma eval_fls_deriv:
   assumes "z \<in> eball 0 (fls_conv_radius F) - {0}"
   shows   "eval_fls (fls_deriv F) z = deriv (eval_fls F) z"
-  by (rule sym, rule DERIV_imp_deriv, rule has_field_derivative_eval_fls, rule assms)
+  by (metis DERIV_imp_deriv assms has_field_derivative_eval_fls)
 
 lemma analytic_on_eval_fls:
   assumes "A \<subseteq> eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
@@ -781,8 +766,12 @@
 lemma is_pole_imp_neg_fls_subdegree:
   assumes F: "(\<lambda>x. f (z + x)) has_laurent_expansion F" and "is_pole f z"
   shows   "fls_subdegree F < 0"
-  apply (rule is_pole_0_imp_neg_fls_subdegree[OF F])
-  using assms(2) is_pole_shift_0 by blast
+proof -
+  have "is_pole (\<lambda>x. f (z + x)) 0"
+    using assms(2) is_pole_shift_0 by blast
+  then show ?thesis
+    using F is_pole_0_imp_neg_fls_subdegree by blast
+qed
 
 lemma is_pole_fls_subdegree_iff:
   assumes "(\<lambda>x. f (z + x)) has_laurent_expansion F"
@@ -2587,13 +2576,8 @@
     by auto
 qed
 
-lemma analytic_on_prod [analytic_intros]:
-  assumes "\<And>x. x \<in> A \<Longrightarrow> f x analytic_on B"
-  shows   "(\<lambda>z. \<Prod>x\<in>A. f x z) analytic_on B"
-  using assms by (induction A rule: infinite_finite_induct) (auto intro!: analytic_intros)
-
 lemma zorder_const [simp]: "c \<noteq> 0 \<Longrightarrow> zorder (\<lambda>_. c) z = 0"
-  by (intro zorder_eqI[where s = UNIV]) auto
+  by (intro zorder_eqI[where S = UNIV]) auto
 
 lemma zorder_prod_analytic:
   assumes "\<And>x. x \<in> A \<Longrightarrow> f x analytic_on {z}"
@@ -2613,12 +2597,7 @@
 lemma zorder_eq_0I:
   assumes "g analytic_on {z}" "g z \<noteq> 0"
   shows   "zorder g z = 0"
-proof -
-  from assms obtain r where r: "r > 0" "g holomorphic_on ball z r"
-    unfolding analytic_on_def by blast
-  thus ?thesis using assms
-    by (intro zorder_eqI[of "ball z r" _ g]) auto
-qed
+  using analytic_at assms zorder_eqI by fastforce
 
 lemma zorder_pos_iff:
   assumes "f holomorphic_on A" "open A" "z \<in> A" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
@@ -2651,12 +2630,7 @@
 lemma zorder_pos_iff':
   assumes "f analytic_on {z}" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
   shows   "zorder f z > 0 \<longleftrightarrow> f z = 0"
-proof -
-  from assms(1) obtain A where A: "open A" "{z} \<subseteq> A" "f holomorphic_on A"
-    unfolding analytic_on_holomorphic by auto
-  with zorder_pos_iff [OF A(3,1), of z] assms show ?thesis
-    by auto
-qed
+  using analytic_at assms zorder_pos_iff by blast
 
 lemma zorder_ge_0:
   assumes "f analytic_on {z}" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
@@ -2673,15 +2647,7 @@
 lemma zorder_eq_0_iff:
   assumes "f analytic_on {z}" "frequently (\<lambda>w. f w \<noteq> 0) (at z)"
   shows   "zorder f z = 0 \<longleftrightarrow> f z \<noteq> 0"
-proof
-  assume "f z \<noteq> 0"
-  thus "zorder f z = 0"
-    using assms zorder_eq_0I by blast
-next
-  assume "zorder f z = 0"
-  thus "f z \<noteq> 0"
-    using assms zorder_pos_iff' by fastforce
-qed
+  using assms zorder_eq_0I zorder_pos_iff' by fastforce
 
 lemma dist_mult_left:
   "dist (a * b) (a * c :: 'a :: real_normed_field) = norm a * dist b c"
@@ -2835,15 +2801,7 @@
 qed
 
 lemma fps_to_fls_eq_fls_const_iff [simp]: "fps_to_fls F = fls_const c \<longleftrightarrow> F = fps_const c"
-proof
-  assume "F = fps_const c"
-  thus "fps_to_fls F = fls_const c"
-    by simp
-next
-  assume "fps_to_fls F = fls_const c"
-  thus "F = fps_const c"
-    by (metis fls_regpart_const fls_regpart_fps_trivial)
-qed
+  using fps_to_fls_eq_iff by fastforce
 
 lemma zorder_compose':
   assumes "isolated_singularity_at f (g z)" "not_essential f (g z)"
--- a/src/HOL/Complex_Analysis/Residue_Theorem.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Complex_Analysis/Residue_Theorem.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -10,19 +10,16 @@
   fixes f :: "real \<Rightarrow> 'a :: real_normed_field"
   assumes "f \<in> O[at_bot](\<lambda>_. 1)"
   assumes "f \<in> O[at_top](\<lambda>_. 1)"
-  assumes "continuous_on UNIV f"
+  assumes cf: "continuous_on UNIV f"
   shows   "bounded (range f)"
 proof -
-  from assms(1) obtain c1 where "eventually (\<lambda>x. norm (f x) \<le> c1) at_bot"
-    by (auto elim!: landau_o.bigE)
-  then obtain x1 where x1: "\<And>x. x \<le> x1 \<Longrightarrow> norm (f x) \<le> c1"
-    by (auto simp: eventually_at_bot_linorder)
-  from assms(2) obtain c2 where "eventually (\<lambda>x. norm (f x) \<le> c2) at_top"
-    by (auto elim!: landau_o.bigE)
-  then obtain x2 where x2: "\<And>x. x \<ge> x2 \<Longrightarrow> norm (f x) \<le> c2"
-    by (auto simp: eventually_at_top_linorder)
+  obtain c1 c2 
+    where "eventually (\<lambda>x. norm (f x) \<le> c1) at_bot" "eventually (\<lambda>x. norm (f x) \<le> c2) at_top"
+    using assms by (auto elim!: landau_o.bigE)
+  then obtain x1 x2 where x1: "\<And>x. x \<le> x1 \<Longrightarrow> norm (f x) \<le> c1" and x2: "\<And>x. x \<ge> x2 \<Longrightarrow> norm (f x) \<le> c2"
+    by (auto simp: eventually_at_bot_linorder eventually_at_top_linorder)
   have "compact (f ` {x1..x2})"
-    by (intro compact_continuous_image continuous_on_subset[OF assms(3)]) auto
+    by (intro compact_continuous_image continuous_on_subset[OF cf]) auto
   hence "bounded (f ` {x1..x2})"
     by (rule compact_imp_bounded)
   then obtain c3 where c3: "\<And>x. x \<in> {x1..x2} \<Longrightarrow> norm (f x) \<le> c3"
@@ -67,7 +64,7 @@
     also have "f \<in> O[at z0'](\<lambda>_. 1)"
       using z0' by (intro insert.prems) auto
     finally show "g \<in> \<dots>" .
-  qed (insert insert.prems g, auto)
+  qed (use insert.prems g in auto)
   then obtain h where "h holomorphic_on S" "\<forall>z\<in>S - X. h z = g z" by blast
   with g have "h holomorphic_on S" "\<forall>z\<in>S - insert z0 X. h z = f z" by auto
   thus ?case by blast
@@ -96,96 +93,94 @@
 subsection \<open>Cauchy's residue theorem\<close>
 
 lemma get_integrable_path:
-  assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\<in>s-pts" "b\<in>s-pts"
+  assumes "open S" "connected (S-pts)" "finite pts" "f holomorphic_on (S-pts) " "a\<in>S-pts" "b\<in>S-pts"
   obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b"
-    "path_image g \<subseteq> s-pts" "f contour_integrable_on g" using assms
-proof (induct arbitrary:s thesis a rule:finite_induct[OF \<open>finite pts\<close>])
+    "path_image g \<subseteq> S-pts" "f contour_integrable_on g" using assms
+proof (induct arbitrary:S thesis a rule:finite_induct[OF \<open>finite pts\<close>])
   case 1
-  obtain g where "valid_path g" "path_image g \<subseteq> s" "pathstart g = a" "pathfinish g = b"
-    using connected_open_polynomial_connected[OF \<open>open s\<close>,of a b ] \<open>connected (s - {})\<close>
+  obtain g where "valid_path g" "path_image g \<subseteq> S" "pathstart g = a" "pathfinish g = b"
+    using connected_open_polynomial_connected[OF \<open>open S\<close>,of a b ] \<open>connected (S - {})\<close>
       valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto
   moreover have "f contour_integrable_on g"
-    using contour_integrable_holomorphic_simple[OF _ \<open>open s\<close> \<open>valid_path g\<close> \<open>path_image g \<subseteq> s\<close>,of f]
-      \<open>f holomorphic_on s - {}\<close>
+    using contour_integrable_holomorphic_simple[OF _ \<open>open S\<close> \<open>valid_path g\<close> \<open>path_image g \<subseteq> S\<close>,of f]
+      \<open>f holomorphic_on S - {}\<close>
     by auto
   ultimately show ?case using "1"(1)[of g] by auto
 next
   case idt:(2 p pts)
-  obtain e where "e>0" and e:"\<forall>w\<in>ball a e. w \<in> s \<and> (w \<noteq> a \<longrightarrow> w \<notin> insert p pts)"
-    using finite_ball_avoid[OF \<open>open s\<close> \<open>finite (insert p pts)\<close>, of a]
-      \<open>a \<in> s - insert p pts\<close>
+  obtain e where "e>0" and e:"\<forall>w\<in>ball a e. w \<in> S \<and> (w \<noteq> a \<longrightarrow> w \<notin> insert p pts)"
+    using finite_ball_avoid[OF \<open>open S\<close> \<open>finite (insert p pts)\<close>, of a]
+      \<open>a \<in> S - insert p pts\<close>
     by auto
   define a' where "a' \<equiv> a+e/2"
-  have "a'\<in>s-{p} -pts"  using e[rule_format,of "a+e/2"] \<open>e>0\<close>
+  have "a'\<in>S-{p} -pts"  using e[rule_format,of "a+e/2"] \<open>e>0\<close>
     by (auto simp add:dist_complex_def a'_def)
   then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b"
-    "path_image g' \<subseteq> s - {p} - pts" "f contour_integrable_on g'"
-    using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1)
+    "path_image g' \<subseteq> S - {p} - pts" "f contour_integrable_on g'"
+    using idt.hyps(3)[of a' "S-{p}"] idt.prems idt.hyps(1)
     by (metis Diff_insert2 open_delete)
   define g where "g \<equiv> linepath a a' +++ g'"
   have "valid_path g" unfolding g_def by (auto intro: valid_path_join)
   moreover have "pathstart g = a" and  "pathfinish g = b" unfolding g_def by auto
-  moreover have "path_image g \<subseteq> s - insert p pts" unfolding g_def
-    proof (rule subset_path_image_join)
-      have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
-        by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
-      then show "path_image (linepath a a') \<subseteq> s - insert p pts" using e idt(9)
-        by auto
-    next
-      show "path_image g' \<subseteq> s - insert p pts" using g'(4) by blast
-    qed
+  moreover have "path_image g \<subseteq> S - insert p pts" 
+    unfolding g_def
+  proof (rule subset_path_image_join)
+    have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
+      by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
+    then show "path_image (linepath a a') \<subseteq> S - insert p pts" using e idt(9)
+      by auto
+  next
+    show "path_image g' \<subseteq> S - insert p pts" using g'(4) by blast
+  qed
   moreover have "f contour_integrable_on g"
-    proof -
-      have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
-        by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
-      then have "continuous_on (closed_segment a a') f"
-        using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)]
-        apply (elim continuous_on_subset)
-        by auto
-      then have "f contour_integrable_on linepath a a'"
-        using contour_integrable_continuous_linepath by auto
-      then show ?thesis unfolding g_def
-        apply (rule contour_integrable_joinI)
-        by (auto simp add: \<open>e>0\<close>)
-    qed
+  proof -
+    have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
+      by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
+    then have "closed_segment a a' \<subseteq> S - insert p pts"
+      using e idt.prems(6) by auto
+    then have "continuous_on (closed_segment a a') f"
+      using holomorphic_on_imp_continuous_on holomorphic_on_subset idt.prems(5) by presburger
+    then show ?thesis 
+      using contour_integrable_continuous_linepath by (simp add: g_def)
+  qed
   ultimately show ?case using idt.prems(1)[of g] by auto
 qed
 
 lemma Cauchy_theorem_aux:
-  assumes "open s" "connected (s-pts)" "finite pts" "pts \<subseteq> s" "f holomorphic_on s-pts"
-          "valid_path g" "pathfinish g = pathstart g" "path_image g \<subseteq> s-pts"
-          "\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0"
-          "\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
+  assumes "open S" "connected (S-pts)" "finite pts" "pts \<subseteq> S" "f holomorphic_on S-pts"
+          "valid_path g" "pathfinish g = pathstart g" "path_image g \<subseteq> S-pts"
+          "\<forall>z. (z \<notin> S) \<longrightarrow> winding_number g z  = 0"
+          "\<forall>p\<in>S. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>S \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
   shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
     using assms
-proof (induct arbitrary:s g rule:finite_induct[OF \<open>finite pts\<close>])
+proof (induct arbitrary:S g rule:finite_induct[OF \<open>finite pts\<close>])
   case 1
   then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique)
 next
   case (2 p pts)
   note fin[simp] = \<open>finite (insert p pts)\<close>
-    and connected = \<open>connected (s - insert p pts)\<close>
+    and connected = \<open>connected (S - insert p pts)\<close>
     and valid[simp] = \<open>valid_path g\<close>
     and g_loop[simp] = \<open>pathfinish g = pathstart g\<close>
-    and holo[simp]= \<open>f holomorphic_on s - insert p pts\<close>
-    and path_img = \<open>path_image g \<subseteq> s - insert p pts\<close>
-    and winding = \<open>\<forall>z. z \<notin> s \<longrightarrow> winding_number g z = 0\<close>
-    and h = \<open>\<forall>pa\<in>s. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s \<and> (w \<noteq> pa \<longrightarrow> w \<notin> insert p pts))\<close>
-  have "h p>0" and "p\<in>s"
-    and h_p: "\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> insert p pts)"
-    using h \<open>insert p pts \<subseteq> s\<close> by auto
+    and holo[simp]= \<open>f holomorphic_on S - insert p pts\<close>
+    and path_img = \<open>path_image g \<subseteq> S - insert p pts\<close>
+    and winding = \<open>\<forall>z. z \<notin> S \<longrightarrow> winding_number g z = 0\<close>
+    and h = \<open>\<forall>pa\<in>S. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> S \<and> (w \<noteq> pa \<longrightarrow> w \<notin> insert p pts))\<close>
+  have "h p>0" and "p\<in>S"
+    and h_p: "\<forall>w\<in>cball p (h p). w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> insert p pts)"
+    using h \<open>insert p pts \<subseteq> S\<close> by auto
   obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p"
-      "path_image pg \<subseteq> s-insert p pts" "f contour_integrable_on pg"
-    proof -
-      have "p + h p\<in>cball p (h p)" using h[rule_format,of p]
-        by (simp add: \<open>p \<in> s\<close> dist_norm)
-      then have "p + h p \<in> s - insert p pts" using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close>
-        by fastforce
-      moreover have "pathstart g \<in> s - insert p pts " using path_img by auto
-      ultimately show ?thesis
-        using get_integrable_path[OF \<open>open s\<close> connected fin holo,of "pathstart g" "p+h p"] that
-        by blast
-    qed
+    "path_image pg \<subseteq> S-insert p pts" "f contour_integrable_on pg"
+  proof -
+    have "p + h p\<in>cball p (h p)" using h[rule_format,of p]
+      by (simp add: \<open>p \<in> S\<close> dist_norm)
+    then have "p + h p \<in> S - insert p pts" using h[rule_format,of p] \<open>insert p pts \<subseteq> S\<close>
+      by fastforce
+    moreover have "pathstart g \<in> S - insert p pts " using path_img by auto
+    ultimately show ?thesis
+      using get_integrable_path[OF \<open>open S\<close> connected fin holo,of "pathstart g" "p+h p"] that
+      by blast
+  qed
   obtain n::int where "n=winding_number g p"
     using integer_winding_number[OF _ g_loop,of p] valid path_img
     by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path)
@@ -193,12 +188,13 @@
   define p_circ_pt where "p_circ_pt \<equiv> linepath (p+h p) (p+h p)"
   define n_circ where "n_circ \<equiv> \<lambda>n. ((+++) p_circ ^^ n) p_circ_pt"
   define cp where "cp \<equiv> if n\<ge>0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))"
+
   have n_circ:"valid_path (n_circ k)"
       "winding_number (n_circ k) p = k"
       "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p"
       "path_image (n_circ k) =  (if k=0 then {p + h p} else sphere p (h p))"
       "p \<notin> path_image (n_circ k)"
-      "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number (n_circ k) p'=0 \<and> p'\<notin>path_image (n_circ k)"
+      "\<And>p'. p'\<notin>S - pts \<Longrightarrow> winding_number (n_circ k) p'=0 \<and> p'\<notin>path_image (n_circ k)"
       "f contour_integrable_on (n_circ k)"
       "contour_integral (n_circ k) f = k *  contour_integral p_circ f"
       for k
@@ -212,7 +208,7 @@
         and "p \<notin> path_image (n_circ 0)"
         unfolding n_circ_def p_circ_pt_def using \<open>h p > 0\<close>
         by (auto simp add: dist_norm)
-      show "winding_number (n_circ 0) p'=0 \<and> p'\<notin>path_image (n_circ 0)" when "p'\<notin>s- pts" for p'
+      show "winding_number (n_circ 0) p'=0 \<and> p'\<notin>path_image (n_circ 0)" when "p'\<notin>S- pts" for p'
         unfolding n_circ_def p_circ_pt_def
         apply (auto intro!:winding_number_trivial)
         by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+
@@ -226,7 +222,7 @@
       have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto
       have pcirc:"p \<notin> path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)"
         using Suc(3) unfolding p_circ_def using \<open>h p > 0\<close> by (auto simp add: p_circ_def)
-      have pcirc_image:"path_image p_circ \<subseteq> s - insert p pts"
+      have pcirc_image:"path_image p_circ \<subseteq> S - insert p pts"
         proof -
           have "path_image p_circ \<subseteq> cball p (h p)" using \<open>0 < h p\<close> p_circ_def by auto
           then show ?thesis using h_p pcirc(1) by auto
@@ -263,59 +259,59 @@
         by (simp add: n_circ_def p_circ_def)
       show "pathfinish (n_circ (Suc k)) = p + h p"
         using Suc(4) unfolding n_circ_def by auto
-      show "winding_number (n_circ (Suc k)) p'=0 \<and>  p'\<notin>path_image (n_circ (Suc k))" when "p'\<notin>s-pts" for p'
+      show "winding_number (n_circ (Suc k)) p'=0 \<and>  p'\<notin>path_image (n_circ (Suc k))" when "p'\<notin>S-pts" for p'
         proof -
-          have " p' \<notin> path_image p_circ" using \<open>p \<in> s\<close> h p_circ_def that using pcirc_image by blast
+          have " p' \<notin> path_image p_circ" using \<open>p \<in> S\<close> h p_circ_def that using pcirc_image by blast
           moreover have "p' \<notin> path_image (n_circ k)"
             using Suc.hyps(7) that by blast
           moreover have "winding_number p_circ p' = 0"
             proof -
               have "path_image p_circ \<subseteq> cball p (h p)"
-                using h unfolding p_circ_def using \<open>p \<in> s\<close> by fastforce
-              moreover have "p'\<notin>cball p (h p)" using \<open>p \<in> s\<close> h that "2.hyps"(2) by fastforce
-              ultimately show ?thesis unfolding p_circ_def
-                apply (intro winding_number_zero_outside)
-                by auto
+                using h unfolding p_circ_def using \<open>p \<in> S\<close> by fastforce
+              moreover have "p'\<notin>cball p (h p)" using \<open>p \<in> S\<close> h that "2.hyps"(2) by fastforce
+              ultimately show ?thesis 
+                unfolding p_circ_def
+                by (intro winding_number_zero_outside) auto
             qed
           ultimately show ?thesis
-            unfolding n_Suc
-            apply (subst winding_number_join)
-            by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that])
+            unfolding n_Suc using Suc.hyps pcirc
+            by (metis add.right_neutral not_in_path_image_join that valid_path_imp_path winding_number_join)
         qed
       show "f contour_integrable_on (n_circ (Suc k))"
         unfolding n_Suc
         by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)])
       show "contour_integral (n_circ (Suc k)) f = (Suc k) *  contour_integral p_circ f"
-        unfolding n_Suc
-        by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]
-          Suc(9) algebra_simps)
+        by (simp add: Rings.ring_distribs(2) Suc.hyps n_Suc pcirc pcirc_integrable)
     qed
   have cp[simp]:"pathstart cp = p + h p"  "pathfinish cp = p + h p"
-         "valid_path cp" "path_image cp \<subseteq> s - insert p pts"
+         "valid_path cp" "path_image cp \<subseteq> S - insert p pts"
          "winding_number cp p = - n"
-         "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number cp p'=0 \<and> p' \<notin> path_image cp"
+         "\<And>p'. p'\<notin>S - pts \<Longrightarrow> winding_number cp p'=0 \<and> p' \<notin> path_image cp"
          "f contour_integrable_on cp"
          "contour_integral cp f = - n * contour_integral p_circ f"
     proof -
       show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp"
         using n_circ unfolding cp_def by auto
     next
-      have "sphere p (h p) \<subseteq>  s - insert p pts"
-        using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> by force
-      moreover  have "p + complex_of_real (h p) \<in> s - insert p pts"
+      have "sphere p (h p) \<subseteq>  S - insert p pts"
+        using h[rule_format,of p] \<open>insert p pts \<subseteq> S\<close> by force
+      moreover  have "p + complex_of_real (h p) \<in> S - insert p pts"
         using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE)
-      ultimately show "path_image cp \<subseteq>  s - insert p pts" unfolding cp_def
+      ultimately show "path_image cp \<subseteq>  S - insert p pts" unfolding cp_def
         using n_circ(5)  by auto
     next
       show "winding_number cp p = - n"
         unfolding cp_def using winding_number_reversepath n_circ \<open>h p>0\<close>
         by (auto simp: valid_path_imp_path)
     next
-      show "winding_number cp p'=0 \<and> p' \<notin> path_image cp" when "p'\<notin>s - pts" for p'
-        unfolding cp_def
-        apply (auto)
-        apply (subst winding_number_reversepath)
-        by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1))
+      show "winding_number cp p'=0 \<and> p' \<notin> path_image cp" when "p'\<notin>S - pts" for p'
+      proof -
+        have "winding_number (reversepath (n_circ (nat n))) p' = 0"
+          using n_circ that
+          by (metis add.inverse_neutral valid_path_imp_path winding_number_reversepath)
+        then show ?thesis
+          using cp_def n_circ(7) that by force
+      qed
     next
       show "f contour_integrable_on cp" unfolding cp_def
         using contour_integrable_reversepath_eq n_circ(1,8) by auto
@@ -326,30 +322,31 @@
     qed
   define g' where "g' \<equiv> g +++ pg +++ cp +++ (reversepath pg)"
   have "contour_integral g' f = (\<Sum>p\<in>pts. winding_number g' p * contour_integral (circlepath p (h p)) f)"
-    proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \<open>finite pts\<close> ])
-      show "connected (s - {p} - pts)" using connected by (metis Diff_insert2)
-      show "open (s - {p})" using \<open>open s\<close> by auto
-      show " pts \<subseteq> s - {p}" using \<open>insert p pts \<subseteq> s\<close> \<open> p \<notin> pts\<close>  by blast
-      show "f holomorphic_on s - {p} - pts" using holo \<open>p \<notin> pts\<close> by (metis Diff_insert2)
+    proof (rule "2.hyps"(3)[of "S-{p}" "g'",OF _ _ \<open>finite pts\<close> ])
+      show "connected (S - {p} - pts)" using connected by (metis Diff_insert2)
+      show "open (S - {p})" using \<open>open S\<close> by auto
+      show " pts \<subseteq> S - {p}" using \<open>insert p pts \<subseteq> S\<close> \<open> p \<notin> pts\<close>  by blast
+      show "f holomorphic_on S - {p} - pts" using holo \<open>p \<notin> pts\<close> by (metis Diff_insert2)
       show "valid_path g'"
         unfolding g'_def cp_def using n_circ valid pg g_loop
-        by (auto intro!:valid_path_join )
+        by (auto intro!:valid_path_join)
       show "pathfinish g' = pathstart g'"
         unfolding g'_def cp_def using pg(2) by simp
-      show "path_image g' \<subseteq> s - {p} - pts"
+      show "path_image g' \<subseteq> S - {p} - pts"
         proof -
-          define s' where "s' \<equiv> s - {p} - pts"
-          have s':"s' = s-insert p pts " unfolding s'_def by auto
+          define s' where "s' \<equiv> S - {p} - pts"
+          have s':"s' = S-insert p pts " unfolding s'_def by auto
           then show ?thesis using path_img pg(4) cp(4)
-            unfolding g'_def
-            apply (fold s'_def s')
-            apply (intro subset_path_image_join)
-            by auto
+            by (simp add: g'_def s'_def subset_path_image_join)
         qed
       note path_join_imp[simp]
-      show "\<forall>z. z \<notin> s - {p} \<longrightarrow> winding_number g' z = 0"
+      show "\<forall>z. z \<notin> S - {p} \<longrightarrow> winding_number g' z = 0"
         proof clarify
-          fix z assume z:"z\<notin>s - {p}"
+          fix z assume z:"z\<notin>S - {p}"
+          have z_notin_cp: "z \<notin> path_image cp"
+            using cp(6) cp_def n_circ(6) z by auto
+          have z_notin_pg: "z \<notin> path_image pg"
+            by (metis Diff_iff Diff_insert2 pg(4) subsetD z)
           have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z
               + winding_number (pg +++ cp +++ (reversepath pg)) z"
             proof (rule winding_number_join)
@@ -358,13 +355,13 @@
               show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp
                 by (simp add: valid_path_imp_path)
             next
-              have "path_image (pg +++ cp +++ reversepath pg) \<subseteq> s - insert p pts"
+              have "path_image (pg +++ cp +++ reversepath pg) \<subseteq> S - insert p pts"
                 using pg(4) cp(4) by (auto simp:subset_path_image_join)
               then show "z \<notin> path_image (pg +++ cp +++ reversepath pg)" using z by auto
             next
               show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto
             qed
-          also have "... = winding_number g z + (winding_number pg z
+          also have "\<dots> = winding_number g z + (winding_number pg z
               + winding_number (cp +++ (reversepath pg)) z)"
             proof (subst add_left_cancel,rule winding_number_join)
               show "path pg" and "path (cp +++ reversepath pg)"
@@ -375,167 +372,129 @@
                 by (metis Diff_iff \<open>z \<notin> path_image pg\<close> contra_subsetD cp(4) insertI1
                   not_in_path_image_join path_image_reversepath singletonD)
             qed
-          also have "... = winding_number g z + (winding_number pg z
+          also have "\<dots> = winding_number g z + (winding_number pg z
               + (winding_number cp z + winding_number (reversepath pg) z))"
-            apply (auto intro!:winding_number_join simp: valid_path_imp_path)
-            apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z)
-            by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z)
-          also have "... = winding_number g z + winding_number cp z"
-            apply (subst winding_number_reversepath)
-            apply (auto simp: valid_path_imp_path)
-            by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z)
+            by (simp add: valid_path_imp_path winding_number_join z_notin_cp z_notin_pg)
+          also have "\<dots> = winding_number g z + winding_number cp z"
+            by (simp add: valid_path_imp_path winding_number_reversepath z_notin_pg)
           finally have "winding_number g' z = winding_number g z + winding_number cp z"
             unfolding g'_def .
           moreover have "winding_number g z + winding_number cp z = 0"
             using winding z \<open>n=winding_number g p\<close> by auto
           ultimately show "winding_number g' z = 0" unfolding g'_def by auto
         qed
-      show "\<forall>pa\<in>s - {p}. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s - {p} \<and> (w \<noteq> pa \<longrightarrow> w \<notin> pts))"
+      show "\<forall>pa \<in> S - {p}. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> S - {p} \<and> (w \<noteq> pa \<longrightarrow> w \<notin> pts))"
         using h by fastforce
     qed
   moreover have "contour_integral g' f = contour_integral g f
       - winding_number g p * contour_integral p_circ f"
-    proof -
-      have "contour_integral g' f =  contour_integral g f
-        + contour_integral (pg +++ cp +++ reversepath pg) f"
-        unfolding g'_def
-        apply (subst contour_integral_join)
-        by (auto simp add:open_Diff[OF \<open>open s\<close>,OF finite_imp_closed[OF fin]]
-          intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img]
-          contour_integrable_reversepath)
-      also have "... = contour_integral g f + contour_integral pg f
-          + contour_integral (cp +++ reversepath pg) f"
-        apply (subst contour_integral_join)
-        by (auto simp add:contour_integrable_reversepath)
-      also have "... = contour_integral g f + contour_integral pg f
+  proof -
+    have *: "f contour_integrable_on g" "f contour_integrable_on pg" "f contour_integrable_on cp"
+        by (auto simp add: open_Diff[OF \<open>open S\<close>,OF finite_imp_closed[OF fin]]
+                 intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img])
+      have "contour_integral g' f = contour_integral g f + contour_integral pg f
           + contour_integral cp f + contour_integral (reversepath pg) f"
-        apply (subst contour_integral_join)
-        by (auto simp add:contour_integrable_reversepath)
-      also have "... = contour_integral g f + contour_integral cp f"
+        using * by (simp add: g'_def contour_integrable_reversepath_eq)
+      also have "\<dots> = contour_integral g f + contour_integral cp f"
         using contour_integral_reversepath
         by (auto simp add:contour_integrable_reversepath)
-      also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f"
+      also have "\<dots> = contour_integral g f - winding_number g p * contour_integral p_circ f"
         using \<open>n=winding_number g p\<close> by auto
       finally show ?thesis .
     qed
   moreover have "winding_number g' p' = winding_number g p'" when "p'\<in>pts" for p'
     proof -
-      have [simp]: "p' \<notin> path_image g" "p' \<notin> path_image pg" "p'\<notin>path_image cp"
-        using "2.prems"(8) that
-        apply blast
-        apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that)
-        by (meson DiffD2 cp(4) rev_subsetD subset_insertI that)
-      have "winding_number g' p' = winding_number g p'
-          + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def
-        apply (subst winding_number_join)
-        apply (simp_all add: valid_path_imp_path)
-        apply (intro not_in_path_image_join)
-        by auto
-      also have "... = winding_number g p' + winding_number pg p'
+      obtain [simp]: "p' \<notin> path_image g" "p' \<notin> path_image pg" "p'\<notin>path_image cp"
+        using "2.prems"(8) that by (metis Diff_iff Diff_insert2 \<open>p' \<in> pts\<close> cp(4) pg(4) subsetD)
+      have "winding_number g' p' = winding_number g p' + winding_number pg p'
           + winding_number (cp +++ reversepath pg) p'"
-        apply (subst winding_number_join)
-        apply (simp_all add: valid_path_imp_path)
-        apply (intro not_in_path_image_join)
-        by auto
-      also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p'
-          + winding_number (reversepath pg) p'"
-        apply (subst winding_number_join)
-        by (simp_all add: valid_path_imp_path)
-      also have "... = winding_number g p' + winding_number cp p'"
-        apply (subst winding_number_reversepath)
-        by (simp_all add: valid_path_imp_path)
-      also have "... = winding_number g p'" using that by auto
+        by (simp add: g'_def not_in_path_image_join valid_path_imp_path winding_number_join)
+      also have "\<dots> = winding_number g p'" using that
+        by (simp add: valid_path_imp_path winding_number_join winding_number_reversepath)
       finally show ?thesis .
     qed
   ultimately show ?case unfolding p_circ_def
     apply (subst (asm) sum.cong[OF refl,
         of pts _ "\<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"])
-    by (auto simp add:sum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps)
+    by (auto simp: sum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps)
 qed
 
 lemma Cauchy_theorem_singularities:
-  assumes "open s" "connected s" "finite pts" and
-          holo:"f holomorphic_on s-pts" and
+  assumes "open S" "connected S" "finite pts" and
+          holo: "f holomorphic_on S-pts" and
           "valid_path g" and
           loop:"pathfinish g = pathstart g" and
-          "path_image g \<subseteq> s-pts" and
-          homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0" and
-          avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
+          "path_image g \<subseteq> S-pts" and
+          homo:"\<forall>z. (z \<notin> S) \<longrightarrow> winding_number g z  = 0" and
+          avoid:"\<forall>p\<in>S. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>S \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
   shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
     (is "?L=?R")
 proof -
   define circ where "circ \<equiv> \<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"
-  define pts1 where "pts1 \<equiv> pts \<inter> s"
+  define pts1 where "pts1 \<equiv> pts \<inter> S"
   define pts2 where "pts2 \<equiv> pts - pts1"
-  have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> s={}" "pts1\<subseteq>s"
+  have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> S={}" "pts1\<subseteq>S"
     unfolding pts1_def pts2_def by auto
   have "contour_integral g f =  (\<Sum>p\<in>pts1. circ p)" unfolding circ_def
-    proof (rule Cauchy_theorem_aux[OF \<open>open s\<close> _ _ \<open>pts1\<subseteq>s\<close> _ \<open>valid_path g\<close> loop _ homo])
+    proof (rule Cauchy_theorem_aux[OF \<open>open S\<close> _ _ \<open>pts1\<subseteq>S\<close> _ \<open>valid_path g\<close> loop _ homo])
       have "finite pts1" unfolding pts1_def using \<open>finite pts\<close> by auto
-      then show "connected (s - pts1)"
-        using \<open>open s\<close> \<open>connected s\<close> connected_open_delete_finite[of s] by auto
+      then show "connected (S - pts1)"
+        using \<open>open S\<close> \<open>connected S\<close> connected_open_delete_finite[of S] by auto
     next
       show "finite pts1" using \<open>pts = pts1 \<union> pts2\<close> assms(3) by auto
-      show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def)
-      show "path_image g \<subseteq> s - pts1" using assms(7) pts1_def by auto
-      show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts1))"
+      show "f holomorphic_on S - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def)
+      show "path_image g \<subseteq> S - pts1" using assms(7) pts1_def by auto
+      show "\<forall>p\<in>S. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts1))"
         by (simp add: avoid pts1_def)
     qed
-  moreover have "sum circ pts2=0"
-    proof -
-      have "winding_number g p=0" when "p\<in>pts2" for p
-        using  \<open>pts2 \<inter> s={}\<close> that homo[rule_format,of p] by auto
-      thus ?thesis unfolding circ_def
-        apply (intro sum.neutral)
-        by auto
-    qed
+  moreover have "sum circ pts2 = 0"
+    by (metis \<open>pts2 \<inter> S = {}\<close> circ_def disjoint_iff_not_equal homo mult_zero_left sum.neutral)
   moreover have "?R=sum circ pts1 + sum circ pts2"
     unfolding circ_def
     using sum.union_disjoint[OF _ _ \<open>pts1 \<inter> pts2 = {}\<close>] \<open>finite pts\<close> \<open>pts=pts1 \<union> pts2\<close>
     by blast
   ultimately show ?thesis
-    apply (fold circ_def)
-    by auto
+    by simp
 qed
 
 theorem Residue_theorem:
-  fixes s pts::"complex set" and f::"complex \<Rightarrow> complex"
+  fixes S pts::"complex set" and f::"complex \<Rightarrow> complex"
     and g::"real \<Rightarrow> complex"
-  assumes "open s" "connected s" "finite pts" and
-          holo:"f holomorphic_on s-pts" and
+  assumes "open S" "connected S" "finite pts" and
+          holo:"f holomorphic_on S-pts" and
           "valid_path g" and
           loop:"pathfinish g = pathstart g" and
-          "path_image g \<subseteq> s-pts" and
-          homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0"
+          "path_image g \<subseteq> S-pts" and
+          homo:"\<forall>z. (z \<notin> S) \<longrightarrow> winding_number g z  = 0"
   shows "contour_integral g f = 2 * pi * \<i> *(\<Sum>p\<in>pts. winding_number g p * residue f p)"
 proof -
   define c where "c \<equiv>  2 * pi * \<i>"
-  obtain h where avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
-    using finite_cball_avoid[OF \<open>open s\<close> \<open>finite pts\<close>] by metis
+  obtain h where avoid:"\<forall>p\<in>S. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>S \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
+    using finite_cball_avoid[OF \<open>open S\<close> \<open>finite pts\<close>] by metis
   have "contour_integral g f
       = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
     using Cauchy_theorem_singularities[OF assms avoid] .
-  also have "... = (\<Sum>p\<in>pts.  c * winding_number g p * residue f p)"
+  also have "\<dots> = (\<Sum>p\<in>pts.  c * winding_number g p * residue f p)"
     proof (intro sum.cong)
       show "pts = pts" by simp
     next
       fix x assume "x \<in> pts"
       show "winding_number g x * contour_integral (circlepath x (h x)) f
           = c * winding_number g x * residue f x"
-        proof (cases "x\<in>s")
+        proof (cases "x\<in>S")
           case False
           then have "winding_number g x=0" using homo by auto
           thus ?thesis by auto
         next
           case True
           have "contour_integral (circlepath x (h x)) f = c* residue f x"
-            using \<open>x\<in>pts\<close> \<open>finite pts\<close> avoid[rule_format,OF True]
-            apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def])
-            by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \<open>open s\<close> finite_imp_closed])
+            using \<open>x\<in>pts\<close> \<open>finite pts\<close> avoid[rule_format, OF True]
+            apply (intro base_residue[of "S-(pts-{x})",THEN contour_integral_unique,folded c_def])
+            by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \<open>open S\<close> finite_imp_closed])
           then show ?thesis by auto
         qed
     qed
-  also have "... = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)"
+  also have "\<dots> = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)"
     by (simp add: sum_distrib_left algebra_simps)
   finally show ?thesis unfolding c_def .
 qed
@@ -543,17 +502,17 @@
 subsection \<open>The argument principle\<close>
 
 theorem argument_principle:
-  fixes f::"complex \<Rightarrow> complex" and poles s:: "complex set"
-  defines "pz \<equiv> {w\<in>s. f w = 0 \<or> w \<in> poles}" \<comment> \<open>\<^term>\<open>pz\<close> is the set of poles and zeros\<close>
-  assumes "open s" "connected s" and
-          f_holo:"f holomorphic_on s-poles" and
-          h_holo:"h holomorphic_on s" and
+  fixes f::"complex \<Rightarrow> complex" and poles S:: "complex set"
+  defines "pz \<equiv> {w\<in>S. f w = 0 \<or> w \<in> poles}" \<comment> \<open>\<^term>\<open>pz\<close> is the set of poles and zeros\<close>
+  assumes "open S" "connected S" and
+          f_holo:"f holomorphic_on S-poles" and
+          h_holo:"h holomorphic_on S" and
           "valid_path g" and
           loop:"pathfinish g = pathstart g" and
-          path_img:"path_image g \<subseteq> s - pz" and
-          homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and
+          path_img:"path_image g \<subseteq> S - pz" and
+          homo:"\<forall>z. (z \<notin> S) \<longrightarrow> winding_number g z = 0" and
           finite:"finite pz" and
-          poles:"\<forall>p\<in>s\<inter>poles. is_pole f p"
+          poles:"\<forall>p\<in>S\<inter>poles. is_pole f p"
   shows "contour_integral g (\<lambda>x. deriv f x * h x / f x) = 2 * pi * \<i> *
           (\<Sum>p\<in>pz. winding_number g p * h p * zorder f p)"
     (is "?L=?R")
@@ -561,12 +520,12 @@
   define c where "c \<equiv> 2 * complex_of_real pi * \<i> "
   define ff where "ff \<equiv> (\<lambda>x. deriv f x * h x / f x)"
   define cont where "cont \<equiv> \<lambda>ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)"
-  define avoid where "avoid \<equiv> \<lambda>p e. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pz)"
+  define avoid where "avoid \<equiv> \<lambda>p e. \<forall>w\<in>cball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> pz)"
 
-  have "\<exists>e>0. avoid p e \<and> (p\<in>pz \<longrightarrow> cont ff p e)" when "p\<in>s" for p
+  have "\<exists>e>0. avoid p e \<and> (p\<in>pz \<longrightarrow> cont ff p e)" when "p\<in>S" for p
   proof -
     obtain e1 where "e1>0" and e1_avoid:"avoid p e1"
-      using finite_cball_avoid[OF \<open>open s\<close> finite] \<open>p\<in>s\<close> unfolding avoid_def by auto
+      using finite_cball_avoid[OF \<open>open S\<close> finite] \<open>p\<in>S\<close> unfolding avoid_def by auto
     have "\<exists>e2>0. cball p e2 \<subseteq> ball p e1 \<and> cont ff p e2" when "p\<in>pz"
     proof -
       define po where "po \<equiv> zorder f p"
@@ -580,9 +539,10 @@
       proof -
         have "isolated_singularity_at f p"
         proof -
-          have "f holomorphic_on ball p e1 - {p}"
-            apply (intro holomorphic_on_subset[OF f_holo])
-            using e1_avoid \<open>p\<in>pz\<close> unfolding avoid_def pz_def by force
+          have "ball p e1 - {p} \<subseteq> S - poles"
+            using avoid_def e1_avoid pz_def by fastforce
+          then have "f holomorphic_on ball p e1 - {p}"
+            by (intro holomorphic_on_subset[OF f_holo])
           then show ?thesis unfolding isolated_singularity_at_def
             using \<open>e1>0\<close> analytic_on_open open_delete by blast
         qed
@@ -592,13 +552,13 @@
           then show ?thesis unfolding not_essential_def by auto
         next
           case False
-          then have "p\<in>s-poles" using \<open>p\<in>s\<close> poles unfolding pz_def by auto
-          moreover have "open (s-poles)"
+          then have "p\<in>S-poles" using \<open>p\<in>S\<close> poles unfolding pz_def by auto
+          moreover have "open (S-poles)"
           proof -
-            have "closed (s \<inter> poles)"
+            have "closed (S \<inter> poles)"
               using finite by (simp add: pz_def finite_imp_closed rev_finite_subset subset_eq)
             then show ?thesis
-              by (metis Diff_Compl Diff_Diff_Int Diff_eq \<open>open s\<close> open_Diff) 
+              by (metis Diff_Compl Diff_Diff_Int Diff_eq \<open>open S\<close> open_Diff) 
           qed
           ultimately have "isCont f p"
             using holomorphic_on_imp_continuous_on[OF f_holo] continuous_on_eq_continuous_at
@@ -611,17 +571,17 @@
           then have "\<forall>\<^sub>F w in at p. f w= 0" unfolding frequently_def by auto
           then obtain r1 where "r1>0" and r1:"\<forall>w\<in>ball p r1 - {p}. f w =0"
             unfolding eventually_at by (auto simp add:dist_commute)
-          obtain r2 where "r2>0" and r2: "ball p r2 \<subseteq> s"
-            using \<open>p\<in>s\<close> \<open>open s\<close> openE by blast
+          obtain r2 where "r2>0" and r2: "ball p r2 \<subseteq> S"
+            using \<open>p\<in>S\<close> \<open>open S\<close> openE by blast
           define rr where "rr=min r1 r2"
 
           from r1 r2
-          have "ball p rr - {p} \<subseteq> {w\<in> s \<inter> ball p rr-{p}. f w=0}"
+          have "ball p rr - {p} \<subseteq> {w\<in> S \<inter> ball p rr-{p}. f w=0}"
             unfolding rr_def by auto
           moreover have "infinite (ball p rr - {p})" 
             using \<open>r1>0\<close> \<open>r2>0\<close> finite_imp_not_open 
             unfolding rr_def by fastforce
-          ultimately have "infinite {w\<in>s \<inter> ball p rr-{p}. f w=0}" using infinite_super by blast
+          ultimately have "infinite {w\<in>S \<inter> ball p rr-{p}. f w=0}" using infinite_super by blast
           then have "infinite pz"
             unfolding pz_def by (smt (verit) infinite_super Collect_mono_iff DiffE Int_iff)
           then show False using \<open>finite pz\<close> by auto
@@ -643,9 +603,9 @@
       define prin where "prin \<equiv> \<lambda>w. po * h w / (w - p)"
       have "((\<lambda>w.  prin w + anal w) has_contour_integral c * po * h p) (circlepath p e2)"
       proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified])
-        have "ball p r \<subseteq> s"
+        have "ball p r \<subseteq> S"
           using \<open>r<e1\<close> avoid_def ball_subset_cball e1_avoid by (simp add: subset_eq)
-        then have "cball p e2 \<subseteq> s"
+        then have "cball p e2 \<subseteq> S"
           using \<open>r>0\<close> unfolding e2_def by auto
         then have "(\<lambda>w. po * h w) holomorphic_on cball p e2"
           using h_holo by (auto intro!: holomorphic_intros)
@@ -653,7 +613,7 @@
           using Cauchy_integral_circlepath_simple[folded c_def, of "\<lambda>w. po * h w"] \<open>e2>0\<close>
           unfolding prin_def by (auto simp add: mult.assoc)
         have "anal holomorphic_on ball p r" unfolding anal_def
-          using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close> \<open>pp p\<noteq>0\<close>
+          using pp_holo h_holo pp_po \<open>ball p r \<subseteq> S\<close> \<open>pp p\<noteq>0\<close>
           by (auto intro!: holomorphic_intros)
         then show "(anal has_contour_integral 0) (circlepath p e2)"
           using e2_def \<open>r>0\<close>
@@ -677,11 +637,8 @@
             by (auto intro!: derivative_eq_intros DERIV_cong[OF has_field_derivative_powr_of_int])
         qed
         ultimately show "prin w + anal w = ff' w"
-          unfolding ff'_def prin_def anal_def
-          apply simp
-          apply (unfold f'_def)
-          apply (fold wp_def)
-          apply (auto simp add:field_simps)
+          unfolding f'_def ff'_def prin_def anal_def
+          apply (simp add: field_simps flip: wp_def)
           by (metis (no_types, lifting) mult.commute power_int_minus_mult)
       qed
       then have "cont ff p e2" unfolding cont_def
@@ -693,12 +650,11 @@
           show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo
             by (auto intro!: holomorphic_intros)
         next
-          have "ball p e1 - {p} \<subseteq> s - poles"
+          have "ball p e1 - {p} \<subseteq> S - poles"
             using ball_subset_cball e1_avoid[unfolded avoid_def] unfolding pz_def
             by auto
-          then have "ball p r - {p} \<subseteq> s - poles"
-            apply (elim dual_order.trans)
-            using \<open>r<e1\<close> by auto
+          then have "ball p r - {p} \<subseteq> S - poles"
+            using \<open>r<e1\<close> by force
           then show "f holomorphic_on ball p r - {p}" using f_holo
             by auto
         next
@@ -728,38 +684,38 @@
       by (auto simp add: e2 e4_def)
     ultimately show ?thesis by auto
   qed
-  then obtain get_e where get_e:"\<forall>p\<in>s. get_e p>0 \<and> avoid p (get_e p)
+  then obtain get_e where get_e:"\<forall>p\<in>S. get_e p>0 \<and> avoid p (get_e p)
       \<and> (p\<in>pz \<longrightarrow> cont ff p (get_e p))"
     by metis
   define ci where "ci \<equiv> \<lambda>p. contour_integral (circlepath p (get_e p)) ff"
   define w where "w \<equiv> \<lambda>p. winding_number g p"
   have "contour_integral g ff = (\<Sum>p\<in>pz. w p * ci p)" unfolding ci_def w_def
-  proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected s\<close> finite _ \<open>valid_path g\<close> loop
+  proof (rule Cauchy_theorem_singularities[OF \<open>open S\<close> \<open>connected S\<close> finite _ \<open>valid_path g\<close> loop
         path_img homo])
-    have "open (s - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto
-    then show "ff holomorphic_on s - pz" unfolding ff_def using f_holo h_holo
+    have "open (S - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open S\<close> by auto
+    then show "ff holomorphic_on S - pz" unfolding ff_def using f_holo h_holo
       by (auto intro!: holomorphic_intros simp add:pz_def)
   next
-    show "\<forall>p\<in>s. 0 < get_e p \<and> (\<forall>w\<in>cball p (get_e p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pz))"
+    show "\<forall>p\<in>S. 0 < get_e p \<and> (\<forall>w\<in>cball p (get_e p). w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> pz))"
       using get_e using avoid_def by blast
   qed
-  also have "... = (\<Sum>p\<in>pz. c * w p * h p * zorder f p)"
+  also have "\<dots> = (\<Sum>p\<in>pz. c * w p * h p * zorder f p)"
   proof (rule sum.cong[of pz pz,simplified])
     fix p assume "p \<in> pz"
     show "w p * ci p = c * w p * h p * (zorder f p)"
-    proof (cases "p\<in>s")
-      assume "p \<in> s"
-      have "ci p = c * h p * (zorder f p)" unfolding ci_def
-        apply (rule contour_integral_unique)
-        using get_e \<open>p\<in>s\<close> \<open>p\<in>pz\<close> unfolding cont_def by (metis mult.assoc mult.commute)
+    proof (cases "p\<in>S")
+      assume "p \<in> S"
+      have "ci p = c * h p * (zorder f p)" 
+        unfolding ci_def
+        using \<open>p \<in> S\<close> \<open>p \<in> pz\<close> cont_def contour_integral_unique get_e by fastforce
       thus ?thesis by auto
     next
-      assume "p\<notin>s"
+      assume "p\<notin>S"
       then have "w p=0" using homo unfolding w_def by auto
       then show ?thesis by auto
     qed
   qed
-  also have "... = c*(\<Sum>p\<in>pz. w p * h p * zorder f p)"
+  also have "\<dots> = c*(\<Sum>p\<in>pz. w p * h p * zorder f p)"
     unfolding sum_distrib_left by (simp add:algebra_simps)
   finally have "contour_integral g ff = c * (\<Sum>p\<in>pz. w p * h p * of_int (zorder f p))" .
   then show ?thesis unfolding ff_def c_def w_def by simp
@@ -945,7 +901,8 @@
       proof -
         have "cmod (g p/f p) <1"
           by (smt (verit) divide_less_eq_1_pos norm_divide norm_ge_zero path_less that)
-        then show ?thesis unfolding h_def by (auto simp add:dist_complex_def)
+        then show ?thesis 
+          unfolding h_def by (auto simp add:dist_complex_def)
       qed
       then have "path_image (h o \<gamma>) \<subseteq> ball 1 1"
         by (simp add: image_subset_iff path_image_compose)
@@ -999,8 +956,7 @@
     then have " ((/) 1 has_contour_integral 0) (h \<circ> \<gamma>)
                   = ((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>"
       unfolding has_contour_integral
-      apply (intro has_integral_spike_eq[OF negligible_finite, OF \<open>finite spikes\<close>])
-      by auto
+      by (force intro!: has_integral_spike_eq[OF negligible_finite, OF \<open>finite spikes\<close>])
     ultimately show ?thesis by auto
   qed
   then have "contour_integral \<gamma> (\<lambda>x. deriv h x / h x) = 0"
@@ -1010,8 +966,8 @@
   proof -
     have "(\<lambda>p. deriv f p / f p) contour_integrable_on \<gamma>"
     proof (rule contour_integrable_holomorphic_simple[OF _ _ \<open>valid_path \<gamma>\<close> path_f])
-      show "open (s - zeros_f)" using finite_imp_closed[OF \<open>finite zeros_f\<close>] \<open>open s\<close>
-        by auto
+      show "open (s - zeros_f)" 
+        using finite_imp_closed[OF \<open>finite zeros_f\<close>] \<open>open s\<close> by auto
       then show "(\<lambda>p. deriv f p / f p) holomorphic_on s - zeros_f"
         using f_holo
         by (auto intro!: holomorphic_intros simp add:zeros_f_def)
@@ -1026,16 +982,15 @@
     moreover have "deriv fg p / fg p =  deriv f p / f p + deriv h p / h p"
                       when "p\<in> path_image \<gamma>" for p
     proof -
-      have "fg p\<noteq>0" and "f p\<noteq>0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def
-        by auto
-      have "h p\<noteq>0"
+      have "fg p \<noteq> 0" and "f p \<noteq> 0" 
+          using path_f path_fg that unfolding zeros_f_def zeros_fg_def by auto
+      have "h p \<noteq> 0"
       proof (rule ccontr)
         assume "\<not> h p \<noteq> 0"
-        then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2)
-        then have "cmod (g p/f p) = 1" by auto
-        moreover have "cmod (g p/f p) <1"
-          by (simp add: \<open>f p \<noteq> 0\<close> norm_divide path_less that)
-        ultimately show False by auto
+        then have "cmod (g p/f p) = 1"
+          by (simp add: add_eq_0_iff2 h_def)
+        then show False
+          by (smt (verit) divide_eq_1_iff norm_divide path_less that)
       qed
       have der_fg:"deriv fg p =  deriv f p + deriv g p" unfolding fg_def
         using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _  \<open>open s\<close>] path_img that
@@ -1050,9 +1005,10 @@
         then show ?thesis unfolding der_def using DERIV_imp_deriv by auto
       qed
       show ?thesis
-        apply (simp only:der_fg der_h)
-        apply (auto simp add:field_simps \<open>h p\<noteq>0\<close> \<open>f p\<noteq>0\<close> \<open>fg p\<noteq>0\<close>)
-        by (auto simp add:field_simps h_def \<open>f p\<noteq>0\<close> fg_def)
+        using \<open>h p\<noteq>0\<close> \<open>f p\<noteq>0\<close> \<open>fg p\<noteq>0\<close>
+       unfolding der_fg der_h
+        apply (simp add: divide_simps h_def fg_def)
+       by (simp add: mult.commute mult.left_commute ring_class.ring_distribs(1))
     qed
     then have "contour_integral \<gamma> (\<lambda>p. deriv fg p / fg p)
                   = contour_integral \<gamma> (\<lambda>p. deriv f p / f p + deriv h p / h p)"
@@ -1061,11 +1017,14 @@
   qed
   moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = c * (\<Sum>p\<in>zeros_fg. w p * zorder fg p)"
   proof -
-    have "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto
+    have "fg holomorphic_on s" 
+      unfolding fg_def using f_holo g_holo holomorphic_on_add by auto
     moreover
-    have "path_image \<gamma> \<subseteq> s - {p\<in>s. fg p = 0}" using path_fg unfolding zeros_fg_def .
+    have "path_image \<gamma> \<subseteq> s - {p\<in>s. fg p = 0}" 
+      using path_fg unfolding zeros_fg_def .
     moreover
-    have " finite {p\<in>s. fg p = 0}" using \<open>finite zeros_fg\<close> unfolding zeros_fg_def .
+    have " finite {p\<in>s. fg p = 0}" 
+      using \<open>finite zeros_fg\<close> unfolding zeros_fg_def .
     ultimately show ?thesis
       unfolding c_def zeros_fg_def w_def
       using argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo, of _ "{}" "\<lambda>_. 1"]
@@ -1075,9 +1034,12 @@
     unfolding c_def zeros_f_def w_def
   proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo
         , of _ "{}" "\<lambda>_. 1",simplified])
-    show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto
-    show "path_image \<gamma> \<subseteq> s - {p\<in>s. f p = 0}" using path_f unfolding zeros_f_def .
-    show " finite {p\<in>s. f p = 0}" using \<open>finite zeros_f\<close> unfolding zeros_f_def .
+    show "f holomorphic_on s" 
+      using f_holo g_holo holomorphic_on_add by auto
+    show "path_image \<gamma> \<subseteq> s - {p\<in>s. f p = 0}" 
+      using path_f unfolding zeros_f_def .
+    show " finite {p\<in>s. f p = 0}" 
+      using \<open>finite zeros_f\<close> unfolding zeros_f_def .
   qed
   ultimately have " c* (\<Sum>p\<in>zeros_fg. w p * (zorder fg p)) = c* (\<Sum>p\<in>zeros_f. w p * (zorder f p))"
     by auto
--- a/src/HOL/Complex_Analysis/Riemann_Mapping.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Complex_Analysis/Riemann_Mapping.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -62,10 +62,9 @@
     then show ?thesis by auto
   qed
   show ?thesis
-  apply (simp add: Moebius_function_def)
-  apply (intro holomorphic_intros)
-  using assms *
-  by (metis complex_cnj_cnj complex_cnj_mult complex_cnj_one complex_mod_cnj mem_ball_0 mult.commute right_minus_eq)
+    unfolding Moebius_function_def
+    apply (intro holomorphic_intros)
+    by (metis "*" mult.commute complex_cnj_cnj complex_cnj_mult complex_cnj_one complex_mod_cnj mem_ball_0 right_minus_eq)
 qed
 
 lemma Moebius_function_compose:
@@ -154,9 +153,7 @@
           if "z \<in> ball 0 1" for z::complex
         proof (rule DERIV_chain' [where g=f])
           show "(f has_field_derivative deriv f (of_real r * z)) (at (of_real r * z))"
-            apply (rule holomorphic_derivI [OF holF \<open>open S\<close>])
-             apply (rule \<open>f \<in> F\<close>)
-            by (meson imageI r01 subset_iff that)
+            by (metis holomorphic_derivI [OF holF \<open>open S\<close>] \<open>f \<in> F\<close> image_subset_iff r01 that)
         qed simp
         have df0: "((\<lambda>w. f (r * w)) has_field_derivative deriv f 0 * r) (at 0)"
           using * [of 0] by simp
@@ -170,16 +167,18 @@
     qed
     define l where "l \<equiv> SUP h\<in>F. norm (deriv h 0)"
     have eql: "norm (deriv f 0) = l" if le: "l \<le> norm (deriv f 0)" and "f \<in> F" for f
-      apply (rule order_antisym [OF _ le])
-      using \<open>f \<in> F\<close> bdd cSUP_upper by (fastforce simp: l_def)
+    proof (rule order_antisym [OF _ le])
+      show "cmod (deriv f 0) \<le> l"
+        using \<open>f \<in> F\<close> bdd cSUP_upper by (fastforce simp: l_def)
+    qed
     obtain \<F> where \<F>in: "\<And>n. \<F> n \<in> F" and \<F>lim: "(\<lambda>n. norm (deriv (\<F> n) 0)) \<longlonglongrightarrow> l"
     proof -
       have "\<exists>f. f \<in> F \<and> \<bar>norm (deriv f 0) - l\<bar> < 1 / (Suc n)" for n
       proof -
         obtain f where "f \<in> F" and f: "l < norm (deriv f 0) + 1/(Suc n)"
-          using cSup_least [OF imF_ne, of "l - 1/(Suc n)"] by (fastforce simp add: l_def)
+          using cSup_least [OF imF_ne, of "l - 1/(Suc n)"] by (fastforce simp: l_def)
         then have "\<bar>norm (deriv f 0) - l\<bar> < 1 / (Suc n)"
-          by (fastforce simp add: abs_if not_less eql)
+          by (fastforce simp: abs_if not_less eql)
         with \<open>f \<in> F\<close> show ?thesis
           by blast
       qed
@@ -197,7 +196,7 @@
           fix n assume "N \<le> n"
           have "dist (norm (deriv (\<F> n) 0)) l < 1 / (Suc n)"
             using fless by (simp add: dist_norm)
-          also have "... < e"
+          also have "\<dots> < e"
             using N \<open>N \<le> n\<close> inverse_of_nat_le le_less_trans by blast
           finally show "dist (norm (deriv (\<F> n) 0)) l < e" .
         qed
@@ -230,18 +229,8 @@
     with LIMSEQ_subseq_LIMSEQ [OF \<F>lim r] have no_df0: "norm(deriv f 0) = l"
       by (force simp: o_def intro: tendsto_unique)
     have nonconstf: "\<not> f constant_on S"
-    proof -
-      have False if "\<And>x. x \<in> S \<Longrightarrow> f x = c" for c
-      proof -
-        have "deriv f 0 = 0"
-          by (metis that \<open>open S\<close> \<open>0 \<in> S\<close> DERIV_imp_deriv [OF has_field_derivative_transform_within_open [OF DERIV_const]])
-        with no_df0 have "l = 0"
-          by auto
-        with eql [OF _ idF] show False by auto
-      qed
-      then show ?thesis
-        by (meson constant_on_def)
-    qed
+      using \<open>open S\<close> \<open>0 \<in> S\<close> no_df0 holomorphic_nonconstant [OF holf] eql [OF _ idF]
+      by force
     show ?thesis
     proof
       show "f \<in> F"
@@ -316,26 +305,14 @@
           show "(\<theta> \<circ> g \<circ> k) holomorphic_on (h \<circ> f) ` S"
           proof (intro holomorphic_on_compose)
             show "k holomorphic_on (h \<circ> f) ` S"
-              apply (rule holomorphic_on_subset [OF holk])
-              using f01 h01 by force
+              using holomorphic_on_subset [OF holk] f01 h01 by force
             show "g holomorphic_on k ` (h \<circ> f) ` S"
-              apply (rule holomorphic_on_subset [OF holg])
-              by (auto simp: kh nf1)
+              using holomorphic_on_subset [OF holg] by (force simp: kh nf1)
             show "\<theta> holomorphic_on g ` k ` (h \<circ> f) ` S"
-              apply (rule holomorphic_on_subset [OF hol\<theta>])
-              by (auto simp: gf kh nf1)
+              using holomorphic_on_subset [OF hol\<theta>] by (force simp: gf kh nf1)
           qed
           show "((\<theta> \<circ> g \<circ> k) (h (f z)))\<^sup>2 = h (f z)" if "z \<in> S" for z
-          proof -
-            have "f z \<in> ball 0 1"
-              by (simp add: nf1 that)
-            then have "(\<theta> (g (k (h (f z)))))\<^sup>2 = (\<theta> (g (f z)))\<^sup>2"
-              by (metis kh)
-            also have "... = h (f z)"
-              using \<theta>2 gf that by auto
-            finally show ?thesis
-              by (simp add: o_def)
-          qed
+            using \<theta>2 gf kh nf1 that by fastforce
         qed
       qed
       have norm\<psi>1: "norm(\<psi> (h (f z))) < 1" if "z \<in> S" for z
@@ -359,18 +336,14 @@
         show "p \<circ> \<psi> \<circ> h \<circ> f holomorphic_on S"
         proof (intro holomorphic_on_compose holf)
           show "h holomorphic_on f ` S"
-            apply (rule holomorphic_on_subset [OF holh])
-            using f01 by force
+            using holomorphic_on_subset [OF holh] f01 by fastforce
           show "\<psi> holomorphic_on h ` f ` S"
-            apply (rule holomorphic_on_subset [OF hol\<psi>])
-            by auto
+            using holomorphic_on_subset [OF hol\<psi>] by fastforce
           show "p holomorphic_on \<psi> ` h ` f ` S"
-            apply (rule holomorphic_on_subset [OF holp])
-            by (auto simp: norm\<psi>1)
+            using holomorphic_on_subset [OF holp] by (simp add: image_subset_iff norm\<psi>1)
         qed
         show "(p \<circ> \<psi> \<circ> h \<circ> f) ` S \<subseteq> ball 0 1"
-          apply clarsimp
-          by (meson norm\<psi>1 p01 image_subset_iff mem_ball_0)
+          using norm\<psi>1 p01 by fastforce
         show "(p \<circ> \<psi> \<circ> h \<circ> f) 0 = 0"
           by (simp add: \<open>p (\<psi> (h (f 0))) = 0\<close>)
         show "inj_on (p \<circ> \<psi> \<circ> h \<circ> f) S"
@@ -385,8 +358,8 @@
           using holomorphic_on_subset holomorphic_on_power
           by (blast intro: holomorphic_on_ident)
         show "k holomorphic_on power2 ` q ` ball 0 1"
-          apply (rule holomorphic_on_subset [OF holk])
-          using q01 by (auto simp: norm_power abs_square_less_1)
+          using q01  holomorphic_on_subset [OF holk] 
+          by (force simp: norm_power abs_square_less_1)
       qed
       have 2: "(k \<circ> power2 \<circ> q) 0 = 0"
         using p0 F_def \<open>f \<in> F\<close> \<psi>01 \<psi>2 \<open>0 \<in> S\<close> kh qp by force
@@ -442,12 +415,12 @@
     then show "a \<in> f ` S"
       by blast
   qed
-  then have "f ` S = ball 0 1"
+  then have fS: "f ` S = ball 0 1"
     using F_def \<open>f \<in> F\<close> by blast
-  then show ?thesis
-    apply (rule_tac x=f in exI)
-    apply (rule_tac x=g in exI)
-    using holf holg derg gf by safe force+
+  then have "\<forall>z\<in>ball 0 1. g z \<in> S \<and> f (g z) = z"
+    by (metis  gf imageE)
+  with fS show ?thesis
+    by (metis gf holf holg image_eqI)
 qed
 
 
@@ -526,7 +499,7 @@
           qed (use \<open>e > 0\<close> in auto)
           with \<open>e > 0\<close> have "inverse (norm (y - x)) * norm (z - f x * (y - x)) \<le> e/2"
             by (simp add: field_split_simps)
-          also have "... < e"
+          also have "\<dots> < e"
             using \<open>e > 0\<close> by simp
           finally show ?thesis
             by (simp add: contour_integral_unique [OF z])
@@ -593,11 +566,10 @@
     by (simp add: openS holf holomorphic_deriv holomorphic_on_divide nz)
   then obtain g where g: "\<And>z. z \<in> S \<Longrightarrow> (g has_field_derivative deriv f z / f z) (at z)"
     using prev [of "\<lambda>z. deriv f z / f z"] by metis
+  have Df: "\<And>x. x \<in> S \<Longrightarrow> DERIV f x :> deriv f x"
+    using holf holomorphic_derivI openS by force
   have hfd: "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>z. exp (g z) / f z) has_field_derivative 0) (at x)"
-    apply (rule derivative_eq_intros g| simp)+
-      apply (subst DERIV_deriv_iff_field_differentiable)
-    using openS holf holomorphic_on_imp_differentiable_at nz apply auto
-    done
+    by (rule derivative_eq_intros Df g nz| simp)+
   obtain c where c: "\<And>x. x \<in> S \<Longrightarrow> exp (g x) / f x = c"
   proof (rule DERIV_zero_connected_constant[OF \<open>connected S\<close> openS finite.emptyI])
     show "continuous_on S (\<lambda>z. exp (g z) / f z)"
@@ -607,9 +579,10 @@
   qed auto
   show ?thesis
   proof (intro exI ballI conjI)
-    show "(\<lambda>z. Ln(inverse c) + g z) holomorphic_on S"
-      apply (intro holomorphic_intros)
+    have "g holomorphic_on S"
       using openS g holomorphic_on_open by blast
+    then show "(\<lambda>z. Ln(inverse c) + g z) holomorphic_on S"
+      by (intro holomorphic_intros)
     fix z :: complex
     assume "z \<in> S"
     then have "exp (g z) / c = f z"
@@ -675,18 +648,14 @@
       by blast
     then obtain w where w: "- g z = g w" "dist a w < \<delta>"
       by auto
-    then have "w \<in> ball a \<delta>"
-      by simp
-    then have "w \<in> S"
-      using \<delta> by blast
+    with \<delta> have "w \<in> S"
+      by force
     then have "w = z"
       by (metis diff_add_cancel eqg power_minus_Bit0 that w(1))
     then have "g z = 0"
       using \<open>- g z = g w\<close> by auto
-    with eqg [OF that] have "z = b"
-      by auto
-    with that \<open>b \<notin> S\<close> show False
-      by simp
+    with eqg that \<open>b \<notin> S\<close> show False
+      by force
   qed
   then have nz: "\<And>z. z \<in> S \<Longrightarrow> g z + g a \<noteq> 0"
     by (metis \<open>0 < r\<close> add.commute add_diff_cancel_left' centre_in_ball diff_0)
@@ -723,23 +692,11 @@
     if holf: "f holomorphic_on h ` S" and nz: "\<And>z. z \<in> h ` S \<Longrightarrow> f z \<noteq> 0" "inj_on f (h ` S)" for f
   proof -
     obtain g where holg: "g holomorphic_on S" and eqg: "\<And>z. z \<in> S \<Longrightarrow> (f \<circ> h) z = (g z)\<^sup>2"
-    proof -
-      have "f \<circ> h holomorphic_on S"
-        by (simp add: holh holomorphic_on_compose holf)
-      moreover have "\<forall>z\<in>S. (f \<circ> h) z \<noteq> 0"
-        by (simp add: nz)
-      ultimately show thesis
-        using prev that by blast
-    qed
+      by (smt (verit) comp_def holf holh holomorphic_on_compose image_eqI nz(1) prev)
     show ?thesis
     proof (intro exI conjI)
       show "g \<circ> k holomorphic_on h ` S"
-      proof -
-        have "k ` h ` S \<subseteq> S"
-          by (simp add: \<open>\<And>z. z \<in> S \<Longrightarrow> k (h z) = z\<close> image_subset_iff)
-        then show ?thesis
-          by (meson holg holk holomorphic_on_compose holomorphic_on_subset)
-      qed
+        by (smt (verit) holg holk holomorphic_on_compose holomorphic_on_subset imageE image_subset_iff kh)
       show "\<forall>z\<in>h ` S. f z = ((g \<circ> k) z)\<^sup>2"
         using eqg kh by auto
     qed
@@ -845,10 +802,15 @@
 
 corollary contractible_eq_simply_connected_2d:
   fixes S :: "complex set"
-  shows "open S \<Longrightarrow> (contractible S \<longleftrightarrow> simply_connected S)"
-  apply safe
-   apply (simp add: contractible_imp_simply_connected)
-  using convex_imp_contractible homeomorphic_contractible_eq simply_connected_eq_homeomorphic_to_disc by auto
+  assumes "open S"
+  shows "contractible S \<longleftrightarrow> simply_connected S"
+proof
+  show "contractible S \<Longrightarrow> simply_connected S"
+    by (simp add: contractible_imp_simply_connected)
+  show "simply_connected S \<Longrightarrow> contractible S"
+    using assms convex_imp_contractible homeomorphic_contractible_eq 
+      simply_connected_eq_homeomorphic_to_disc by auto
+qed
 
 subsection\<open>A further chain of equivalences about components of the complement of a simply connected set\<close>
 
@@ -867,9 +829,7 @@
   then show ?thesis
   proof
     assume "S = {}"
-    then have "bounded S"
-      by simp
-    with \<open>S = {}\<close> show ?thesis
+    then show ?thesis
       by simp
   next
     assume S01: "S homeomorphic ball (0::complex) 1"
@@ -889,12 +849,11 @@
       by (simp add: X_def)
     have Xsubclo: "X n \<subseteq> closure S" for n
       unfolding X_def by (metis A01 closure_mono fim image_mono)
-    have connX: "connected(X n)" for n
+    have "connected (A n)" for n
+      using connected_annulus [of _ "0::complex"] by (simp add: A_def)
+    then have connX: "connected(X n)" for n
       unfolding X_def
-      apply (rule connected_imp_connected_closure)
-      apply (rule connected_continuous_image)
-      apply (simp add: continuous_on_subset [OF contf A01])
-      using connected_annulus [of _ "0::complex"] by (simp add: A_def)
+      by (metis A01 connected_continuous_image connected_imp_connected_closure contf continuous_on_subset)
     have nestX: "X n \<subseteq> X m" if "m \<le> n" for m n
     proof -
       have "1 - 1 / (real m + 2) \<le> 1 - 1 / (real n + 2)"
@@ -923,22 +882,13 @@
         have *: "(f ` cball 0 (1 - 1 / (real n + 2))) \<subseteq> S"
           by (force simp: D_def Seq)
         show "x \<in> X n"
-          using \<open>x \<in> closure S\<close> unfolding X_def Seq
-          using \<open>x \<notin> S\<close> * D_def clo_fim by auto
+          using Seq X_def \<open>x \<in> closure S\<close> \<open>x \<notin> S\<close> clo_fim by fastforce
       qed
     qed
     moreover have "(\<Inter>n. X n) \<subseteq> closure S - S"
     proof -
       have "(\<Inter>n. X n) \<subseteq> closure S"
-      proof -
-        have "(\<Inter>n. X n) \<subseteq> X 0"
-          by blast
-        also have "... \<subseteq> closure S"
-          apply (simp add: X_def fim [symmetric])
-          apply (rule closure_mono)
-          by (auto simp: A_def)
-        finally show "(\<Inter>n. X n) \<subseteq> closure S" .
-      qed
+        using Xsubclo by blast
       moreover have "(\<Inter>n. X n) \<inter> S \<subseteq> {}"
       proof (clarify, clarsimp simp: X_def fim [symmetric])
         fix x assume x [rule_format]: "\<forall>n. f x \<in> closure (f ` A n)" and "cmod x < 1"
@@ -984,13 +934,9 @@
     proof (cases "bounded S")
       case True
       have bouX: "bounded (X n)" for n
-        apply (simp add: X_def)
-        apply (rule bounded_closure)
-        by (metis A01 fim image_mono bounded_subset [OF True])
+        by (meson True Xsubclo bounded_closure bounded_subset)
       have compaX: "compact (X n)" for n
-        apply (simp add: compact_eq_bounded_closed bouX)
-        apply (auto simp: X_def)
-        done
+        by (simp add: bouX cloX compact_eq_bounded_closed)
       have "connected (\<Inter>n. X n)"
         by (metis nestX compaX connX connected_nest)
       then show ?thesis
@@ -1085,10 +1031,8 @@
                 by (simp add: j_def \<open>finite J\<close>)
               have "\<Inter> ((\<lambda>n. X n \<inter> closure U) ` J) = X j \<inter> closure U"
                 using False jmax nestX \<open>j \<in> J\<close> by auto
-              then have "X j \<inter> closure U = X j \<inter> U"
-                apply safe
-                using DiffI J empty apply auto[1]
-                using closure_subset by blast
+              then have XU: "X j \<inter> closure U = X j \<inter> U"
+                using J closure_subset empty by fastforce
               then have "openin (top_of_set (X j)) (X j \<inter> closure U)"
                 by (simp add: openin_open_Int \<open>open U\<close>)
               moreover have "closedin (top_of_set (X j)) (X j \<inter> closure U)"
@@ -1096,13 +1040,7 @@
               moreover have "X j \<inter> closure U \<noteq> X j"
                 by (metis unboundedX \<open>compact (closure U)\<close> bounded_subset compact_eq_bounded_closed inf.order_iff)
               moreover have "X j \<inter> closure U \<noteq> {}"
-              proof -
-                have "C \<noteq> {}"
-                  using C in_components_nonempty by blast
-                moreover have "C \<subseteq> X j \<inter> closure U"
-                  using \<open>C \<subseteq> K\<close> \<open>K \<subseteq> U\<close> Ksub closure_subset by blast
-                ultimately show ?thesis by blast
-              qed
+                by (metis Cco Ksub UNIV_I \<open>C \<subseteq> K\<close> \<open>K \<subseteq> U\<close> XU bot.extremum_uniqueI in_components_maximal le_INF_iff le_inf_iff)
               ultimately show False
                 using connX [of j] by (force simp: connected_clopen)
             qed
@@ -1115,7 +1053,7 @@
           have "x \<notin> V"
             using \<open>U \<inter> V = {}\<close> \<open>open V\<close> closure_iff_nhds_not_empty that(2) by blast
           then show ?thesis
-            by (metis (no_types) Diff_iff INT_I V \<open>K \<subseteq> U\<close> contra_subsetD that(1))
+            by (metis (no_types) Diff_iff INT_I V \<open>K \<subseteq> U\<close> subsetD that(1))
         qed
         ultimately show False
           by (auto simp: open_Int_closure_eq_empty [OF \<open>open V\<close>, of U])
@@ -1160,8 +1098,11 @@
         proof -
           have "C \<inter> frontier S = {}"
             using that by (simp add: C_ccsw)
-          then show False
-            by (metis C_ccsw ComplI Compl_eq_Compl_iff Diff_subset False \<open>w \<notin> S\<close> clo_ccs closure_closed compl_bot_eq connected_component_eq_UNIV connected_component_eq_empty empty_subsetI frontier_complement frontier_def frontier_not_empty frontier_of_connected_component_subset le_inf_iff subset_antisym)
+          moreover have "closed C"
+            using C_ccsw clo_ccs by blast
+          ultimately show False
+            by (metis C False \<open>S \<noteq> UNIV\<close> C_ccsw bot_eq_sup_iff connected_component_eq_UNIV frontier_Int_closed
+                frontier_closed frontier_complement frontier_eq_empty frontier_of_components_subset in_components_maximal inf.orderE)
         qed
         then show "connected_component_set (- S) w \<inter> frontier S \<noteq> {}"
           by auto
@@ -1176,15 +1117,13 @@
           have "\<not> bounded (-S)"
             by (simp add: True cobounded_imp_unbounded)
           then have "connected_component_set (- S) z \<noteq> {}"
-            apply (simp only: connected_component_eq_empty)
+            unfolding connected_component_eq_empty
             using confr openS \<open>bounded C\<close> \<open>w \<notin> S\<close>
             apply (simp add: frontier_def interior_open C_ccsw)
             by (metis ComplI Compl_eq_Diff_UNIV connected_UNIV closed_closure closure_subset connected_component_eq_self
                       connected_diff_open_from_closed subset_UNIV)
           then show "frontier (connected_component_set (- S) z) \<noteq> {}"
-            apply (simp add: frontier_eq_empty connected_component_eq_UNIV)
-            apply (metis False compl_top_eq double_compl)
-            done
+            by (metis False \<open>S \<noteq> UNIV\<close> connected_component_eq_UNIV frontier_complement frontier_eq_empty)
         qed
         then show "connected_component_set (- S) z \<inter> frontier S \<noteq> {}"
           by auto
@@ -1205,15 +1144,14 @@
     by (metis C C_ccsw False bounded_empty compl_top_eq connected_component_eq_UNIV double_compl frontier_not_empty in_components_nonempty)
   ultimately obtain z where zin: "z \<in> frontier S" and z: "z \<in> connected_component_set (- S) w"
     by blast
-  have *: "connected_component_set (frontier S) z \<in> components(frontier S)"
+  have "connected_component_set (frontier S) z \<in> components(frontier S)"
     by (simp add: \<open>z \<in> frontier S\<close> componentsI)
   with prev False have "\<not> bounded (connected_component_set (frontier S) z)"
     by simp
   moreover have "connected_component (- S) w = connected_component (- S) z"
     using connected_component_eq [OF z] by force
   ultimately show ?thesis
-    by (metis C_ccsw * zin bounded_subset closed_Compl closure_closed connected_component_maximal
-              connected_component_refl connected_connected_component frontier_closures in_components_subset le_inf_iff mem_Collect_eq openS)
+    by (metis C_ccsw SC_Chain.openS SC_Chain_axioms bounded_subset closed_Compl connected_component_mono frontier_complement frontier_subset_eq)
 qed
 
 lemma empty_inside:
@@ -1245,7 +1183,8 @@
   interpret SC_Chain
     using assms by (simp add: SC_Chain_def)
   have "?fp \<and> ?ucc \<and> ?ei"
-    using empty_inside empty_inside_imp_simply_connected frontier_properties unbounded_complement_components winding_number_zero by blast
+    using empty_inside empty_inside_imp_simply_connected frontier_properties 
+      unbounded_complement_components winding_number_zero by blast
   then show ?fp ?ucc ?ei
     by blast+
 qed
@@ -1253,16 +1192,20 @@
 lemma simply_connected_iff_simple:
   fixes S :: "complex set"
   assumes "open S" "bounded S"
-  shows "simply_connected S \<longleftrightarrow> connected S \<and> connected(- S)"
-  apply (simp add: simply_connected_eq_unbounded_complement_components assms, safe)
-   apply (metis DIM_complex assms(2) cobounded_has_bounded_component double_compl order_refl)
-  by (meson assms inside_bounded_complement_connected_empty simply_connected_eq_empty_inside simply_connected_eq_unbounded_complement_components)
+  shows "simply_connected S \<longleftrightarrow> connected S \<and> connected(- S)" (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (metis DIM_complex assms cobounded_has_bounded_component double_complement dual_order.order_iff_strict
+        simply_connected_eq_unbounded_complement_components)
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (simp add: assms connected_frontier_simple simply_connected_eq_frontier_properties)
+qed
 
 lemma subset_simply_connected_imp_inside_subset:
   fixes A :: "complex set"
   assumes "simply_connected A" "open A" "B \<subseteq> A"
   shows   "inside B \<subseteq> A" 
-by (metis assms Diff_eq_empty_iff inside_mono subset_empty simply_connected_eq_empty_inside)
+  by (metis assms Diff_eq_empty_iff inside_mono subset_empty simply_connected_eq_empty_inside)
 
 subsection\<open>Further equivalences based on continuous logs and sqrts\<close>
 
@@ -1305,7 +1248,7 @@
 lemma continuous_sqrt:
   fixes f :: "complex\<Rightarrow>complex"
   assumes contf: "continuous_on S f" and nz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
-  and prev: "\<And>f::complex\<Rightarrow>complex.
+    and prev: "\<And>f::complex\<Rightarrow>complex.
                 \<lbrakk>continuous_on S f; \<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0\<rbrakk>
                   \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>z \<in> S. f z = exp(g z))"
   shows "\<exists>g. continuous_on S g \<and> (\<forall>z \<in> S. f z = (g z)\<^sup>2)"
@@ -1313,8 +1256,8 @@
   obtain g where contg: "continuous_on S g" and geq: "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"
     using contf nz prev by metis
   show ?thesis
-proof (intro exI ballI conjI)
-  show "continuous_on S (\<lambda>z. exp(g z/2))"
+  proof (intro exI ballI conjI)
+    show "continuous_on S (\<lambda>z. exp(g z/2))"
       by (intro continuous_intros) (auto simp: contg)
     show "\<And>z. z \<in> S \<Longrightarrow> f z = (exp (g z/2))\<^sup>2"
       by (metis (no_types, lifting) divide_inverse exp_double geq mult.left_commute mult.right_neutral right_inverse zero_neq_numeral)
@@ -1343,12 +1286,14 @@
         using contg [unfolded continuous_on_iff] by (metis \<open>g z \<noteq> 0\<close> \<open>z \<in> S\<close> zero_less_norm_iff)
       then have \<delta>: "\<And>w. \<lbrakk>w \<in> S; w \<in> ball z \<delta>\<rbrakk> \<Longrightarrow> g w + g z \<noteq> 0"
         apply (clarsimp simp: dist_norm)
-        by (metis \<open>g z \<noteq> 0\<close> add_diff_cancel_left' diff_0_right norm_eq_zero norm_increases_online norm_minus_commute norm_not_less_zero not_less_iff_gr_or_eq)
+        by (metis add_diff_cancel_left' dist_0_norm dist_complex_def less_le_not_le norm_increases_online norm_minus_commute)
       have *: "(\<lambda>x. (f x - f z) / (x - z) / (g x + g z)) \<midarrow>z\<rightarrow> deriv f z / (g z + g z)"
-        apply (intro tendsto_intros)
-        using SC_Chain.openS SC_Chain_axioms \<open>f holomorphic_on S\<close> \<open>z \<in> S\<close> has_field_derivativeD holomorphic_derivI apply fastforce
-        using \<open>z \<in> S\<close> contg continuous_on_eq_continuous_at isCont_def openS apply blast
-        by (simp add: \<open>g z \<noteq> 0\<close>)
+      proof (intro tendsto_intros)
+        show "(\<lambda>x. (f x - f z) / (x - z)) \<midarrow>z\<rightarrow> deriv f z"
+          using \<open>f holomorphic_on S\<close> \<open>z \<in> S\<close> has_field_derivative_iff holomorphic_derivI openS by blast
+        show "g \<midarrow>z\<rightarrow> g z"
+          using \<open>z \<in> S\<close> contg continuous_on_eq_continuous_at isCont_def openS by blast
+      qed (simp add: \<open>g z \<noteq> 0\<close>)
       then have "(g has_field_derivative deriv f z / (g z + g z)) (at z)"
         unfolding has_field_derivative_iff
       proof (rule Lim_transform_within_open)
@@ -1386,20 +1331,9 @@
 proof -
   interpret SC_Chain
     using assms by (simp add: SC_Chain_def)
-  have "?log \<and> ?sqrt"
-proof -
-  have *: "\<lbrakk>\<alpha> \<Longrightarrow> \<beta>; \<beta> \<Longrightarrow> \<gamma>; \<gamma> \<Longrightarrow> \<alpha>\<rbrakk>
-           \<Longrightarrow> (\<alpha> \<longleftrightarrow> \<beta>) \<and> (\<alpha> \<longleftrightarrow> \<gamma>)" for \<alpha> \<beta> \<gamma>
-    by blast
-  show ?thesis
-    apply (rule *)
-    apply (simp add: local.continuous_log winding_number_zero)
-    apply (simp add: continuous_sqrt)
-    apply (simp add: continuous_sqrt_imp_simply_connected)
-    done
-qed
-  then show ?log ?sqrt
-    by safe
+  show ?log ?sqrt
+    using local.continuous_log winding_number_zero continuous_sqrt continuous_sqrt_imp_simply_connected 
+    by auto
 qed
 
 
@@ -1550,21 +1484,24 @@
             apply (clarsimp simp add: path_image_subpath_gen)
             by (metis subsetD le_less_trans \<open>dist u t < d\<close> d dist_commute dist_in_closed_segment)
           have *: "path (g \<circ> subpath t u p)"
-            apply (rule path_continuous_image)
-            using \<open>path p\<close> t that apply auto[1]
-            using piB contg continuous_on_subset by blast
+          proof (rule path_continuous_image)
+            show "path (subpath t u p)"
+              using \<open>path p\<close> t that by auto
+            show "continuous_on (path_image (subpath t u p)) g"
+              using piB contg continuous_on_subset by blast
+          qed
           have "(g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \<i>)
               =  winding_number (exp \<circ> g \<circ> subpath t u p) 0"
             using winding_number_compose_exp [OF *]
             by (simp add: pathfinish_def pathstart_def o_assoc)
-          also have "... = winding_number (\<lambda>w. subpath t u p w - \<zeta>) 0"
+          also have "\<dots> = winding_number (\<lambda>w. subpath t u p w - \<zeta>) 0"
           proof (rule winding_number_cong)
             have "exp(g y) = y - \<zeta>" if "y \<in> (subpath t u p) ` {0..1}" for y
               by (metis that geq path_image_def piB subset_eq)
             then show "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> (exp \<circ> g \<circ> subpath t u p) x = subpath t u p x - \<zeta>"
               by auto
           qed
-          also have "... = winding_number (\<lambda>w. subpath 0 u p w - \<zeta>) 0 -
+          also have "\<dots> = winding_number (\<lambda>w. subpath 0 u p w - \<zeta>) 0 -
                            winding_number (\<lambda>w. subpath 0 t p w - \<zeta>) 0"
             apply (simp add: winding_number_offset [symmetric])
             using winding_number_subpath_combine [OF \<open>path p\<close> \<zeta>, of 0 t u] \<open>t \<in> {0..1}\<close> \<open>u \<in> {0..1}\<close>
@@ -1637,9 +1574,8 @@
   then obtain a where "homotopic_loops (-{\<zeta>}) p (\<lambda>t. a)" ..
   then have "winding_number p \<zeta> = winding_number (\<lambda>t. a) \<zeta>" "a \<noteq> \<zeta>"
     using winding_number_homotopic_loops homotopic_loops_imp_subset by (force simp:)+
-  moreover have "winding_number (\<lambda>t. a) \<zeta> = 0"
-    by (metis winding_number_zero_const \<open>a \<noteq> \<zeta>\<close>)
-  ultimately show ?lhs by metis
+  then show ?lhs
+    using winding_number_zero_const by auto
 qed
 
 lemma winding_number_homotopic_paths_null_explicit_eq:
@@ -1650,7 +1586,7 @@
   assume ?lhs
   then show ?rhs
     using homotopic_loops_imp_homotopic_paths_null 
-    by (force simp add: linepath_refl winding_number_homotopic_loops_null_eq [OF assms])
+    by (force simp: linepath_refl winding_number_homotopic_loops_null_eq [OF assms])
 next
   assume ?rhs
   then show ?lhs
@@ -1688,7 +1624,7 @@
     using winding_number_homotopic_paths_null_explicit_eq by blast
   then show ?rhs
     using homotopic_paths_imp_pathstart assms
-    by (fastforce simp add: dest: homotopic_paths_imp_homotopic_loops homotopic_paths_loop_parts)
+    by (fastforce simp: dest: homotopic_paths_imp_homotopic_loops homotopic_paths_loop_parts)
 qed (simp add: winding_number_homotopic_paths)
 
 lemma winding_number_homotopic_loops_eq:
--- a/src/HOL/Number_Theory/Eratosthenes.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Number_Theory/Eratosthenes.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -109,8 +109,7 @@
     with q have a: "a = Suc n * q - 2" by simp
     with B have "q + n * q < n + n + 2" by auto
     then have "m * q < m * 2" by (simp add: m_def)
-    with \<open>m > 0\<close> have "q < 2" by simp
-    with \<open>q > 0\<close> have "q = 1" by simp
+    with \<open>m > 0\<close> \<open>q > 0\<close> have "q = 1" by simp
     with a have "a = n - 1" by simp
     with \<open>n > 0\<close> C show False by simp
   qed
@@ -212,7 +211,6 @@
     from \<open>m dvd Suc n\<close> obtain q where "Suc n = m * q" ..
     with \<open>Suc (Suc n) \<le> m\<close> have "Suc (m * q) \<le> m" by simp
     then have "m * q < m" by arith
-    then have "q = 0" by simp
     with \<open>Suc n = m * q\<close> show ?thesis by simp
   qed
   have aux2: "m dvd q"
@@ -303,15 +301,7 @@
 
 lemma primes_upto_sieve [code]:
   "primes_upto n = map fst (filter snd (enumerate 2 (sieve 1 (replicate (n - 1) True))))"
-proof -
-  have "primes_upto n = sorted_list_of_set (numbers_of_marks 2 (sieve 1 (replicate (n - 1) True)))"
-    apply (rule sorted_distinct_set_unique)
-    apply (simp_all only: set_primes_upto_sieve numbers_of_marks_def)
-    apply auto
-    done
-  then show ?thesis
-    by (simp add: sorted_list_of_set_numbers_of_marks)
-qed
+  using primes_upto_def set_primes_upto set_primes_upto_sieve sorted_list_of_set_numbers_of_marks by presburger
 
 lemma prime_in_primes_upto: "prime n \<longleftrightarrow> n \<in> set (primes_upto n)"
   by (simp add: set_primes_upto)
--- a/src/HOL/Number_Theory/Residues.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Number_Theory/Residues.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -59,12 +59,18 @@
 qed
 
 lemma comm_monoid: "comm_monoid R"
-  unfolding R_m_def residue_ring_def
-  apply (rule comm_monoidI)
-    using m_gt_one  apply auto
-  apply (metis mod_mult_right_eq mult.assoc mult.commute)
-  apply (metis mult.commute)
-  done
+proof -
+  have "\<And>x y z. \<lbrakk>x \<in> carrier R; y \<in> carrier R; z \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> y \<otimes> z = x \<otimes> (y \<otimes> z)"
+    "\<And>x y. \<lbrakk>x \<in> carrier R; y \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> y = y \<otimes> x"
+    unfolding R_m_def residue_ring_def
+    by (simp_all add: algebra_simps mod_mult_right_eq)
+  then show ?thesis
+    unfolding R_m_def residue_ring_def
+    by unfold_locales (use m_gt_one in simp_all)
+qed
+
+interpretation comm_monoid R
+  using comm_monoid by blast
 
 lemma cring: "cring R"
   apply (intro cringI abelian_group comm_monoid)
@@ -101,21 +107,31 @@
 lemma res_one_eq: "\<one> = 1"
   by (auto simp: R_m_def residue_ring_def units_of_def)
 
-lemma res_units_eq: "Units R = {x. 0 < x \<and> x < m \<and> coprime x m}"
-  using m_gt_one
-  apply (auto simp add: Units_def R_m_def residue_ring_def ac_simps invertible_coprime intro: ccontr)
-  apply (subst (asm) coprime_iff_invertible'_int)
-   apply (auto simp add: cong_def)
-  done
+lemma res_units_eq: "Units R = {x. 0 < x \<and> x < m \<and> coprime x m}" (is "_ = ?rhs")
+proof
+  show "Units R \<subseteq> ?rhs"
+    using zero_less_mult_iff invertible_coprime 
+    by (fastforce simp: Units_def R_m_def residue_ring_def)
+next
+  show "?rhs \<subseteq> Units R"
+    unfolding Units_def R_m_def residue_ring_def 
+    by (force simp add: cong_def coprime_iff_invertible'_int mult.commute)
+qed
 
 lemma res_neg_eq: "\<ominus> x = (- x) mod m"
-  using m_gt_one unfolding R_m_def a_inv_def m_inv_def residue_ring_def
-  apply simp
-  apply (rule the_equality)
-   apply (simp add: mod_add_right_eq)
-   apply (simp add: add.commute mod_add_right_eq)
-  apply (metis add.right_neutral minus_add_cancel mod_add_right_eq mod_pos_pos_trivial)
-  done
+proof -
+  have "\<ominus> x = (THE y. 0 \<le> y \<and> y < m \<and> (x + y) mod m = 0 \<and> (y + x) mod m = 0)"
+    by (simp add: R_m_def a_inv_def m_inv_def residue_ring_def)
+  also have "\<dots> = (- x) mod m"
+  proof -
+    have "\<And>y. 0 \<le> y \<and> y < m \<and> (x + y) mod m = 0 \<and> (y + x) mod m = 0 \<Longrightarrow>
+         y = - x mod m"
+      by (metis minus_add_cancel mod_add_eq plus_int_code(1) zmod_trivial_iff)
+    then show ?thesis
+      by (intro the_equality) (use m_gt_one  in \<open>simp add: add.commute mod_add_right_eq\<close>)
+  qed
+  finally show ?thesis .
+qed
 
 lemma finite [iff]: "finite (carrier R)"
   by (simp add: res_carrier_eq)
@@ -148,10 +164,15 @@
 (* FIXME revise algebra library to use 1? *)
 lemma pow_cong: "(x mod m) [^] n = x^n mod m"
   using m_gt_one
-  apply (induct n)
-  apply (auto simp add: nat_pow_def one_cong)
-  apply (metis mult.commute mult_cong)
-  done
+proof (induct n)
+  case 0
+  then show ?case
+    by (simp add: one_cong) 
+next
+  case (Suc n)
+  then show ?case
+    by (simp add: mult_cong power_commutes) 
+qed
 
 lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m"
   by (metis mod_minus_eq res_neg_eq)
@@ -189,10 +210,7 @@
 
 text \<open>Other useful facts about the residue ring.\<close>
 lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2"
-  apply (simp add: res_one_eq res_neg_eq)
-  apply (metis add.commute add_diff_cancel mod_mod_trivial one_add_one uminus_add_conv_diff
-    zero_neq_one zmod_zminus1_eq_if)
-  done
+  using one_cong res_neg_eq res_one_eq zmod_zminus1_eq_if by fastforce
 
 end
 
@@ -206,9 +224,7 @@
 
 sublocale residues_prime < residues p
   unfolding R_def residues_def
-  using p_prime apply auto
-  apply (metis (full_types) of_nat_1 of_nat_less_iff prime_gt_1_nat)
-  done
+  by (auto simp: p_prime prime_gt_1_int)
 
 context residues_prime
 begin
@@ -227,7 +243,7 @@
 
 lemma p_coprime_right_int:
   "coprime a (int p) \<longleftrightarrow> \<not> int p dvd a"
-  using p_coprime_left_int [of a] by (simp add: ac_simps)
+  using coprime_commute p_coprime_left_int by blast
 
 lemma is_field: "field R"
 proof -
@@ -239,9 +255,7 @@
 qed
 
 lemma res_prime_units_eq: "Units R = {1..p - 1}"
-  apply (subst res_units_eq)
-  apply (auto simp add: p_coprime_right_int zdvd_not_zless)
-  done
+  by (auto simp add: res_units_eq p_coprime_right_int zdvd_not_zless)
 
 end
 
@@ -277,7 +291,7 @@
 qed
 
 lemma (in residues_prime) prime_totient_eq: "totient p = p - 1"
-  using totient_eq by (simp add: res_prime_units_eq)
+  using p_prime totient_prime by blast
 
 lemma (in residues) euler_theorem:
   assumes "coprime a m"
@@ -322,9 +336,8 @@
 
 lemma (in field) inv_pair_lemma: "x \<in> Units R \<Longrightarrow> y \<in> Units R \<Longrightarrow>
     {x, inv x} \<noteq> {y, inv y} \<Longrightarrow> {x, inv x} \<inter> {y, inv y} = {}"
-  apply auto
-  apply (metis Units_inv_inv)+
-  done
+  by auto
+
 
 lemma (in residues_prime) wilson_theorem1:
   assumes a: "p > 2"
@@ -333,27 +346,20 @@
   let ?Inverse_Pairs = "{{x, inv x}| x. x \<in> Units R - {\<one>, \<ominus> \<one>}}"
   have UR: "Units R = {\<one>, \<ominus> \<one>} \<union> \<Union>?Inverse_Pairs"
     by auto
+  have 11: "\<one> \<noteq> \<ominus> \<one>"
+    using a one_eq_neg_one by force
   have "(\<Otimes>i\<in>Units R. i) = (\<Otimes>i\<in>{\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i\<in>\<Union>?Inverse_Pairs. i)"
     apply (subst UR)
     apply (subst finprod_Un_disjoint)
-         apply (auto intro: funcsetI)
-    using inv_one apply auto[1]
-    using inv_eq_neg_one_eq apply auto
+    using inv_one inv_eq_neg_one_eq apply (auto intro!: funcsetI)+
     done
   also have "(\<Otimes>i\<in>{\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>"
-    apply (subst finprod_insert)
-        apply auto
-    apply (frule one_eq_neg_one)
-    using a apply force
-    done
+    by (simp add: 11)
   also have "(\<Otimes>i\<in>(\<Union>?Inverse_Pairs). i) = (\<Otimes>A\<in>?Inverse_Pairs. (\<Otimes>y\<in>A. y))"
-    apply (subst finprod_Union_disjoint)
-       apply (auto simp: pairwise_def disjnt_def)
-     apply (metis Units_inv_inv)+
-    done
+    by (rule finprod_Union_disjoint) (auto simp: pairwise_def disjnt_def dest!: inv_eq_imp_eq)
   also have "\<dots> = \<one>"
     apply (rule finprod_one_eqI)
-     apply auto
+    apply clarsimp
     apply (subst finprod_insert)
         apply auto
     apply (metis inv_eq_self)
@@ -365,11 +371,8 @@
   also have "\<dots> = (\<Prod>i\<in>Units R. i) mod p"
     by (rule prod_cong) auto
   also have "\<dots> = fact (p - 1) mod p"
-    apply (simp add: fact_prod)
     using assms
-    apply (subst res_prime_units_eq)
-    apply (simp add: int_prod zmod_int prod_int_eq)
-    done
+    by (simp add: res_prime_units_eq int_prod zmod_int prod_int_eq fact_prod)
   finally have "fact (p - 1) mod p = \<ominus> \<one>" .
   then show ?thesis
     by (simp add: cong_def res_neg_eq res_one_eq zmod_int)
@@ -396,7 +399,7 @@
 lemma mod_nat_int_pow_eq:
   fixes n :: nat and p a :: int
   shows "a \<ge> 0 \<Longrightarrow> p \<ge> 0 \<Longrightarrow> (nat a ^ n) mod (nat p) = nat ((a ^ n) mod p)"
-  by (simp add: int_one_le_iff_zero_less nat_mod_distrib order_less_imp_le nat_power_eq[symmetric])
+  by (simp add: nat_mod_as_int)
 
 theorem residue_prime_mult_group_has_gen:
  fixes p :: nat
--- a/src/HOL/Probability/Giry_Monad.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Probability/Giry_Monad.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -31,7 +31,7 @@
 qed
 
 lemma (in subprob_space) emeasure_subprob_space_less_top: "emeasure M A \<noteq> top"
-  using emeasure_finite[of A] .
+  by simp
 
 lemma prob_space_imp_subprob_space:
   "prob_space M \<Longrightarrow> subprob_space M"
@@ -44,10 +44,10 @@
   by (rule subprob_spaceI) (simp_all add: emeasure_space_1 not_empty)
 
 lemma subprob_space_sigma [simp]: "\<Omega> \<noteq> {} \<Longrightarrow> subprob_space (sigma \<Omega> X)"
-by(rule subprob_spaceI)(simp_all add: emeasure_sigma space_measure_of_conv)
+  by(rule subprob_spaceI)(simp_all add: emeasure_sigma space_measure_of_conv)
 
 lemma subprob_space_null_measure: "space M \<noteq> {} \<Longrightarrow> subprob_space (null_measure M)"
-by(simp add: null_measure_def)
+  by(simp add: null_measure_def)
 
 lemma (in subprob_space) subprob_space_distr:
   assumes f: "f \<in> measurable M M'" and "space M' \<noteq> {}" shows "subprob_space (distr M M' f)"
@@ -343,7 +343,7 @@
   assumes [measurable]: "f \<in> measurable M N"
   shows "(\<lambda>M'. distr M' N f) \<in> measurable (subprob_algebra M) (subprob_algebra N)"
 proof (cases "space N = {}")
-  assume not_empty: "space N \<noteq> {}"
+  case False
   show ?thesis
   proof (rule measurable_subprob_algebra)
     fix A assume A: "A \<in> sets N"
@@ -355,8 +355,8 @@
     also have "\<dots>"
       using A by (intro measurable_emeasure_subprob_algebra) simp
     finally show "(\<lambda>M'. emeasure (distr M' N f) A) \<in> borel_measurable (subprob_algebra M)" .
-  qed (auto intro!: subprob_space.subprob_space_distr simp: space_subprob_algebra not_empty cong: measurable_cong_sets)
-qed (insert assms, auto simp: measurable_empty_iff space_subprob_algebra_empty_iff)
+  qed (auto intro!: subprob_space.subprob_space_distr simp: space_subprob_algebra False cong: measurable_cong_sets)
+qed (use assms in \<open>auto simp: measurable_empty_iff space_subprob_algebra_empty_iff\<close>)
 
 lemma emeasure_space_subprob_algebra[measurable]:
   "(\<lambda>a. emeasure a (space a)) \<in> borel_measurable (subprob_algebra N)"
@@ -565,10 +565,7 @@
 
 lemma subprob_space_return_ne:
   assumes "space M \<noteq> {}" shows "subprob_space (return M x)"
-proof
-  show "emeasure (return M x) (space (return M x)) \<le> 1"
-    by (subst emeasure_return) (auto split: split_indicator)
-qed (simp, fact)
+  by (metis assms emeasure_return indicator_simps(2) sets.top space_return subprob_spaceI subprob_space_return zero_le)
 
 lemma measure_return: assumes X: "X \<in> sets M" shows "measure (return M x) X = indicator X x"
   unfolding measure_def emeasure_return[OF X, of x] by (simp split: split_indicator)
@@ -654,8 +651,7 @@
         by (auto simp add: emeasure_distr f_M' cong: measurable_cong_sets)
       also have "\<dots> = (\<integral>\<^sup>+M''. emeasure (g x) (f M'' -` A \<inter> space M) \<partial>?R)"
         apply (subst emeasure_pair_measure_alt)
-        apply (rule measurable_sets[OF _ A])
-        apply (auto simp add: f_M' cong: measurable_cong_sets)
+        apply (force simp add: f_M' cong: measurable_cong_sets intro!: measurable_sets[OF _ A])
         apply (intro nn_integral_cong arg_cong[where f="emeasure (g x)"])
         apply (auto simp: space_subprob_algebra space_pair_measure)
         done
@@ -668,10 +664,10 @@
     qed
   qed
   also have "\<dots>"
-    apply (intro measurable_compose[OF measurable_pair_measure measurable_distr])
-    apply (rule return_measurable)
-    apply measurable
-    done
+  proof (intro measurable_compose[OF measurable_pair_measure measurable_distr])
+    show "return L \<in> L \<rightarrow>\<^sub>M subprob_algebra L"
+      by (rule return_measurable)
+  qed measurable
   finally show ?thesis .
 qed
 
@@ -703,19 +699,15 @@
 
   have *: "\<And>x. fst x \<in> space M \<Longrightarrow> snd x \<in> A (fst x) \<longleftrightarrow> x \<in> (SIGMA x:space M. A x)"
     by (auto simp: fun_eq_iff)
-  have "(\<lambda>(x, y). indicator (A x) y::ennreal) \<in> borel_measurable (M \<Otimes>\<^sub>M N)"
+  have MN: "Measurable.pred (M \<Otimes>\<^sub>M N) (\<lambda>w. w \<in> Sigma (space M) A)"
+    by auto
+  then have "(\<lambda>(x, y). indicator (A x) y::ennreal) \<in> borel_measurable (M \<Otimes>\<^sub>M N)"
     apply measurable
-    apply (subst measurable_cong)
-    apply (rule *)
-    apply (auto simp: space_pair_measure)
-    done
+    by (smt (verit, best) MN measurable_cong mem_Sigma_iff prod.collapse space_pair_measure)
   then have "(\<lambda>x. integral\<^sup>N (L x) (indicator (A x))) \<in> borel_measurable M"
     by (intro nn_integral_measurable_subprob_algebra2[where N=N] L)
   then show "(\<lambda>x. emeasure (L x) (A x)) \<in> borel_measurable M"
-    apply (rule measurable_cong[THEN iffD1, rotated])
-    apply (rule nn_integral_indicator)
-    apply (simp add: subprob_measurableD[OF L] **)
-    done
+    by (smt (verit) "**" L measurable_cong_simp nn_integral_indicator sets_kernel)
 qed
 
 lemma measure_measurable_subprob_algebra2:
@@ -751,7 +743,7 @@
   next
     assume "space (subprob_algebra N) \<noteq> {}"
     with eq show ?thesis
-      by (fastforce simp add: space_subprob_algebra)
+      by (smt (verit) equals0I mem_Collect_eq space_subprob_algebra)
   qed
 qed
 
@@ -807,10 +799,11 @@
   assume [simp]: "space N \<noteq> {}"
   fix M assume M: "M \<in> space (subprob_algebra (subprob_algebra N))"
   then have "(\<integral>\<^sup>+M'. emeasure M' (space N) \<partial>M) \<le> (\<integral>\<^sup>+M'. 1 \<partial>M)"
-    apply (intro nn_integral_mono)
-    apply (auto simp: space_subprob_algebra
-                 dest!: sets_eq_imp_space_eq subprob_space.emeasure_space_le_1)
-    done
+  proof (intro nn_integral_mono)
+    show "\<And>x. \<lbrakk>M \<in> space (subprob_algebra (subprob_algebra N)); x \<in> space M\<rbrakk>
+         \<Longrightarrow> emeasure x (space N) \<le> 1"
+      by (smt (verit) mem_Collect_eq sets_eq_imp_space_eq space_subprob_algebra subprob_space.subprob_emeasure_le_1)
+  qed
   with M show "subprob_space (join M)"
     by (intro subprob_spaceI)
        (auto simp: emeasure_join space_subprob_algebra M dest: subprob_space.emeasure_space_le_1)
@@ -888,7 +881,7 @@
 lemma measurable_join1:
   "\<lbrakk> f \<in> measurable N K; sets M = sets (subprob_algebra N) \<rbrakk>
   \<Longrightarrow> f \<in> measurable (join M) K"
-by(simp add: measurable_def)
+  by(simp add: measurable_def)
 
 lemma
   fixes f :: "_ \<Rightarrow> real"
@@ -1051,12 +1044,15 @@
 lemma join_return':
   assumes "sets N = sets M"
   shows "join (distr M (subprob_algebra N) (return N)) = M"
-apply (rule measure_eqI)
-apply (simp add: assms)
-apply (subgoal_tac "return N \<in> measurable M (subprob_algebra N)")
-apply (simp add: emeasure_join nn_integral_distr measurable_emeasure_subprob_algebra assms)
-apply (subst measurable_cong_sets, rule assms[symmetric], rule refl, rule return_measurable)
-done
+proof (rule measure_eqI)
+  fix A
+  have "return N \<in> measurable M (subprob_algebra N)"
+    using assms by auto
+  moreover
+  assume "A \<in> sets (join (distr M (subprob_algebra N) (return N)))"
+  ultimately show "emeasure (join (distr M (subprob_algebra N) (return N))) A = emeasure M A"
+    by (simp add: emeasure_join nn_integral_distr measurable_emeasure_subprob_algebra assms)
+qed (simp add: assms)
 
 lemma join_distr_distr:
   fixes f :: "'a \<Rightarrow> 'b" and M :: "'a measure measure" and N :: "'b measure"
@@ -1107,7 +1103,7 @@
   by (simp add: bind_def)
 
 lemma sets_bind_empty: "sets M = {} \<Longrightarrow> sets (bind M f) = {{}}"
-  by (auto simp: bind_def)
+  by auto
 
 lemma space_bind_empty: "space M = {} \<Longrightarrow> space (bind M f) = {}"
   by (simp add: bind_def)
@@ -1139,11 +1135,12 @@
 lemma bind_nonempty':
   assumes "f \<in> measurable M (subprob_algebra N)" "x \<in> space M"
   shows "bind M f = join (distr M (subprob_algebra N) f)"
-  using assms
-  apply (subst bind_nonempty, blast)
-  apply (subst subprob_algebra_cong[OF sets_kernel[OF assms(1) someI_ex]], blast)
-  apply (simp add: subprob_algebra_cong[OF sets_kernel[OF assms]])
-  done
+proof -
+  have "join (distr M (subprob_algebra (f (SOME x. x \<in> space M))) f) = join (distr M (subprob_algebra N) f)"
+    by (metis assms someI_ex subprob_algebra_cong subprob_measurableD(2))
+  with assms show ?thesis
+    by (metis bind_nonempty empty_iff)
+qed
 
 lemma bind_nonempty'':
   assumes "f \<in> measurable M (subprob_algebra N)" "space M \<noteq> {}"
@@ -1182,14 +1179,15 @@
   have "(AE x in M \<bind> N. P x) \<longleftrightarrow> (\<integral>\<^sup>+ x. integral\<^sup>N (N x) (indicator {x \<in> space B. \<not> P x}) \<partial>M) = 0"
     by (simp add: AE_iff_nn_integral sets_bind[OF _ M] space_bind[OF _ M] * nn_integral_bind[where B=B]
              del: nn_integral_indicator)
-  also have "\<dots> = (AE x in M. AE y in N x. P y)"
-    apply (subst nn_integral_0_iff_AE)
+  also have "... = (AE x in M. integral\<^sup>N (N x) (indicator {x \<in> space B. \<not> P x}) = 0)"
+  proof (rule nn_integral_0_iff_AE)
+    show "(\<lambda>x. integral\<^sup>N (N x) (indicator {x \<in> space B. \<not> P x})) \<in> borel_measurable M"
     apply (rule measurable_compose[OF N nn_integral_measurable_subprob_algebra])
-    apply measurable
+      by measurable
+  qed
+  also have "\<dots> = (AE x in M. AE y in N x. P y)"
     apply (intro eventually_subst AE_I2)
-    apply (auto simp add: subprob_measurableD(1)[OF N]
-                intro!: AE_iff_measurable[symmetric])
-    done
+    by (auto simp add: subprob_measurableD(1)[OF N] intro!: AE_iff_measurable[symmetric])
   finally show ?thesis .
 qed
 
@@ -1351,13 +1349,14 @@
   assumes N: "N \<in> measurable M (subprob_algebra K)" "space M \<noteq> {}"
   assumes f: "f \<in> measurable K R"
   shows "distr (M \<bind> N) R f = (M \<bind> (\<lambda>x. distr (N x) R f))"
-  unfolding bind_nonempty''[OF N]
-  apply (subst bind_nonempty''[OF measurable_compose[OF N(1) measurable_distr] N(2)])
-  apply (rule f)
-  apply (simp add: join_distr_distr[OF _ f, symmetric])
-  apply (subst distr_distr[OF measurable_distr, OF f N(1)])
-  apply (simp add: comp_def)
-  done
+proof -
+  have "distr (join (distr M (subprob_algebra K) N)) R f =
+       join (distr M (subprob_algebra R) (\<lambda>x. distr (N x) R f))"
+    by (simp add: assms distr_distr[OF measurable_distr] comp_def flip: join_distr_distr)
+  with assms show ?thesis
+    unfolding bind_nonempty''[OF N]
+    by (smt (verit) bind_nonempty sets_distr subprob_algebra_cong)
+qed
 
 lemma bind_distr:
   assumes f[measurable]: "f \<in> measurable M X"
@@ -1393,16 +1392,20 @@
   show "sets (restrict_space (bind M N) X) = sets (bind M (\<lambda>x. restrict_space (N x) X))"
     by (simp add: sets_restrict_space assms(2) sets_bind[OF sets_kernel[OF restrict_space_measurable[OF assms(4,3,1)]]])
   fix A assume "A \<in> sets (restrict_space (M \<bind> N) X)"
-  with X have "A \<in> sets K" "A \<subseteq> X"
+  with X have A: "A \<in> sets K" "A \<subseteq> X"
     by (auto simp: sets_restrict_space)
-  then show "emeasure (restrict_space (M \<bind> N) X) A = emeasure (M \<bind> (\<lambda>x. restrict_space (N x) X)) A"
-    using assms
-    apply (subst emeasure_restrict_space)
-    apply (simp_all add: emeasure_bind[OF assms(2,1)])
-    apply (subst emeasure_bind[OF _ restrict_space_measurable[OF _ _ N]])
-    apply (auto simp: sets_restrict_space emeasure_restrict_space space_subprob_algebra
-                intro!: nn_integral_cong dest!: measurable_space)
+  then have "emeasure (restrict_space (M \<bind> N) X) A = emeasure (M \<bind> N) A"
+    by (simp add: emeasure_restrict_space)
+  also have "\<dots> = \<integral>\<^sup>+ x. emeasure (N x) A \<partial>M"
+    by (metis \<open>A \<in> sets K\<close> N \<open>space M \<noteq> {}\<close> emeasure_bind)
+  also have "... = \<integral>\<^sup>+ x. emeasure (restrict_space (N x) X) A \<partial>M"
+    using A assms by (smt (verit, best) emeasure_restrict_space nn_integral_cong sets.Int_space_eq2 subprob_measurableD(2))
+  also have "\<dots> = emeasure (M \<bind> (\<lambda>x. restrict_space (N x) X)) A"
+    using A assms
+    apply (subst emeasure_bind[OF _ restrict_space_measurable])
+    apply (auto simp: sets_restrict_space)
     done
+  finally show "emeasure (restrict_space (M \<bind> N) X) A = emeasure (M \<bind> (\<lambda>x. restrict_space (N x) X)) A" .
 qed
 
 lemma bind_restrict_space:
@@ -1442,13 +1445,18 @@
      (simp_all add: space_subprob_algebra prob_space.not_empty emeasure_bind_const_prob_space)
 
 lemma bind_return_distr:
-    "space M \<noteq> {} \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> bind M (return N \<circ> f) = distr M N f"
-  apply (simp add: bind_nonempty)
-  apply (subst subprob_algebra_cong)
-  apply (rule sets_return)
-  apply (subst distr_distr[symmetric])
-  apply (auto intro!: return_measurable simp: distr_distr[symmetric] join_return')
-  done
+  assumes "space M \<noteq> {}" "f \<in> measurable M N"
+  shows "bind M (return N \<circ> f) = distr M N f"
+proof -
+  have "bind M (return N \<circ> f)
+      = join (distr M (subprob_algebra (return N (f (SOME x. x \<in> space M)))) (return N \<circ> f))"
+    by (simp add: Giry_Monad.bind_def assms)
+  also have "\<dots> = join (distr M (subprob_algebra N) (return N \<circ> f))"
+    by (metis sets_return subprob_algebra_cong)
+  also have "\<dots> = distr M N f"
+    by (metis assms(2) distr_distr join_return' return_measurable sets_distr)
+  finally show ?thesis .
+qed
 
 lemma bind_return_distr':
   "space M \<noteq> {} \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> bind M (\<lambda>x. return N (f x)) = distr M N f"
@@ -1469,6 +1477,9 @@
                          sets_kernel[OF M2 someI_ex[OF ex_in[OF \<open>space N \<noteq> {}\<close>]]]
   note space_some[simp] = sets_eq_imp_space_eq[OF this(1)] sets_eq_imp_space_eq[OF this(2)]
 
+
+  have *: "(\<lambda>x. distr x (subprob_algebra R) g) \<circ> f \<in> M \<rightarrow>\<^sub>M subprob_algebra (subprob_algebra R)"
+    using M1 M2 measurable_comp measurable_distr by blast
   have "bind M (\<lambda>x. bind (f x) g) =
         join (distr M (subprob_algebra R) (join \<circ> (\<lambda>x. (distr x (subprob_algebra R) g)) \<circ> f))"
     by (simp add: sets_eq_imp_space_eq[OF sets_fx] bind_nonempty o_def
@@ -1478,10 +1489,7 @@
                           (subprob_algebra (subprob_algebra R))
                           (\<lambda>x. distr x (subprob_algebra R) g))
                    (subprob_algebra R) join"
-      apply (subst distr_distr,
-             (blast intro: measurable_comp measurable_distr measurable_join M1 M2)+)+
-      apply (simp add: o_assoc)
-      done
+    by (simp add: distr_distr M1 M2 measurable_distr measurable_join fun.map_comp *)
   also have "join ... = bind (bind M f) g"
       by (simp add: join_assoc join_distr_distr M2 bind_nonempty cong: subprob_algebra_cong)
   finally show ?thesis ..
@@ -1637,7 +1645,7 @@
       using measurable_space[OF g]
     by (auto simp: measurable_restrict_space2_iff prob_algebra_def space_pair_measure Pi_iff
                 intro!: prob_space.prob_space_bind[where S=R] AE_I2)
-qed (insert g, simp)
+qed (use g in simp)
 
 
 lemma measurable_prob_algebra_generated:
@@ -1659,7 +1667,7 @@
       by (intro measurable_cong) auto
     then show "(\<lambda>a. emeasure (K a) \<Omega>) \<in> borel_measurable M" by simp
   qed
-qed (insert subsp, auto)
+qed (use subsp in auto)
 
 lemma in_space_prob_algebra:
   "x \<in> space (prob_algebra M) \<Longrightarrow> emeasure x (space M) = 1"
@@ -1668,13 +1676,7 @@
 
 lemma prob_space_pair:
   assumes "prob_space M" "prob_space N" shows "prob_space (M \<Otimes>\<^sub>M N)"
-proof -
-  interpret M: prob_space M by fact
-  interpret N: prob_space N by fact
-  interpret P: pair_prob_space M N proof qed
-  show ?thesis
-    by unfold_locales
-qed
+  by (metis assms measurable_fst prob_space.distr_pair_fst prob_space_distrD)
 
 lemma measurable_pair_prob[measurable]:
   "f \<in> M \<rightarrow>\<^sub>M prob_algebra N \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M prob_algebra L \<Longrightarrow> (\<lambda>x. f x \<Otimes>\<^sub>M g x) \<in> M \<rightarrow>\<^sub>M prob_algebra (N \<Otimes>\<^sub>M L)"
@@ -1738,7 +1740,7 @@
   also from assms(3) x have "... = emeasure (distr (density M f') (count_space A) g) {x}"
     by (subst emeasure_distr) simp_all
   finally show "f x = emeasure (distr (density M f') (count_space A) g) {x}" .
-qed (insert assms, auto)
+qed (use assms in auto)
 
 lemma bind_cong_AE:
   assumes "M = N"
@@ -1796,7 +1798,6 @@
     by eventually_elim auto
   thus "y \<in> space M"
     by simp
-
   show "M = return M y"
   proof (rule measure_eqI)
     fix X assume X: "X \<in> sets M"
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -102,13 +102,11 @@
 
 lemma sets_Collect_single':
   "i \<in> I \<Longrightarrow> {x\<in>space (M i). P x} \<in> sets (M i) \<Longrightarrow> {x\<in>space (PiM I M). P (x i)} \<in> sets (PiM I M)"
-  using sets_Collect_single[of i I "{x\<in>space (M i). P x}" M]
-  by (simp add: space_PiM PiE_iff cong: conj_cong)
+  by auto
 
 lemma (in finite_product_prob_space) finite_measure_PiM_emb:
   "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))"
-  using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets.sets_into_space, of I A M]
-  by auto
+  by (rule prob_times)
 
 lemma (in product_prob_space) PiM_component:
   assumes "i \<in> I"
--- a/src/HOL/Probability/Information.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Probability/Information.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -11,10 +11,10 @@
 begin
 
 lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y"
-  by (subst log_le_cancel_iff) auto
+  by simp
 
 lemma log_less: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x < y \<Longrightarrow> log a x < log a y"
-  by (subst log_less_cancel_iff) auto
+  by simp
 
 lemma sum_cartesian_product':
   "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. sum (\<lambda>y. f (x, y)) B)"
@@ -94,13 +94,11 @@
     using f nn by (intro density_RN_deriv_density) auto
   then have eq: "AE x in M. RN_deriv M (density M f) x = f x"
     using f nn by (intro density_unique) auto
-  show "(\<integral>x. f x * entropy_density b M (density M (\<lambda>x. ennreal (f x))) x \<partial>M) = (\<integral>x. f x * log b (f x) \<partial>M)"
-    apply (intro integral_cong_AE)
-    apply measurable
-    using eq nn
-    apply eventually_elim
-    apply (auto simp: entropy_density_def)
-    done
+  have "AE x in M. f x * entropy_density b M (density M (\<lambda>x. ennreal (f x))) x =
+               f x * log b (f x)"
+    using eq nn by (auto simp: entropy_density_def)
+  then show "(\<integral>x. f x * entropy_density b M (density M (\<lambda>x. ennreal (f x))) x \<partial>M) = (\<integral>x. f x * log b (f x) \<partial>M)"
+    by (intro integral_cong_AE) measurable
 qed fact+
 
 lemma (in sigma_finite_measure) KL_density_density:
@@ -240,10 +238,7 @@
   have "AE x in M. 1 = RN_deriv M M x"
   proof (rule RN_deriv_unique)
     show "density M (\<lambda>x. 1) = M"
-      apply (auto intro!: measure_eqI emeasure_density)
-      apply (subst emeasure_density)
-      apply auto
-      done
+      by (simp add: density_1)
   qed auto
   then have "AE x in M. log b (enn2real (RN_deriv M M x)) = 0"
     by (elim AE_mp) simp
@@ -373,9 +368,7 @@
       by (simp add: emeasure_pair_measure_Times) }
   then show ?thesis
     unfolding absolutely_continuous_def
-    apply (auto simp: null_sets_distr_iff)
-    apply (auto simp: null_sets_def intro!: measurable_sets)
-    done
+    by (metis emeasure_distr measurable_fst null_setsD1 null_setsD2 null_setsI sets_distr subsetI)
 qed
 
 lemma ac_snd:
@@ -390,9 +383,7 @@
       by (simp add: emeasure_pair_measure_Times) }
   then show ?thesis
     unfolding absolutely_continuous_def
-    apply (auto simp: null_sets_distr_iff)
-    apply (auto simp: null_sets_def intro!: measurable_sets)
-    done
+    by (metis emeasure_distr measurable_snd null_setsD1 null_setsD2 null_setsI sets_distr subsetI)
 qed
 
 lemma (in information_space) finite_entropy_integrable:
@@ -473,7 +464,7 @@
         unfolding \<open>Q = P\<close> by (intro measure_eqI) (auto simp: emeasure_density)
     qed auto
     then have ae_0: "AE x in P. entropy_density b P Q x = 0"
-      by eventually_elim (auto simp: entropy_density_def)
+      by (auto simp: entropy_density_def)
     then have "integrable P (entropy_density b P Q) \<longleftrightarrow> integrable Q (\<lambda>x. 0::real)"
       using ed unfolding \<open>Q = P\<close> by (intro integrable_cong_AE) auto
     then show "integrable Q (entropy_density b P Q)" by simp
@@ -579,7 +570,7 @@
   have B: "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)"
     by (rule subdensity_real[OF measurable_snd Pxy Py]) (insert Py_nn Pxy_nn, auto simp: space_pair_measure)
   ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
-    by eventually_elim auto
+    by auto
 
   show "?M = ?R"
     unfolding M f_def using Pxy_nn Px_nn Py_nn
@@ -599,9 +590,8 @@
       intro!: borel_measurable_times borel_measurable_log borel_measurable_divide)
   ultimately have int: "integrable (S \<Otimes>\<^sub>M T) f"
     apply (rule integrable_cong_AE_imp)
-    using A B AE_space
-    by eventually_elim
-       (auto simp: f_def log_divide_eq log_mult_eq field_simps space_pair_measure Px_nn Py_nn Pxy_nn
+    using A B
+    by  (auto simp: f_def log_divide_eq log_mult_eq field_simps space_pair_measure Px_nn Py_nn Pxy_nn
                   less_le)
 
   show "0 \<le> ?M" unfolding M
@@ -683,7 +673,7 @@
   have "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)"
     by (rule subdensity_real[OF measurable_snd Pxy Py]) (insert Py_nn Pxy_nn, auto simp: space_pair_measure)
   ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
-    by eventually_elim auto
+    by auto
 
   show "?M = ?R"
     unfolding M f_def
@@ -784,17 +774,10 @@
   assumes X: "distributed M S X Px"
   shows "AE x in S. RN_deriv S (density S Px) x = Px x"
 proof -
-  note D = distributed_measurable[OF X] distributed_borel_measurable[OF X]
-  interpret X: prob_space "distr M S X"
-    using D(1) by (rule prob_space_distr)
-
-  have sf: "sigma_finite_measure (distr M S X)" by standard
-  show ?thesis
-    using D
-    apply (subst eq_commute)
-    apply (intro RN_deriv_unique_sigma_finite)
-    apply (auto simp: distributed_distr_eq_density[symmetric, OF X] sf)
-    done
+  have "distributed M S X (RN_deriv S (density S Px))"
+    by (metis RN_derivI assms borel_measurable_RN_deriv distributed_def)
+  then show ?thesis
+    using assms distributed_unique by blast
 qed
 
 lemma (in information_space)
@@ -806,14 +789,9 @@
   note ae = distributed_RN_deriv[OF X]
   note distributed_real_measurable[OF nn X, measurable]
 
-  have ae_eq: "AE x in distr M MX X. log b (enn2real (RN_deriv MX (distr M MX X) x)) =
-    log b (f x)"
+  have ae_eq: "AE x in distr M MX X. log b (enn2real (RN_deriv MX (distr M MX X) x)) = log b (f x)"
     unfolding distributed_distr_eq_density[OF X]
-    apply (subst AE_density)
-    using D apply simp
-    using ae apply eventually_elim
-    apply auto
-    done
+    using D ae by (auto simp: AE_density)
 
   have int_eq: "(\<integral> x. f x * log b (f x) \<partial>MX) = (\<integral> x. log b (f x) \<partial>distr M MX X)"
     unfolding distributed_distr_eq_density[OF X]
@@ -1078,14 +1056,16 @@
     by (subst STP.nn_integral_snd[symmetric])
        (auto simp add: split_beta' ennreal_mult[symmetric] space_pair_measure intro!: nn_integral_cong)
   also have "\<dots> = (\<integral>\<^sup>+x. ennreal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)"
-    apply (rule nn_integral_cong_AE)
-    using aeX1 aeX2 aeX3 AE_space
-    apply eventually_elim
-  proof (case_tac x, simp add: space_pair_measure)
-    fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "a \<in> space T \<and> b \<in> space P"
-      "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) \<partial>S) = ennreal (Pz b)"
-    then show "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \<partial>S) = ennreal (Pyz (a, b))"
-      by (subst nn_integral_multc) (auto split: prod.split simp: ennreal_mult[symmetric])
+  proof -
+    have D: "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \<partial>S) = ennreal (Pyz (a, b))"
+      if "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "a \<in> space T \<and> b \<in> space P"
+        "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) \<partial>S) = ennreal (Pz b)"
+      for a b
+      using that by (subst nn_integral_multc) (auto split: prod.split simp: ennreal_mult[symmetric])
+    show ?thesis
+      apply (rule nn_integral_cong_AE)
+      using aeX1 aeX2 aeX3 
+      by (force simp add: space_pair_measure D)
   qed
   also have "\<dots> = 1"
     using Q.emeasure_space_1 distributed_distr_eq_density[OF Pyz]
@@ -1103,8 +1083,8 @@
     then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x = 0"
       by (intro nn_integral_0_iff_AE[THEN iffD1]) auto
     then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxyz x = 0"
-      using ae1 ae2 ae3 ae4 AE_space
-      by eventually_elim (auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure)
+      using ae1 ae2 ae3 ae4 
+      by (auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure)
     then have "(\<integral>\<^sup>+ x. ennreal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0"
       by (subst nn_integral_cong_AE[of _ "\<lambda>x. 0"]) auto
     with P.emeasure_space_1 show False
@@ -1112,11 +1092,7 @@
   qed
 
   have neg: "(\<integral>\<^sup>+ x. - ?f x \<partial>?P) = 0"
-    apply (rule nn_integral_0_iff_AE[THEN iffD2])
-    apply simp
-    apply (subst AE_density)
-    apply (auto simp: space_pair_measure ennreal_neg)
-    done
+    by (subst nn_integral_0_iff_AE) (auto simp: space_pair_measure ennreal_neg)
 
   have I3: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
     apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ Bochner_Integration.integrable_diff[OF I1 I2]])
@@ -1136,8 +1112,6 @@
     qed simp
     then have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) = (\<integral>x. ?f x \<partial>?P)"
       apply (rule nn_integral_eq_integral)
-      apply (subst AE_density)
-      apply simp
       apply (auto simp: space_pair_measure ennreal_neg)
       done
     with pos le1
@@ -1148,35 +1122,31 @@
   proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"])
     show "AE x in ?P. ?f x \<in> {0<..}"
       unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]]
-      using ae1 ae2 ae3 ae4 AE_space
-      by eventually_elim (auto simp: space_pair_measure less_le)
+      using ae1 ae2 ae3 ae4 
+      by (auto simp: space_pair_measure less_le)
     show "integrable ?P ?f"
       unfolding real_integrable_def
       using fin neg by (auto simp: split_beta')
-    show "integrable ?P (\<lambda>x. - log b (?f x))"
-      apply (subst integrable_real_density)
-      apply simp
-      apply (auto simp: space_pair_measure) []
-      apply simp
+    have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. Pxyz x * - log b (?f x))"
       apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3])
-      apply simp
-      apply simp
-      using ae1 ae2 ae3 ae4 AE_space
-      apply eventually_elim
+      using ae1 ae2 ae3 ae4 
       apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps
-        less_le space_pair_measure)
+          less_le space_pair_measure)
       done
+    then
+    show "integrable ?P (\<lambda>x. - log b (?f x))"
+      by (subst integrable_real_density) (auto simp: space_pair_measure) 
   qed (auto simp: b_gt_1 minus_log_convex)
   also have "\<dots> = conditional_mutual_information b S T P X Y Z"
     unfolding \<open>?eq\<close>
     apply (subst integral_real_density)
-    apply simp
-    apply (auto simp: space_pair_measure) []
-    apply simp
+       apply simp
+      apply (force simp: space_pair_measure) 
+     apply simp
     apply (intro integral_cong_AE)
     using ae1 ae2 ae3 ae4
-    apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps
-      space_pair_measure less_le)
+      apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps
+        space_pair_measure less_le)
     done
   finally show ?nonneg
     by simp
@@ -1276,10 +1246,8 @@
     using Pxyz Px Pyz by simp
   ultimately have I1: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
     apply (rule integrable_cong_AE_imp)
-    using ae1 ae4 AE_space
-    by eventually_elim
-       (insert Px_nn Pyz_nn Pxyz_nn,
-        auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff space_pair_measure less_le)
+    using ae1 ae4 Px_nn Pyz_nn Pxyz_nn 
+    by (auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff space_pair_measure less_le)
 
   have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)
     (\<lambda>x. Pxyz x * log b (Pxz (fst x, snd (snd x))) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pz (snd (snd x))))"
@@ -1292,20 +1260,14 @@
     by auto
   ultimately have I2: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
     apply (rule integrable_cong_AE_imp)
-    using ae1 ae2 ae3 ae4 AE_space
-    by eventually_elim
-       (insert Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn,
-         auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff less_le space_pair_measure)
+    using ae1 ae2 ae3 ae4  Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn
+    by (auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff less_le space_pair_measure)
 
   from ae I1 I2 show ?eq
-    unfolding conditional_mutual_information_def
-    apply (subst mi_eq)
-    apply (subst mutual_information_distr[OF S TP Px Px_nn Pyz Pyz_nn Pxyz Pxyz_nn])
-    apply simp
-    apply simp
-    apply (simp add: space_pair_measure)
+    unfolding conditional_mutual_information_def mi_eq
+    apply (subst mutual_information_distr[OF S TP Px Px_nn Pyz Pyz_nn Pxyz Pxyz_nn]; simp add: space_pair_measure)
     apply (subst Bochner_Integration.integral_diff[symmetric])
-    apply (auto intro!: integral_cong_AE simp: split_beta' simp del: Bochner_Integration.integral_diff)
+      apply (auto intro!: integral_cong_AE simp: split_beta' simp del: Bochner_Integration.integral_diff)
     done
 
   let ?P = "density (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) Pxyz"
@@ -1335,15 +1297,15 @@
     by (subst STP.nn_integral_snd[symmetric])
        (auto simp add: split_beta' ennreal_mult[symmetric] space_pair_measure intro!: nn_integral_cong)
   also have "\<dots> = (\<integral>\<^sup>+x. ennreal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)"
-    apply (rule nn_integral_cong_AE)
-    using aeX1 aeX2 aeX3 AE_space
-    apply eventually_elim
-  proof (case_tac x, simp add: space_pair_measure)
-    fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P"
-      "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) \<partial>S) = ennreal (Pz b)"
-    then show "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \<partial>S) = ennreal (Pyz (a, b))"
-      using Pyz_nn[of "(a,b)"]
+  proof -
+    have *: "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \<partial>S) = ennreal (Pyz (a, b))"
+      if "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P"
+        "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) \<partial>S) = ennreal (Pz b)" for a b
+      using Pyz_nn[of "(a,b)"] that
       by (subst nn_integral_multc) (auto simp: space_pair_measure ennreal_mult[symmetric])
+    show ?thesis
+      using aeX1 aeX2 aeX3 AE_space
+      by (force simp: * space_pair_measure intro: nn_integral_cong_AE)
   qed
   also have "\<dots> = 1"
     using Q.emeasure_space_1 Pyz_nn distributed_distr_eq_density[OF Pyz]
@@ -1362,9 +1324,8 @@
     then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x = 0"
       by (intro nn_integral_0_iff_AE[THEN iffD1]) auto
     then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxyz x = 0"
-      using ae1 ae2 ae3 ae4 AE_space
-      by eventually_elim
-         (insert Px_nn Pz_nn Pxz_nn Pyz_nn,
+      using ae1 ae2 ae3 ae4 
+      by  (insert Px_nn Pz_nn Pxz_nn Pyz_nn,
            auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure)
     then have "(\<integral>\<^sup>+ x. ennreal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0"
       by (subst nn_integral_cong_AE[of _ "\<lambda>x. 0"]) auto
@@ -1382,19 +1343,12 @@
   have I3: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
     apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ Bochner_Integration.integrable_diff[OF I1 I2]])
     using ae
-    apply (auto simp: split_beta')
-    done
+    by (auto simp: split_beta')
 
   have "- log b 1 \<le> - log b (integral\<^sup>L ?P ?f)"
   proof (intro le_imp_neg_le log_le[OF b_gt_1])
     have If: "integrable ?P ?f"
-      unfolding real_integrable_def
-    proof (intro conjI)
-      from neg show "(\<integral>\<^sup>+ x. - ?f x \<partial>?P) \<noteq> \<infinity>"
-        by simp
-      from fin show "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>"
-        by simp
-    qed simp
+      using neg fin by (force simp add: real_integrable_def)
     then have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) = (\<integral>x. ?f x \<partial>?P)"
       using Pz_nn Pxz_nn Pyz_nn Pxyz_nn
       by (intro nn_integral_eq_integral)
@@ -1407,25 +1361,19 @@
   proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"])
     show "AE x in ?P. ?f x \<in> {0<..}"
       unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]]
-      using ae1 ae2 ae3 ae4 AE_space
-      by eventually_elim (insert Pxyz_nn Pyz_nn Pz_nn Pxz_nn, auto simp: space_pair_measure less_le)
+      using ae1 ae2 ae3 ae4 
+      by (insert Pxyz_nn Pyz_nn Pz_nn Pxz_nn, auto simp: space_pair_measure less_le)
     show "integrable ?P ?f"
       unfolding real_integrable_def
       using fin neg by (auto simp: split_beta')
-    show "integrable ?P (\<lambda>x. - log b (?f x))"
-      using Pz_nn Pxz_nn Pyz_nn Pxyz_nn
-      apply (subst integrable_real_density)
-      apply simp
-      apply simp
-      apply simp
+    have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. Pxyz x * - log b (?f x))"
       apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3])
-      apply simp
-      apply simp
-      using ae1 ae2 ae3 ae4 AE_space
-      apply eventually_elim
-      apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff
-                        zero_less_divide_iff field_simps space_pair_measure less_le)
-      done
+      using Pz_nn Pxz_nn Pyz_nn Pxyz_nn ae2 ae3 ae4 
+      by (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff
+          zero_less_divide_iff field_simps space_pair_measure less_le)
+    then
+    show "integrable ?P (\<lambda>x. - log b (?f x))"
+      using Pxyz_nn by (auto simp: integrable_real_density)
   qed (auto simp: b_gt_1 minus_log_convex)
   also have "\<dots> = conditional_mutual_information b S T P X Y Z"
     unfolding \<open>?eq\<close>
@@ -1435,9 +1383,9 @@
     apply simp
     apply simp
     apply (intro integral_cong_AE)
-    using ae1 ae2 ae3 ae4 AE_space
+    using ae1 ae2 ae3 ae4 
     apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff
-                      field_simps space_pair_measure less_le)
+                      field_simps space_pair_measure less_le integral_cong_AE)
     done
   finally show ?nonneg
     by simp
@@ -1504,24 +1452,18 @@
   note sd = simple_distributedI[OF _ _ refl]
   note sp = simple_function_Pair
   show ?thesis
-   apply (rule conditional_mutual_information_generic_nonneg[OF sf[OF X] sf[OF Y] sf[OF Z]])
-   apply (rule simple_distributed[OF sd[OF X]])
-   apply simp
-   apply simp
-   apply (rule simple_distributed[OF sd[OF Z]])
-   apply simp
-   apply simp
-   apply (rule simple_distributed_joint[OF sd[OF sp[OF Y Z]]])
-   apply simp
-   apply simp
-   apply (rule simple_distributed_joint[OF sd[OF sp[OF X Z]]])
-   apply simp
-   apply simp
-   apply (rule simple_distributed_joint2[OF sd[OF sp[OF X sp[OF Y Z]]]])
-   apply simp
-   apply simp
-   apply (auto intro!: integrable_count_space simp: X Y Z simple_functionD)
-   done
+    apply (rule conditional_mutual_information_generic_nonneg[OF sf[OF X] sf[OF Y] sf[OF Z]])
+               apply (force intro: simple_distributed[OF sd[OF X]])
+              apply simp
+             apply (force intro: simple_distributed[OF sd[OF Z]])
+            apply simp
+           apply (force intro: simple_distributed_joint[OF sd[OF sp[OF Y Z]]])
+          apply simp
+         apply (force intro: simple_distributed_joint[OF sd[OF sp[OF X Z]]])
+        apply simp
+       apply (fastforce intro:  simple_distributed_joint2[OF sd[OF sp[OF X sp[OF Y Z]]]])
+      apply (auto intro!: integrable_count_space simp: X Y Z simple_functionD)
+    done
 qed
 
 subsection \<open>Conditional Entropy\<close>
@@ -1560,11 +1502,8 @@
   have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x)). Py (snd x) = enn2real (RN_deriv T (distr M T Y) (snd x))"
     unfolding AE_density[OF distributed_borel_measurable, OF Pxy]
     unfolding distributed_distr_eq_density[OF Py]
-    apply (rule ST.AE_pair_measure)
-    apply auto
     using distributed_RN_deriv[OF Py]
-    apply auto
-    done
+    by (force intro: ST.AE_pair_measure)
   ultimately
   have "conditional_entropy b S T X Y = - (\<integral>x. Pxy x * log b (Pxy x / Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
     unfolding conditional_entropy_def neg_equal_iff_equal
@@ -1613,17 +1552,14 @@
     using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'')
   ultimately have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Pxy x \<and> 0 \<le> Py (snd x) \<and>
     (Pxy x = 0 \<or> (Pxy x \<noteq> 0 \<longrightarrow> 0 < Pxy x \<and> 0 < Py (snd x)))"
-    using AE_space by eventually_elim (auto simp: space_pair_measure less_le)
+    by (auto simp: space_pair_measure less_le)
   then have ae: "AE x in S \<Otimes>\<^sub>M T.
      Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / Py (snd x))"
-    by eventually_elim (auto simp: log_simps field_simps b_gt_1)
+    by (auto simp: log_simps field_simps b_gt_1)
   have "conditional_entropy b S T X Y =
     - (\<integral>x. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
     unfolding conditional_entropy_generic_eq[OF S T Py Py_nn Pxy Pxy_nn, simplified] neg_equal_iff_equal
-    apply (intro integral_cong_AE)
-    using ae
-    apply auto
-    done
+    using ae by (force intro: integral_cong_AE)
   also have "\<dots> = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^sub>M T)) - - (\<integral>x.  Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
     by (simp add: Bochner_Integration.integral_diff[OF I1 I2])
   finally show ?thesis
@@ -1671,7 +1607,7 @@
   from Y show "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) =
     - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
     by (auto intro!: sum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite eq sum.If_cases split_beta')
-qed (insert Y XY, auto)
+qed (use Y XY in auto)
 
 lemma (in information_space) conditional_mutual_information_eq_conditional_entropy:
   assumes X: "simple_function M X" and Y: "simple_function M Y"
@@ -1891,7 +1827,7 @@
   have "0 \<le> mutual_information b S T X Y"
     by (rule mutual_information_nonneg') fact+
   also have "\<dots> = entropy b S X - conditional_entropy b S T X Y"
-    apply (rule mutual_information_eq_entropy_conditional_entropy')
+    apply (intro mutual_information_eq_entropy_conditional_entropy')
     using assms
     by (auto intro!: finite_entropy_integrable finite_entropy_distributed
       finite_entropy_integrable_transform[OF Px]
@@ -1933,14 +1869,22 @@
   have eq: "(\<lambda>x. ((f \<circ> X) x, X x)) ` space M = (\<lambda>x. (f x, x)) ` X ` space M" by auto
   have inj: "\<And>A. inj_on (\<lambda>x. (f x, x)) A"
     by (auto simp: inj_on_def)
-  show ?thesis
-    apply (subst entropy_chain_rule[symmetric, OF fX X])
-    apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] measure_nonneg refl]])
-    apply (subst entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]])
+
+  have "\<H>(X) = - (\<Sum>x\<in>X ` space M. prob (X -` {x} \<inter> space M) * log b (prob (X -` {x} \<inter> space M)))"
+    by (simp add: entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]])
+  also have "\<dots> =  - (\<Sum>x\<in>(\<lambda>x. ((f \<circ> X) x, X x)) ` space M.
+                       prob ((\<lambda>x. ((f \<circ> X) x, X x)) -` {x} \<inter> space M) *
+                       log b (prob ((\<lambda>x. ((f \<circ> X) x, X x)) -` {x} \<inter> space M)))"
     unfolding eq
     apply (subst sum.reindex[OF inj])
     apply (auto intro!: sum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
     done
+  also have "... = \<H>(\<lambda>x. ((f \<circ> X) x, X x))"
+    using entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] measure_nonneg refl]]
+    by fastforce
+  also have "\<dots> = \<H>(f \<circ> X) + \<H>(X|f \<circ> X)"
+    using X entropy_chain_rule by blast
+  finally show ?thesis .
 qed
 
 corollary (in information_space) entropy_data_processing:
--- a/src/HOL/Probability/Levy.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Probability/Levy.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -32,13 +32,8 @@
       by (simp add: norm_divide norm_mult)
     also have "cmod ?one / abs t + cmod ?two / abs t \<le>
         ((- (a * t))^2 / 2) / abs t + ((- (b * t))^2 / 2) / abs t"
-      apply (rule add_mono)
-      apply (rule divide_right_mono)
-      using iexp_approx1 [of "-(t * a)" 1] apply (simp add: field_simps eval_nat_numeral)
-      apply force
-      apply (rule divide_right_mono)
-      using iexp_approx1 [of "-(t * b)" 1] apply (simp add: field_simps eval_nat_numeral)
-      by force
+      using iexp_approx1 [of "-(t * _)" 1]
+      by (intro add_mono divide_right_mono abs_ge_zero) (auto simp: field_simps eval_nat_numeral)
     also have "\<dots> = a^2 / 2 * abs t + b^2 / 2 * abs t"
       using \<open>t \<noteq> 0\<close> apply (case_tac "t \<ge> 0", simp add: field_simps power2_eq_square)
       using \<open>t \<noteq> 0\<close> by (subst (1 2) abs_of_neg, auto simp add: field_simps power2_eq_square)
@@ -88,6 +83,8 @@
     assume "T \<ge> 0"
     let ?f' = "\<lambda>(t, x). indicator {-T<..<T} t *\<^sub>R ?f t x"
     { fix x
+      have int: "interval_lebesgue_integrable lborel (ereal 0) (ereal T) (\<lambda>t. 2 * (sin (t * (x-w)) / t))" for w
+        using integrable_sinc' interval_lebesgue_integrable_mult_right by blast
       have 1: "complex_interval_lebesgue_integrable lborel u v (\<lambda>t. ?f t x)" for u v :: real
         using Levy_Inversion_aux2[of "x - b" "x - a"]
         apply (simp add: interval_lebesgue_integrable_def set_integrable_def del: times_divide_eq_left)
@@ -110,21 +107,18 @@
       also have "\<dots> = (CLBINT t=(0::real)..T. complex_of_real(
           2 * (sin (t * (x - a)) / t) - 2 * (sin (t * (x - b)) / t)))"
         using \<open>T \<ge> 0\<close>
-        apply (intro interval_integral_cong)
-        apply (simp add: field_simps cis.ctr Im_divide Re_divide Im_exp Re_exp complex_eq_iff)
-        unfolding minus_diff_eq[symmetric, of "y * x" "y * a" for y a] sin_minus cos_minus
-        apply (simp add: field_simps power2_eq_square)
-        done
+        by (intro interval_integral_cong) (simp add: divide_simps Im_divide Re_divide Im_exp Re_exp complex_eq_iff)
       also have "\<dots> = complex_of_real (LBINT t=(0::real)..T.
           2 * (sin (t * (x - a)) / t) - 2 * (sin (t * (x - b)) / t))"
         by (rule interval_lebesgue_integral_of_real)
+      also have "\<dots> = complex_of_real ((LBINT t=ereal 0..ereal T. 2 * (sin (t * (x - a)) / t)) -
+                                       (LBINT t=ereal 0..ereal T. 2 * (sin (t * (x - b)) / t)))"
+        unfolding interval_lebesgue_integral_diff
+        using int by auto
       also have "\<dots> = complex_of_real (2 * (sgn (x - a) * Si (T * abs (x - a)) -
           sgn (x - b) * Si (T * abs (x - b))))"
-        apply (subst interval_lebesgue_integral_diff)
-        apply (rule interval_lebesgue_integrable_mult_right, rule integrable_sinc')+
-        apply (subst interval_lebesgue_integral_mult_right)+
-        apply (simp add: zero_ereal_def[symmetric] LBINT_I0c_sin_scale_divide[OF \<open>T \<ge> 0\<close>])
-        done
+        unfolding interval_lebesgue_integral_mult_right
+        by (simp add: zero_ereal_def[symmetric] LBINT_I0c_sin_scale_divide[OF \<open>T \<ge> 0\<close>])
       finally have "(CLBINT t. ?f' (t, x)) =
           2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))" .
     } note main_eq = this
@@ -287,22 +281,15 @@
     assume "u > 0" and "x \<noteq> 0"
     hence "(CLBINT t:{-u..u}. 1 - iexp (t * x)) = (CLBINT t=-u..u. 1 - iexp (t * x))"
       by (subst interval_integral_Icc, auto)
-    also have "\<dots> = (CLBINT t=-u..0. 1 - iexp (t * x)) + (CLBINT t=0..u. 1 - iexp (t * x))"
+    also have "\<dots> = (CLBINT t=-u..ereal 0. 1 - iexp (t * x)) + (CLBINT t= ereal 0..u. 1 - iexp (t * x))"
       using \<open>u > 0\<close>
-      apply (subst interval_integral_sum)
-      apply (simp add: min_absorb1 min_absorb2 max_absorb1 max_absorb2)
-      apply (rule interval_integrable_isCont)
-      apply auto
-      done
+      by (subst interval_integral_sum; force simp add: interval_integrable_isCont)
     also have "\<dots> = (CLBINT t=ereal 0..u. 1 - iexp (t * -x)) + (CLBINT t=ereal 0..u. 1 - iexp (t * x))"
-      apply (subgoal_tac "0 = ereal 0", erule ssubst)
       by (subst interval_integral_reflect, auto)
+    also have "... = CLBINT xa=ereal 0..ereal u. 1 - iexp (xa * - x) + (1 - iexp (xa * x))"
+      by (subst interval_lebesgue_integral_add (2) [symmetric]) (auto simp: interval_integrable_isCont)
     also have "\<dots> = (LBINT t=ereal 0..u. 2 - 2 * cos (t * x))"
-      apply (subst interval_lebesgue_integral_add (2) [symmetric])
-      apply ((rule interval_integrable_isCont, auto)+) [2]
-      unfolding exp_Euler cos_of_real
-      apply (simp add: of_real_mult interval_lebesgue_integral_of_real[symmetric])
-      done
+      unfolding exp_Euler cos_of_real by (simp flip: interval_lebesgue_integral_of_real)
     also have "\<dots> = 2 * u - 2 * sin (u * x) / x"
       by (subst interval_lebesgue_integral_diff)
          (auto intro!: interval_integrable_isCont
@@ -375,13 +362,11 @@
   proof -
     fix \<epsilon> :: real
     assume "\<epsilon> > 0"
-    note M'.isCont_char [of 0]
-    hence "\<exists>d>0. \<forall>t. abs t < d \<longrightarrow> cmod (char M' t - 1) < \<epsilon> / 4"
-      apply (subst (asm) continuous_at_eps_delta)
-      apply (drule_tac x = "\<epsilon> / 4" in spec)
-      using \<open>\<epsilon> > 0\<close> by (auto simp add: dist_real_def dist_complex_def M'.char_zero)
-    then obtain d where "d > 0 \<and> (\<forall>t. (abs t < d \<longrightarrow> cmod (char M' t - 1) < \<epsilon> / 4))" ..
-    hence d0: "d > 0" and d1: "\<And>t. abs t < d \<Longrightarrow> cmod (char M' t - 1) < \<epsilon> / 4" by auto
+    with M'.isCont_char [of 0]
+    obtain d where d0: "d>0" and "\<forall>x'. dist x' 0 < d \<longrightarrow> dist (char M' x') (char M' 0) < \<epsilon>/4"
+      unfolding continuous_at_eps_delta by (metis \<open>0 < \<epsilon>\<close> divide_pos_pos zero_less_numeral)
+    then have d1: "\<And>t. abs t < d \<Longrightarrow> cmod (char M' t - 1) < \<epsilon> / 4"
+      by (simp add: M'.char_zero dist_norm)
     have 1: "\<And>x. cmod (1 - char M' x) \<le> 2"
       by (rule order_trans [OF norm_triangle_ineq4], auto simp add: M'.cmod_char_le_1)
     then have 2: "\<And>u v. complex_set_integrable lborel {u..v} (\<lambda>x. 1 - char M' x)"
@@ -395,15 +380,18 @@
       using integral_norm_bound[of _ "\<lambda>x. indicator {u..v} x *\<^sub>R (1 - char M' x)" for u v] by simp
     also have 4: "\<dots> \<le> LBINT t:{-d/2..d/2}. \<epsilon> / 4"
       unfolding set_lebesgue_integral_def
-      apply (rule integral_mono [OF 3])
-       apply (simp add: emeasure_lborel_Icc_eq)
-      apply (case_tac "x \<in> {-d/2..d/2}")
-       apply auto
-      apply (subst norm_minus_commute)
-      apply (rule less_imp_le)
-      apply (rule d1 [simplified])
-      using d0 apply auto
-      done
+    proof (rule integral_mono [OF 3])
+
+      show "indicat_real {- d / 2..d / 2} x *\<^sub>R cmod (1 - char M' x)
+         \<le> indicat_real {- d / 2..d / 2} x *\<^sub>R (\<epsilon> / 4)"
+        if "x \<in> space lborel" for x
+      proof (cases "x \<in> {-d/2..d/2}")
+        case True
+        show ?thesis
+          using d0 d1 that True [simplified]
+          by (smt (verit, best) field_sum_of_halves minus_diff_eq norm_minus_cancel indicator_pos_le scaleR_left_mono)
+      qed auto
+    qed (simp add: emeasure_lborel_Icc_eq)
     also from d0 4 have "\<dots> = d * \<epsilon> / 4"
       unfolding set_lebesgue_integral_def by simp
     finally have bound: "cmod (CLBINT t:{-d/2..d/2}. 1 - char M' t) \<le> d * \<epsilon> / 4" .
@@ -450,14 +438,7 @@
         apply (subst Mn.borel_UNIV [symmetric])
         by (subst Mn.prob_compl, auto)
       also have "UNIV - {x. abs x \<ge> 2 / (d / 2)} = {x. -(4 / d) < x \<and> x < (4 / d)}"
-        using d0 apply (auto simp add: field_simps)
-        (* very annoying -- this should be automatic *)
-        apply (case_tac "x \<ge> 0", auto simp add: field_simps)
-        apply (subgoal_tac "0 \<le> x * d", arith, rule mult_nonneg_nonneg, auto)
-        apply (case_tac "x \<ge> 0", auto simp add: field_simps)
-        apply (subgoal_tac "x * d \<le> 0", arith)
-        apply (rule mult_nonpos_nonneg, auto)
-        by (case_tac "x \<ge> 0", auto simp add: field_simps)
+        using d0  by (simp add: set_eq_iff divide_simps abs_if) (smt (verit, best) mult_less_0_iff)
       finally have "measure (M n) {x. -(4 / d) < x \<and> x < (4 / d)} > 1 - \<epsilon>"
         by auto
     } note 6 = this
@@ -470,8 +451,7 @@
       hence "(\<lambda>k. measure (M n) {- real k<..real k}) \<longlonglongrightarrow> 1"
         using Mn.prob_space unfolding * Mn.borel_UNIV by simp
       hence "eventually (\<lambda>k. measure (M n) {- real k<..real k} > 1 - \<epsilon>) sequentially"
-        apply (elim order_tendstoD (1))
-        using \<open>\<epsilon> > 0\<close> by auto
+        using \<open>\<epsilon> > 0\<close> order_tendstoD by fastforce
     } note 7 = this
     { fix n :: nat
       have "eventually (\<lambda>k. \<forall>m < n. measure (M m) {- real k<..real k} > 1 - \<epsilon>) sequentially"
@@ -490,20 +470,18 @@
     hence K: "\<And>m. m < N \<Longrightarrow> 1 - \<epsilon> < Sigma_Algebra.measure (M m) {- real K<..real K}"
       by auto
     let ?K' = "max K (4 / d)"
-    have "-?K' < ?K' \<and> (\<forall>n. 1 - \<epsilon> < measure (M n) {-?K'<..?K'})"
-      using d0 apply auto
-      apply (rule max.strict_coboundedI2, auto)
-    proof -
-      fix n
-      show " 1 - \<epsilon> < measure (M n) {- max (real K) (4 / d)<..max (real K) (4 / d)}"
-        apply (case_tac "n < N")
-        apply (rule order_less_le_trans)
-        apply (erule K)
-        apply (rule Mn.finite_measure_mono, auto)
-        apply (rule order_less_le_trans)
-        apply (rule 6, erule leI)
-        by (rule Mn.finite_measure_mono, auto)
+    have "1 - \<epsilon> < measure (M n) {- max (real K) (4 / d)<..max (real K) (4 / d)}" for n
+    proof (cases "n < N")
+      case True
+      then show ?thesis
+        by (force intro: order_less_le_trans [OF K Mn.finite_measure_mono])
+    next
+      case False
+      then show ?thesis
+        by (force intro: order_less_le_trans [OF 6 Mn.finite_measure_mono])
     qed
+    then have "-?K' < ?K' \<and> (\<forall>n. 1 - \<epsilon> < measure (M n) {-?K'<..?K'})"
+      using d0 by (simp add: less_max_iff_disj minus_less_iff)
     thus "\<exists>a b. a < b \<and> (\<forall>n. 1 - \<epsilon> < measure (M n) {a<..b})" by (intro exI)
   qed
   have tight: "tight M"
--- a/src/HOL/Probability/SPMF.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Probability/SPMF.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -11,20 +11,20 @@
 subsection \<open>Auxiliary material\<close>
 
 lemma cSUP_singleton [simp]: "(SUP x\<in>{x}. f x :: _ :: conditionally_complete_lattice) = f x"
-by (metis cSup_singleton image_empty image_insert)
+  by (metis cSup_singleton image_empty image_insert)
 
 subsubsection \<open>More about extended reals\<close>
 
 lemma [simp]:
   shows ennreal_max_0: "ennreal (max 0 x) = ennreal x"
-  and ennreal_max_0': "ennreal (max x 0) = ennreal x"
-by(simp_all add: max_def ennreal_eq_0_iff)
+    and ennreal_max_0': "ennreal (max x 0) = ennreal x"
+  by(simp_all add: max_def ennreal_eq_0_iff)
 
 lemma e2ennreal_0 [simp]: "e2ennreal 0 = 0"
-by(simp add: zero_ennreal_def)
+  by(simp add: zero_ennreal_def)
 
 lemma enn2real_bot [simp]: "enn2real \<bottom> = 0"
-by(simp add: bot_ennreal_def)
+  by(simp add: bot_ennreal_def)
 
 lemma continuous_at_ennreal[continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. ennreal (f x))"
   unfolding continuous_def by auto
@@ -42,24 +42,24 @@
 
 lemma ennreal_SUP:
   "\<lbrakk> (SUP a\<in>A. ennreal (f a)) \<noteq> \<top>; A \<noteq> {} \<rbrakk> \<Longrightarrow> ennreal (SUP a\<in>A. f a) = (SUP a\<in>A. ennreal (f a))"
-using ennreal_Sup[of "f ` A"] by (auto simp add: image_comp)
+  using ennreal_Sup[of "f ` A"] by (auto simp: image_comp)
 
 lemma ennreal_lt_0: "x < 0 \<Longrightarrow> ennreal x = 0"
-by(simp add: ennreal_eq_0_iff)
+  by(simp add: ennreal_eq_0_iff)
 
 subsubsection \<open>More about \<^typ>\<open>'a option\<close>\<close>
 
 lemma None_in_map_option_image [simp]: "None \<in> map_option f ` A \<longleftrightarrow> None \<in> A"
-by auto
+  by auto
 
 lemma Some_in_map_option_image [simp]: "Some x \<in> map_option f ` A \<longleftrightarrow> (\<exists>y. x = f y \<and> Some y \<in> A)"
-by(auto intro: rev_image_eqI dest: sym)
+  by (smt (verit, best) imageE imageI map_option_eq_Some)
 
 lemma case_option_collapse: "case_option x (\<lambda>_. x) = (\<lambda>_. x)"
-by(simp add: fun_eq_iff split: option.split)
+  by(simp add: fun_eq_iff split: option.split)
 
 lemma case_option_id: "case_option None Some = id"
-by(rule ext)(simp split: option.split)
+  by(rule ext)(simp split: option.split)
 
 inductive ord_option :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> bool"
   for ord :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
@@ -78,65 +78,66 @@
   "ord_option (=) (Some x) y"
 
 lemma ord_option_reflI: "(\<And>y. y \<in> set_option x \<Longrightarrow> ord y y) \<Longrightarrow> ord_option ord x x"
-by(cases x) simp_all
+  by(cases x) simp_all
 
 lemma reflp_ord_option: "reflp ord \<Longrightarrow> reflp (ord_option ord)"
-by(simp add: reflp_def ord_option_reflI)
+  by(simp add: reflp_def ord_option_reflI)
 
 lemma ord_option_trans:
   "\<lbrakk> ord_option ord x y; ord_option ord y z;
     \<And>a b c. \<lbrakk> a \<in> set_option x; b \<in> set_option y; c \<in> set_option z; ord a b; ord b c \<rbrakk> \<Longrightarrow> ord a c \<rbrakk>
   \<Longrightarrow> ord_option ord x z"
-by(auto elim!: ord_option.cases)
+  by(auto elim!: ord_option.cases)
 
 lemma transp_ord_option: "transp ord \<Longrightarrow> transp (ord_option ord)"
-unfolding transp_def by(blast intro: ord_option_trans)
+  unfolding transp_def by(blast intro: ord_option_trans)
 
 lemma antisymp_ord_option: "antisymp ord \<Longrightarrow> antisymp (ord_option ord)"
-by(auto intro!: antisympI elim!: ord_option.cases dest: antisympD)
+  by(auto intro!: antisympI elim!: ord_option.cases dest: antisympD)
 
 lemma ord_option_chainD:
   "Complete_Partial_Order.chain (ord_option ord) Y
   \<Longrightarrow> Complete_Partial_Order.chain ord {x. Some x \<in> Y}"
-by(rule chainI)(auto dest: chainD)
+  by(rule chainI)(auto dest: chainD)
 
 definition lub_option :: "('a set \<Rightarrow> 'b) \<Rightarrow> 'a option set \<Rightarrow> 'b option"
-where "lub_option lub Y = (if Y \<subseteq> {None} then None else Some (lub {x. Some x \<in> Y}))"
+  where "lub_option lub Y = (if Y \<subseteq> {None} then None else Some (lub {x. Some x \<in> Y}))"
 
 lemma map_lub_option: "map_option f (lub_option lub Y) = lub_option (f \<circ> lub) Y"
-by(simp add: lub_option_def)
+  by(simp add: lub_option_def)
 
 lemma lub_option_upper:
   assumes "Complete_Partial_Order.chain (ord_option ord) Y" "x \<in> Y"
-  and lub_upper: "\<And>Y x. \<lbrakk> Complete_Partial_Order.chain ord Y; x \<in> Y \<rbrakk> \<Longrightarrow> ord x (lub Y)"
+    and lub_upper: "\<And>Y x. \<lbrakk> Complete_Partial_Order.chain ord Y; x \<in> Y \<rbrakk> \<Longrightarrow> ord x (lub Y)"
   shows "ord_option ord x (lub_option lub Y)"
-using assms(1-2)
-by(cases x)(auto simp add: lub_option_def intro: lub_upper[OF ord_option_chainD])
+  using assms(1-2)
+  by(cases x)(auto simp: lub_option_def intro: lub_upper[OF ord_option_chainD])
 
 lemma lub_option_least:
   assumes Y: "Complete_Partial_Order.chain (ord_option ord) Y"
-  and upper: "\<And>x. x \<in> Y \<Longrightarrow> ord_option ord x y"
+    and upper: "\<And>x. x \<in> Y \<Longrightarrow> ord_option ord x y"
   assumes lub_least: "\<And>Y y. \<lbrakk> Complete_Partial_Order.chain ord Y; \<And>x. x \<in> Y \<Longrightarrow> ord x y \<rbrakk> \<Longrightarrow> ord (lub Y) y"
   shows "ord_option ord (lub_option lub Y) y"
-using Y
-by(cases y)(auto 4 3 simp add: lub_option_def intro: lub_least[OF ord_option_chainD] dest: upper)
+  using Y
+  by(cases y)(auto 4 3 simp add: lub_option_def intro: lub_least[OF ord_option_chainD] dest: upper)
 
 lemma lub_map_option: "lub_option lub (map_option f ` Y) = lub_option (lub \<circ> (`) f) Y"
-apply(auto simp add: lub_option_def)
-apply(erule notE)
-apply(rule arg_cong[where f=lub])
-apply(auto intro: rev_image_eqI dest: sym)
-done
+proof -
+  have "\<And>u y. \<lbrakk>Some u \<in> Y; y \<in> Y\<rbrakk> \<Longrightarrow> {f y |y. Some y \<in> Y} = f ` {x. Some x \<in> Y}"
+    by blast
+  then show ?thesis
+    by (auto simp: lub_option_def)
+qed
 
 lemma ord_option_mono: "\<lbrakk> ord_option A x y; \<And>x y. A x y \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> ord_option B x y"
-by(auto elim: ord_option.cases)
+  by(auto elim: ord_option.cases)
 
 lemma ord_option_mono' [mono]:
   "(\<And>x y. A x y \<longrightarrow> B x y) \<Longrightarrow> ord_option A x y \<longrightarrow> ord_option B x y"
-by(blast intro: ord_option_mono)
+  by(blast intro: ord_option_mono)
 
 lemma ord_option_compp: "ord_option (A OO B) = ord_option A OO ord_option B"
-by(auto simp add: fun_eq_iff elim!: ord_option.cases intro: ord_option.intros)
+  by(auto simp: fun_eq_iff elim!: ord_option.cases intro: ord_option.intros)
 
 lemma ord_option_inf: "inf (ord_option A) (ord_option B) = ord_option (inf A B)" (is "?lhs = ?rhs")
 proof(rule antisym)
@@ -144,13 +145,13 @@
 qed(auto elim: ord_option_mono)
 
 lemma ord_option_map2: "ord_option ord x (map_option f y) = ord_option (\<lambda>x y. ord x (f y)) x y"
-by(auto elim: ord_option.cases)
+  by(auto elim: ord_option.cases)
 
 lemma ord_option_map1: "ord_option ord (map_option f x) y = ord_option (\<lambda>x y. ord (f x) y) x y"
-by(auto elim: ord_option.cases)
+  by(auto elim: ord_option.cases)
 
 lemma option_ord_Some1_iff: "option_ord (Some x) y \<longleftrightarrow> y = Some x"
-by(auto simp add: flat_ord_def)
+  by(auto simp: flat_ord_def)
 
 subsubsection \<open>A relator for sets that treats sets like predicates\<close>
 
@@ -158,46 +159,46 @@
 begin
 
 definition rel_pred :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool"
-where "rel_pred R A B = (R ===> (=)) (\<lambda>x. x \<in> A) (\<lambda>y. y \<in> B)"
+  where "rel_pred R A B = (R ===> (=)) (\<lambda>x. x \<in> A) (\<lambda>y. y \<in> B)"
 
 lemma rel_predI: "(R ===> (=)) (\<lambda>x. x \<in> A) (\<lambda>y. y \<in> B) \<Longrightarrow> rel_pred R A B"
-by(simp add: rel_pred_def)
+  by(simp add: rel_pred_def)
 
 lemma rel_predD: "\<lbrakk> rel_pred R A B; R x y \<rbrakk> \<Longrightarrow> x \<in> A \<longleftrightarrow> y \<in> B"
-by(simp add: rel_pred_def rel_fun_def)
+  by(simp add: rel_pred_def rel_fun_def)
 
 lemma Collect_parametric: "((A ===> (=)) ===> rel_pred A) Collect Collect"
   \<comment> \<open>Declare this rule as @{attribute transfer_rule} only locally
       because it blows up the search space for @{method transfer}
       (in combination with @{thm [source] Collect_transfer})\<close>
-by(simp add: rel_funI rel_predI)
+  by(simp add: rel_funI rel_predI)
 
 end
 
 subsubsection \<open>Monotonicity rules\<close>
 
 lemma monotone_gfp_eadd1: "monotone (\<ge>) (\<ge>) (\<lambda>x. x + y :: enat)"
-by(auto intro!: monotoneI)
+  by(auto intro!: monotoneI)
 
 lemma monotone_gfp_eadd2: "monotone (\<ge>) (\<ge>) (\<lambda>y. x + y :: enat)"
-by(auto intro!: monotoneI)
+  by(auto intro!: monotoneI)
 
 lemma mono2mono_gfp_eadd[THEN gfp.mono2mono2, cont_intro, simp]:
   shows monotone_eadd: "monotone (rel_prod (\<ge>) (\<ge>)) (\<ge>) (\<lambda>(x, y). x + y :: enat)"
-by(simp add: monotone_gfp_eadd1 monotone_gfp_eadd2)
+  by(simp add: monotone_gfp_eadd1 monotone_gfp_eadd2)
 
 lemma eadd_gfp_partial_function_mono [partial_function_mono]:
   "\<lbrakk> monotone (fun_ord (\<ge>)) (\<ge>) f; monotone (fun_ord (\<ge>)) (\<ge>) g \<rbrakk>
   \<Longrightarrow> monotone (fun_ord (\<ge>)) (\<ge>) (\<lambda>x. f x + g x :: enat)"
-by(rule mono2mono_gfp_eadd)
+  by(rule mono2mono_gfp_eadd)
 
 lemma mono2mono_ereal[THEN lfp.mono2mono]:
   shows monotone_ereal: "monotone (\<le>) (\<le>) ereal"
-by(rule monotoneI) simp
+  by(rule monotoneI) simp
 
 lemma mono2mono_ennreal[THEN lfp.mono2mono]:
   shows monotone_ennreal: "monotone (\<le>) (\<le>) ennreal"
-by(rule monotoneI)(simp add: ennreal_leI)
+  by(rule monotoneI)(simp add: ennreal_leI)
 
 subsubsection \<open>Bijections\<close>
 
@@ -207,20 +208,17 @@
   shows "\<exists>f. bij_betw f A B \<and> (\<forall>x\<in>A. R x (f x))"
 proof -
   from assms obtain f where f: "\<And>x. x \<in> A \<Longrightarrow> R x (f x)" and B: "\<And>x. x \<in> A \<Longrightarrow> f x \<in> B"
-    apply(atomize_elim)
-    apply(fold all_conj_distrib)
-    apply(subst choice_iff[symmetric])
-    apply(auto dest: rel_setD1)
-    done
-  have "inj_on f A" by(rule inj_onI)(auto dest!: f dest: bi_uniqueDl[OF unique])
+    by (metis bi_unique_rel_set_lemma image_eqI)
+  have "inj_on f A"
+    by (metis (no_types, lifting) bi_unique_def f inj_on_def unique) 
   moreover have "f ` A = B" using rel
-    by(auto 4 3 intro: B dest: rel_setD2 f bi_uniqueDr[OF unique])
+    by (smt (verit) bi_unique_def bi_unique_rel_set_lemma f image_cong unique)
   ultimately have "bij_betw f A B" unfolding bij_betw_def ..
   thus ?thesis using f by blast
 qed
 
 lemma bij_betw_rel_setD: "bij_betw f A B \<Longrightarrow> rel_set (\<lambda>x y. y = f x) A B"
-by(rule rel_setI)(auto dest: bij_betwE bij_betw_imp_surj_on[symmetric])
+  by(rule rel_setI)(auto dest: bij_betwE bij_betw_imp_surj_on[symmetric])
 
 subsection \<open>Subprobability mass function\<close>
 
@@ -228,33 +226,29 @@
 translations (type) "'a spmf" \<leftharpoondown> (type) "'a option pmf"
 
 definition measure_spmf :: "'a spmf \<Rightarrow> 'a measure"
-where "measure_spmf p = distr (restrict_space (measure_pmf p) (range Some)) (count_space UNIV) the"
+  where "measure_spmf p = distr (restrict_space (measure_pmf p) (range Some)) (count_space UNIV) the"
 
 abbreviation spmf :: "'a spmf \<Rightarrow> 'a \<Rightarrow> real"
-where "spmf p x \<equiv> pmf p (Some x)"
+  where "spmf p x \<equiv> pmf p (Some x)"
 
 lemma space_measure_spmf: "space (measure_spmf p) = UNIV"
-by(simp add: measure_spmf_def)
+  by(simp add: measure_spmf_def)
 
 lemma sets_measure_spmf [simp, measurable_cong]: "sets (measure_spmf p) = sets (count_space UNIV)"
-by(simp add: measure_spmf_def)
+  by(simp add: measure_spmf_def)
 
 lemma measure_spmf_not_bot [simp]: "measure_spmf p \<noteq> \<bottom>"
-proof
-  assume "measure_spmf p = \<bottom>"
-  hence "space (measure_spmf p) = space \<bottom>" by simp
-  thus False by(simp add: space_measure_spmf)
-qed
+  by (metis empty_not_UNIV space_bot space_measure_spmf)
 
 lemma measurable_the_measure_pmf_Some [measurable, simp]:
   "the \<in> measurable (restrict_space (measure_pmf p) (range Some)) (count_space UNIV)"
-by(auto simp add: measurable_def sets_restrict_space space_restrict_space integral_restrict_space)
+  by(auto simp: measurable_def sets_restrict_space space_restrict_space integral_restrict_space)
 
 lemma measurable_spmf_measure1[simp]: "measurable (measure_spmf M) N = UNIV \<rightarrow> space N"
-by(auto simp: measurable_def space_measure_spmf)
+  by(auto simp: measurable_def space_measure_spmf)
 
 lemma measurable_spmf_measure2[simp]: "measurable N (measure_spmf M) = measurable N (count_space UNIV)"
-by(intro measurable_cong_sets) simp_all
+  by(intro measurable_cong_sets) simp_all
 
 lemma subprob_space_measure_spmf [simp, intro!]: "subprob_space (measure_spmf p)"
 proof
@@ -263,36 +257,37 @@
 qed(simp add: space_measure_spmf)
 
 interpretation measure_spmf: subprob_space "measure_spmf p" for p
-by(rule subprob_space_measure_spmf)
+  by(rule subprob_space_measure_spmf)
 
 lemma finite_measure_spmf [simp]: "finite_measure (measure_spmf p)"
-by unfold_locales
+  by unfold_locales
 
 lemma spmf_conv_measure_spmf: "spmf p x = measure (measure_spmf p) {x}"
-by(auto simp add: measure_spmf_def measure_distr measure_restrict_space pmf.rep_eq space_restrict_space intro: arg_cong2[where f=measure])
+  by(auto simp: measure_spmf_def measure_distr measure_restrict_space pmf.rep_eq space_restrict_space intro: arg_cong2[where f=measure])
 
 lemma emeasure_measure_spmf_conv_measure_pmf:
   "emeasure (measure_spmf p) A = emeasure (measure_pmf p) (Some ` A)"
-by(auto simp add: measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space intro: arg_cong2[where f=emeasure])
+  by(auto simp: measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space intro: arg_cong2[where f=emeasure])
 
 lemma measure_measure_spmf_conv_measure_pmf:
   "measure (measure_spmf p) A = measure (measure_pmf p) (Some ` A)"
-using emeasure_measure_spmf_conv_measure_pmf[of p A]
-by(simp add: measure_spmf.emeasure_eq_measure measure_pmf.emeasure_eq_measure)
+  using emeasure_measure_spmf_conv_measure_pmf[of p A]
+  by(simp add: measure_spmf.emeasure_eq_measure measure_pmf.emeasure_eq_measure)
 
 lemma emeasure_spmf_map_pmf_Some [simp]:
   "emeasure (measure_spmf (map_pmf Some p)) A = emeasure (measure_pmf p) A"
-by(auto simp add: measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space intro: arg_cong2[where f=emeasure])
+  by(auto simp: measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space intro: arg_cong2[where f=emeasure])
 
 lemma measure_spmf_map_pmf_Some [simp]:
   "measure (measure_spmf (map_pmf Some p)) A = measure (measure_pmf p) A"
-using emeasure_spmf_map_pmf_Some[of p A] by(simp add: measure_spmf.emeasure_eq_measure measure_pmf.emeasure_eq_measure)
+  using emeasure_spmf_map_pmf_Some[of p A] by(simp add: measure_spmf.emeasure_eq_measure measure_pmf.emeasure_eq_measure)
 
 lemma nn_integral_measure_spmf: "(\<integral>\<^sup>+ x. f x \<partial>measure_spmf p) = \<integral>\<^sup>+ x. ennreal (spmf p x) * f x \<partial>count_space UNIV"
   (is "?lhs = ?rhs")
 proof -
   have "?lhs = \<integral>\<^sup>+ x. pmf p x * f (the x) \<partial>count_space (range Some)"
-    by(simp add: measure_spmf_def nn_integral_distr nn_integral_restrict_space nn_integral_measure_pmf nn_integral_count_space_indicator ac_simps times_ereal.simps(1)[symmetric] del: times_ereal.simps(1))
+    by(simp add: measure_spmf_def nn_integral_distr nn_integral_restrict_space nn_integral_measure_pmf nn_integral_count_space_indicator ac_simps 
+        flip: times_ereal.simps [symmetric])
   also have "\<dots> = \<integral>\<^sup>+ x. ennreal (spmf p (the x)) * f (the x) \<partial>count_space (range Some)"
     by(rule nn_integral_cong) auto
   also have "\<dots> = \<integral>\<^sup>+ x. spmf p (the (Some x)) * f (the (Some x)) \<partial>count_space UNIV"
@@ -312,69 +307,65 @@
 qed
 
 lemma emeasure_spmf_single: "emeasure (measure_spmf p) {x} = spmf p x"
-by(simp add: measure_spmf.emeasure_eq_measure spmf_conv_measure_spmf)
+  by(simp add: measure_spmf.emeasure_eq_measure spmf_conv_measure_spmf)
 
 lemma measurable_measure_spmf[measurable]:
   "(\<lambda>x. measure_spmf (M x)) \<in> measurable (count_space UNIV) (subprob_algebra (count_space UNIV))"
-by (auto simp: space_subprob_algebra)
+  by (auto simp: space_subprob_algebra)
 
 lemma nn_integral_measure_spmf_conv_measure_pmf:
   assumes [measurable]: "f \<in> borel_measurable (count_space UNIV)"
   shows "nn_integral (measure_spmf p) f = nn_integral (restrict_space (measure_pmf p) (range Some)) (f \<circ> the)"
-by(simp add: measure_spmf_def nn_integral_distr o_def)
+  by(simp add: measure_spmf_def nn_integral_distr o_def)
 
 lemma measure_spmf_in_space_subprob_algebra [simp]:
   "measure_spmf p \<in> space (subprob_algebra (count_space UNIV))"
-by(simp add: space_subprob_algebra)
+  by(simp add: space_subprob_algebra)
 
 lemma nn_integral_spmf_neq_top: "(\<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV) \<noteq> \<top>"
-using nn_integral_measure_spmf[where f="\<lambda>_. 1", of p, symmetric] by simp
+  using nn_integral_measure_spmf[where f="\<lambda>_. 1", of p, symmetric] 
+  by simp
 
 lemma SUP_spmf_neq_top': "(SUP p\<in>Y. ennreal (spmf p x)) \<noteq> \<top>"
-proof(rule neq_top_trans)
-  show "(SUP p\<in>Y. ennreal (spmf p x)) \<le> 1" by(rule SUP_least)(simp add: pmf_le_1)
-qed simp
+  by (metis SUP_least ennreal_le_1 ennreal_one_neq_top neq_top_trans pmf_le_1)
 
 lemma SUP_spmf_neq_top: "(SUP i. ennreal (spmf (Y i) x)) \<noteq> \<top>"
-proof(rule neq_top_trans)
-  show "(SUP i. ennreal (spmf (Y i) x)) \<le> 1" by(rule SUP_least)(simp add: pmf_le_1)
-qed simp
+  by (meson SUP_eq_top_iff ennreal_le_1 ennreal_one_less_top linorder_not_le pmf_le_1)
 
 lemma SUP_emeasure_spmf_neq_top: "(SUP p\<in>Y. emeasure (measure_spmf p) A) \<noteq> \<top>"
-proof(rule neq_top_trans)
-  show "(SUP p\<in>Y. emeasure (measure_spmf p) A) \<le> 1"
-    by(rule SUP_least)(simp add: measure_spmf.subprob_emeasure_le_1)
-qed simp
+  by (metis ennreal_one_less_top less_SUP_iff linorder_not_le measure_spmf.subprob_emeasure_le_1)
 
 subsection \<open>Support\<close>
 
 definition set_spmf :: "'a spmf \<Rightarrow> 'a set"
-where "set_spmf p = set_pmf p \<bind> set_option"
+  where "set_spmf p = set_pmf p \<bind> set_option"
 
 lemma set_spmf_rep_eq: "set_spmf p = {x. measure (measure_spmf p) {x} \<noteq> 0}"
 proof -
   have "\<And>x :: 'a. the -` {x} \<inter> range Some = {Some x}" by auto
   then show ?thesis
-    by(auto simp add: set_spmf_def set_pmf.rep_eq measure_spmf_def measure_distr measure_restrict_space space_restrict_space intro: rev_image_eqI)
+    unfolding set_spmf_def measure_spmf_def
+    by(auto simp: set_pmf.rep_eq  measure_distr measure_restrict_space space_restrict_space)
 qed
 
 lemma in_set_spmf: "x \<in> set_spmf p \<longleftrightarrow> Some x \<in> set_pmf p"
-by(simp add: set_spmf_def)
+  by(simp add: set_spmf_def)
 
 lemma AE_measure_spmf_iff [simp]: "(AE x in measure_spmf p. P x) \<longleftrightarrow> (\<forall>x\<in>set_spmf p. P x)"
-by(auto 4 3 simp add: measure_spmf_def AE_distr_iff AE_restrict_space_iff AE_measure_pmf_iff set_spmf_def cong del: AE_cong)
+  unfolding set_spmf_def measure_spmf_def
+  by(force simp: AE_distr_iff AE_restrict_space_iff AE_measure_pmf_iff cong del: AE_cong)
 
 lemma spmf_eq_0_set_spmf: "spmf p x = 0 \<longleftrightarrow> x \<notin> set_spmf p"
-by(auto simp add: pmf_eq_0_set_pmf set_spmf_def intro: rev_image_eqI)
+  by(auto simp: pmf_eq_0_set_pmf set_spmf_def)
 
 lemma in_set_spmf_iff_spmf: "x \<in> set_spmf p \<longleftrightarrow> spmf p x \<noteq> 0"
-by(auto simp add: set_spmf_def set_pmf_iff intro: rev_image_eqI)
+  by(auto simp: set_spmf_def set_pmf_iff)
 
 lemma set_spmf_return_pmf_None [simp]: "set_spmf (return_pmf None) = {}"
-by(auto simp add: set_spmf_def)
+  by(auto simp: set_spmf_def)
 
 lemma countable_set_spmf [simp]: "countable (set_spmf p)"
-by(simp add: set_spmf_def bind_UNION)
+  by(simp add: set_spmf_def bind_UNION)
 
 lemma spmf_eqI:
   assumes "\<And>i. spmf p i = spmf q i"
@@ -389,147 +380,146 @@
     case None
     have "ennreal (pmf p i) = measure (measure_pmf p) {i}" by(simp add: pmf_def)
     also have "{i} = space (measure_pmf p) - range Some"
-      by(auto simp add: None intro: ccontr)
+      by(auto simp: None intro: ccontr)
     also have "measure (measure_pmf p) \<dots> = ennreal 1 - measure (measure_pmf p) (range Some)"
       by(simp add: measure_pmf.prob_compl ennreal_minus[symmetric] del: space_measure_pmf)
     also have "range Some = (\<Union>x\<in>set_spmf p. {Some x}) \<union> Some ` (- set_spmf p)"
       by auto
     also have "measure (measure_pmf p) \<dots> = measure (measure_pmf p) (\<Union>x\<in>set_spmf p. {Some x})"
-      by(rule measure_pmf.measure_zero_union)(auto simp add: measure_pmf.prob_eq_0 AE_measure_pmf_iff in_set_spmf_iff_spmf set_pmf_iff)
+      by(rule measure_pmf.measure_zero_union)(auto simp: measure_pmf.prob_eq_0 AE_measure_pmf_iff in_set_spmf_iff_spmf set_pmf_iff)
     also have "ennreal \<dots> = \<integral>\<^sup>+ x. measure (measure_pmf p) {Some x} \<partial>count_space (set_spmf p)"
       unfolding measure_pmf.emeasure_eq_measure[symmetric]
       by(simp_all add: emeasure_UN_countable disjoint_family_on_def)
     also have "\<dots> = \<integral>\<^sup>+ x. spmf p x \<partial>count_space (set_spmf p)" by(simp add: pmf_def)
     also have "\<dots> = \<integral>\<^sup>+ x. spmf q x \<partial>count_space (set_spmf p)" by(simp add: assms)
-    also have "set_spmf p = set_spmf q" by(auto simp add: in_set_spmf_iff_spmf assms)
+    also have "set_spmf p = set_spmf q" by(auto simp: in_set_spmf_iff_spmf assms)
     also have "(\<integral>\<^sup>+ x. spmf q x \<partial>count_space (set_spmf q)) = \<integral>\<^sup>+ x. measure (measure_pmf q) {Some x} \<partial>count_space (set_spmf q)"
       by(simp add: pmf_def)
     also have "\<dots> = measure (measure_pmf q) (\<Union>x\<in>set_spmf q. {Some x})"
       unfolding measure_pmf.emeasure_eq_measure[symmetric]
       by(simp_all add: emeasure_UN_countable disjoint_family_on_def)
     also have "\<dots> = measure (measure_pmf q) ((\<Union>x\<in>set_spmf q. {Some x}) \<union> Some ` (- set_spmf q))"
-      by(rule ennreal_cong measure_pmf.measure_zero_union[symmetric])+(auto simp add: measure_pmf.prob_eq_0 AE_measure_pmf_iff in_set_spmf_iff_spmf set_pmf_iff)
+      by(rule ennreal_cong measure_pmf.measure_zero_union[symmetric])+(auto simp: measure_pmf.prob_eq_0 AE_measure_pmf_iff in_set_spmf_iff_spmf set_pmf_iff)
     also have "((\<Union>x\<in>set_spmf q. {Some x}) \<union> Some ` (- set_spmf q)) = range Some" by auto
     also have "ennreal 1 - measure (measure_pmf q) \<dots> = measure (measure_pmf q) (space (measure_pmf q) - range Some)"
       by(simp add: one_ereal_def measure_pmf.prob_compl ennreal_minus[symmetric] del: space_measure_pmf)
     also have "space (measure_pmf q) - range Some = {i}"
-      by(auto simp add: None intro: ccontr)
+      by(auto simp: None intro: ccontr)
     also have "measure (measure_pmf q) \<dots> = pmf q i" by(simp add: pmf_def)
     finally show ?thesis by simp
   qed
 qed
 
 lemma integral_measure_spmf_restrict:
-  fixes f ::  "'a \<Rightarrow> 'b :: {banach, second_countable_topology}" shows
-  "(\<integral> x. f x \<partial>measure_spmf M) = (\<integral> x. f x \<partial>restrict_space (measure_spmf M) (set_spmf M))"
-by(auto intro!: integral_cong_AE simp add: integral_restrict_space)
+  fixes f ::  "'a \<Rightarrow> 'b :: {banach, second_countable_topology}"
+  shows "(\<integral> x. f x \<partial>measure_spmf M) = (\<integral> x. f x \<partial>restrict_space (measure_spmf M) (set_spmf M))"
+  by(auto intro!: integral_cong_AE simp add: integral_restrict_space)
 
 lemma nn_integral_measure_spmf':
   "(\<integral>\<^sup>+ x. f x \<partial>measure_spmf p) = \<integral>\<^sup>+ x. ennreal (spmf p x) * f x \<partial>count_space (set_spmf p)"
-by(auto simp add: nn_integral_measure_spmf nn_integral_count_space_indicator in_set_spmf_iff_spmf intro!: nn_integral_cong split: split_indicator)
+  by(auto simp: nn_integral_measure_spmf nn_integral_count_space_indicator in_set_spmf_iff_spmf intro!: nn_integral_cong split: split_indicator)
 
 subsection \<open>Functorial structure\<close>
 
 abbreviation map_spmf :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a spmf \<Rightarrow> 'b spmf"
-where "map_spmf f \<equiv> map_pmf (map_option f)"
+  where "map_spmf f \<equiv> map_pmf (map_option f)"
 
 context begin
 local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "spmf")\<close>
 
 lemma map_comp: "map_spmf f (map_spmf g p) = map_spmf (f \<circ> g) p"
-by(simp add: pmf.map_comp o_def option.map_comp)
+  by(simp add: pmf.map_comp o_def option.map_comp)
 
 lemma map_id0: "map_spmf id = id"
-by(simp add: pmf.map_id option.map_id0)
+  by(simp add: pmf.map_id option.map_id0)
 
 lemma map_id [simp]: "map_spmf id p = p"
-by(simp add: map_id0)
+  by(simp add: map_id0)
 
 lemma map_ident [simp]: "map_spmf (\<lambda>x. x) p = p"
-by(simp add: id_def[symmetric])
+  by(simp add: id_def[symmetric])
 
 end
 
 lemma set_map_spmf [simp]: "set_spmf (map_spmf f p) = f ` set_spmf p"
-by(simp add: set_spmf_def image_bind bind_image o_def Option.option.set_map)
+  by(simp add: set_spmf_def image_bind bind_image o_def Option.option.set_map)
 
 lemma map_spmf_cong:
-  "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q \<Longrightarrow> f x = g x \<rbrakk>
-  \<Longrightarrow> map_spmf f p = map_spmf g q"
-by(auto intro: pmf.map_cong option.map_cong simp add: in_set_spmf)
+  "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow> map_spmf f p = map_spmf g q"
+  by(auto intro: pmf.map_cong option.map_cong simp add: in_set_spmf)
 
 lemma map_spmf_cong_simp:
   "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q =simp=> f x = g x \<rbrakk>
   \<Longrightarrow> map_spmf f p = map_spmf g q"
-unfolding simp_implies_def by(rule map_spmf_cong)
+  unfolding simp_implies_def by(rule map_spmf_cong)
 
 lemma map_spmf_idI: "(\<And>x. x \<in> set_spmf p \<Longrightarrow> f x = x) \<Longrightarrow> map_spmf f p = p"
-by(rule map_pmf_idI map_option_idI)+(simp add: in_set_spmf)
+  by(rule map_pmf_idI map_option_idI)+(simp add: in_set_spmf)
 
 lemma emeasure_map_spmf:
   "emeasure (measure_spmf (map_spmf f p)) A = emeasure (measure_spmf p) (f -` A)"
-by(auto simp add: measure_spmf_def emeasure_distr measurable_restrict_space1 space_restrict_space emeasure_restrict_space intro: arg_cong2[where f=emeasure])
+  by(auto simp: measure_spmf_def emeasure_distr measurable_restrict_space1 space_restrict_space emeasure_restrict_space intro: arg_cong2[where f=emeasure])
 
 lemma measure_map_spmf: "measure (measure_spmf (map_spmf f p)) A = measure (measure_spmf p) (f -` A)"
-using emeasure_map_spmf[of f p A] by(simp add: measure_spmf.emeasure_eq_measure)
+  using emeasure_map_spmf[of f p A] by(simp add: measure_spmf.emeasure_eq_measure)
 
 lemma measure_map_spmf_conv_distr:
   "measure_spmf (map_spmf f p) = distr (measure_spmf p) (count_space UNIV) f"
-by(rule measure_eqI)(simp_all add: emeasure_map_spmf emeasure_distr)
+  by(rule measure_eqI)(simp_all add: emeasure_map_spmf emeasure_distr)
 
 lemma spmf_map_pmf_Some [simp]: "spmf (map_pmf Some p) i = pmf p i"
-by(simp add: pmf_map_inj')
+  by(simp add: pmf_map_inj')
 
 lemma spmf_map_inj: "\<lbrakk> inj_on f (set_spmf M); x \<in> set_spmf M \<rbrakk> \<Longrightarrow> spmf (map_spmf f M) (f x) = spmf M x"
-by(subst option.map(2)[symmetric, where f=f])(rule pmf_map_inj, auto simp add: in_set_spmf inj_on_def elim!: option.inj_map_strong[rotated])
+  by (smt (verit) elem_set in_set_spmf inj_on_def option.inj_map_strong option.map(2) pmf_map_inj)
 
 lemma spmf_map_inj': "inj f \<Longrightarrow> spmf (map_spmf f M) (f x) = spmf M x"
-by(subst option.map(2)[symmetric, where f=f])(rule pmf_map_inj'[OF option.inj_map])
+  by(subst option.map(2)[symmetric, where f=f])(rule pmf_map_inj'[OF option.inj_map])
 
 lemma spmf_map_outside: "x \<notin> f ` set_spmf M \<Longrightarrow> spmf (map_spmf f M) x = 0"
-unfolding spmf_eq_0_set_spmf by simp
+  unfolding spmf_eq_0_set_spmf by simp
 
 lemma ennreal_spmf_map: "ennreal (spmf (map_spmf f p) x) = emeasure (measure_spmf p) (f -` {x})"
-by(auto simp add: ennreal_pmf_map measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space intro: arg_cong2[where f=emeasure])
+  by (metis emeasure_map_spmf emeasure_spmf_single)
 
 lemma spmf_map: "spmf (map_spmf f p) x = measure (measure_spmf p) (f -` {x})"
-using ennreal_spmf_map[of f p x] by(simp add: measure_spmf.emeasure_eq_measure)
+  using ennreal_spmf_map[of f p x] by(simp add: measure_spmf.emeasure_eq_measure)
 
 lemma ennreal_spmf_map_conv_nn_integral:
   "ennreal (spmf (map_spmf f p) x) = integral\<^sup>N (measure_spmf p) (indicator (f -` {x}))"
-by(auto simp add: ennreal_pmf_map measure_spmf_def emeasure_distr space_restrict_space emeasure_restrict_space intro: arg_cong2[where f=emeasure])
+  by (simp add: ennreal_spmf_map)
 
 subsection \<open>Monad operations\<close>
 
 subsubsection \<open>Return\<close>
 
 abbreviation return_spmf :: "'a \<Rightarrow> 'a spmf"
-where "return_spmf x \<equiv> return_pmf (Some x)"
+  where "return_spmf x \<equiv> return_pmf (Some x)"
 
 lemma pmf_return_spmf: "pmf (return_spmf x) y = indicator {y} (Some x)"
-by(fact pmf_return)
+  by(fact pmf_return)
 
 lemma measure_spmf_return_spmf: "measure_spmf (return_spmf x) = Giry_Monad.return (count_space UNIV) x"
-by(rule measure_eqI)(simp_all add: measure_spmf_def emeasure_distr space_restrict_space emeasure_restrict_space indicator_def)
+  by(rule measure_eqI)(simp_all add: measure_spmf_def emeasure_distr space_restrict_space emeasure_restrict_space indicator_def)
 
 lemma measure_spmf_return_pmf_None [simp]: "measure_spmf (return_pmf None) = null_measure (count_space UNIV)"
-by(rule measure_eqI)(auto simp add: measure_spmf_def emeasure_distr space_restrict_space emeasure_restrict_space indicator_eq_0_iff)
+  by (simp add: emeasure_measure_spmf_conv_measure_pmf measure_eqI)
 
 lemma set_return_spmf [simp]: "set_spmf (return_spmf x) = {x}"
-by(auto simp add: set_spmf_def)
+  by(auto simp: set_spmf_def)
 
 subsubsection \<open>Bind\<close>
 
 definition bind_spmf :: "'a spmf \<Rightarrow> ('a \<Rightarrow> 'b spmf) \<Rightarrow> 'b spmf"
-where "bind_spmf x f = bind_pmf x (\<lambda>a. case a of None \<Rightarrow> return_pmf None | Some a' \<Rightarrow> f a')"
+  where "bind_spmf x f = bind_pmf x (\<lambda>a. case a of None \<Rightarrow> return_pmf None | Some a' \<Rightarrow> f a')"
 
 adhoc_overloading Monad_Syntax.bind bind_spmf
 
 lemma return_None_bind_spmf [simp]: "return_pmf None \<bind> (f :: 'a \<Rightarrow> _) = return_pmf None"
-by(simp add: bind_spmf_def bind_return_pmf)
+  by(simp add: bind_spmf_def bind_return_pmf)
 
 lemma return_bind_spmf [simp]: "return_spmf x \<bind> f = f x"
-by(simp add: bind_spmf_def bind_return_pmf)
+  by(simp add: bind_spmf_def bind_return_pmf)
 
 lemma bind_return_spmf [simp]: "x \<bind> return_spmf = x"
 proof -
@@ -542,7 +532,8 @@
 lemma bind_spmf_assoc [simp]:
   fixes x :: "'a spmf" and f :: "'a \<Rightarrow> 'b spmf" and g :: "'b \<Rightarrow> 'c spmf"
   shows "(x \<bind> f) \<bind> g = x \<bind> (\<lambda>y. f y \<bind> g)"
-by(auto simp add: bind_spmf_def bind_assoc_pmf fun_eq_iff bind_return_pmf split: option.split intro: arg_cong[where f="bind_pmf x"])
+  unfolding bind_spmf_def
+  by (smt (verit, best) bind_assoc_pmf bind_pmf_cong bind_return_pmf option.case_eq_if)
 
 lemma pmf_bind_spmf_None: "pmf (p \<bind> f) None = pmf p None + \<integral> x. pmf (f x) None \<partial>measure_spmf p"
   (is "?lhs = ?rhs")
@@ -551,28 +542,40 @@
   have "?lhs = \<integral> x. ?f x \<partial>measure_pmf p"
     by(simp add: bind_spmf_def pmf_bind)
   also have "\<dots> = \<integral> x. ?f None * indicator {None} x + ?f x * indicator (range Some) x \<partial>measure_pmf p"
-    by(rule Bochner_Integration.integral_cong)(auto simp add: indicator_def)
+    by(rule Bochner_Integration.integral_cong)(auto simp: indicator_def)
   also have "\<dots> = (\<integral> x. ?f None * indicator {None} x \<partial>measure_pmf p) + (\<integral> x. ?f x * indicator (range Some) x \<partial>measure_pmf p)"
     by(rule Bochner_Integration.integral_add)(auto 4 3 intro: integrable_real_mult_indicator measure_pmf.integrable_const_bound[where B=1] simp add: AE_measure_pmf_iff pmf_le_1)
   also have "\<dots> = pmf p None + \<integral> x. indicator (range Some) x * pmf (f (the x)) None \<partial>measure_pmf p"
-    by(auto simp add: measure_measure_pmf_finite indicator_eq_0_iff intro!: Bochner_Integration.integral_cong)
-  also have "\<dots> = ?rhs" unfolding measure_spmf_def
-    by(subst integral_distr)(auto simp add: integral_restrict_space)
+    by(auto simp: measure_measure_pmf_finite indicator_eq_0_iff intro!: Bochner_Integration.integral_cong)
+  also have "\<dots> = ?rhs" 
+    unfolding measure_spmf_def
+    by(subst integral_distr)(auto simp: integral_restrict_space)
   finally show ?thesis .
 qed
 
 lemma spmf_bind: "spmf (p \<bind> f) y = \<integral> x. spmf (f x) y \<partial>measure_spmf p"
-unfolding measure_spmf_def
-by(subst integral_distr)(auto simp add: bind_spmf_def pmf_bind integral_restrict_space indicator_eq_0_iff intro!: Bochner_Integration.integral_cong split: option.split)
+proof -
+  have "\<And>x. spmf (case x of None \<Rightarrow> return_pmf None | Some x \<Rightarrow> f x) y =
+         indicat_real (range Some) x * spmf (f (the x)) y"
+    by (simp add: split: option.split)
+  then show ?thesis
+    by (simp add: measure_spmf_def integral_distr bind_spmf_def pmf_bind integral_restrict_space)
+qed
 
 lemma ennreal_spmf_bind: "ennreal (spmf (p \<bind> f) x) = \<integral>\<^sup>+ y. spmf (f y) x \<partial>measure_spmf p"
-by(auto simp add: bind_spmf_def ennreal_pmf_bind nn_integral_measure_spmf_conv_measure_pmf nn_integral_restrict_space intro: nn_integral_cong split: split_indicator option.split)
+proof -
+  have "\<And>y. ennreal (spmf (case y of None \<Rightarrow> return_pmf None | Some x \<Rightarrow> f x) x) =
+             ennreal (spmf (f (the y)) x) * indicator (range Some) y"
+    by (simp add: split: option.split)
+  then show ?thesis
+    by (simp add: bind_spmf_def ennreal_pmf_bind nn_integral_measure_spmf_conv_measure_pmf nn_integral_restrict_space)
+qed
 
 lemma measure_spmf_bind_pmf: "measure_spmf (p \<bind> f) = measure_pmf p \<bind> measure_spmf \<circ> f"
   (is "?lhs = ?rhs")
 proof(rule measure_eqI)
   show "sets ?lhs = sets ?rhs"
-    by(simp add: sets_bind[where N="count_space UNIV"] space_measure_spmf)
+    by (simp add: Giry_Monad.bind_def)
 next
   fix A :: "'a set"
   have "emeasure ?lhs A = \<integral>\<^sup>+ x. emeasure (measure_spmf (f x)) A \<partial>measure_pmf p"
@@ -602,41 +605,43 @@
 qed
 
 lemma map_spmf_bind_spmf: "map_spmf f (bind_spmf p g) = bind_spmf p (map_spmf f \<circ> g)"
-by(auto simp add: bind_spmf_def map_bind_pmf fun_eq_iff split: option.split intro: arg_cong2[where f=bind_pmf])
+  by(auto simp: bind_spmf_def map_bind_pmf fun_eq_iff split: option.split intro: arg_cong2[where f=bind_pmf])
 
 lemma bind_map_spmf: "map_spmf f p \<bind> g = p \<bind> g \<circ> f"
-by(simp add: bind_spmf_def bind_map_pmf o_def cong del: option.case_cong_weak)
+  by(simp add: bind_spmf_def bind_map_pmf o_def cong del: option.case_cong_weak)
 
 lemma spmf_bind_leI:
   assumes "\<And>y. y \<in> set_spmf p \<Longrightarrow> spmf (f y) x \<le> r"
   and "0 \<le> r"
   shows "spmf (bind_spmf p f) x \<le> r"
 proof -
-  have "ennreal (spmf (bind_spmf p f) x) = \<integral>\<^sup>+ y. spmf (f y) x \<partial>measure_spmf p" by(rule ennreal_spmf_bind)
-  also have "\<dots> \<le> \<integral>\<^sup>+ y. r \<partial>measure_spmf p" by(rule nn_integral_mono_AE)(simp add: assms)
-  also have "\<dots> \<le> r" using assms measure_spmf.emeasure_space_le_1
-    by(auto simp add: measure_spmf.emeasure_eq_measure intro!: mult_left_le)
+  have "ennreal (spmf (bind_spmf p f) x) = \<integral>\<^sup>+ y. spmf (f y) x \<partial>measure_spmf p" 
+    by(rule ennreal_spmf_bind)
+  also have "\<dots> \<le> \<integral>\<^sup>+ y. r \<partial>measure_spmf p" 
+    by(rule nn_integral_mono_AE)(simp add: assms)
+  also have "\<dots> \<le> r" 
+    using assms measure_spmf.emeasure_space_le_1
+    by(auto simp: measure_spmf.emeasure_eq_measure intro!: mult_left_le)
   finally show ?thesis using assms(2) by(simp)
 qed
 
 lemma map_spmf_conv_bind_spmf: "map_spmf f p = (p \<bind> (\<lambda>x. return_spmf (f x)))"
-by(simp add: map_pmf_def bind_spmf_def)(rule bind_pmf_cong, simp_all split: option.split)
+  by(simp add: map_pmf_def bind_spmf_def)(rule bind_pmf_cong, simp_all split: option.split)
 
 lemma bind_spmf_cong:
-  "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q \<Longrightarrow> f x = g x \<rbrakk>
-  \<Longrightarrow> bind_spmf p f = bind_spmf q g"
-by(auto simp add: bind_spmf_def in_set_spmf intro: bind_pmf_cong option.case_cong)
+  "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow> bind_spmf p f = bind_spmf q g"
+  by(auto simp: bind_spmf_def in_set_spmf intro: bind_pmf_cong option.case_cong)
 
 lemma bind_spmf_cong_simp:
   "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q =simp=> f x = g x \<rbrakk>
   \<Longrightarrow> bind_spmf p f = bind_spmf q g"
-by(simp add: simp_implies_def cong: bind_spmf_cong)
+  by(simp add: simp_implies_def cong: bind_spmf_cong)
 
 lemma set_bind_spmf: "set_spmf (M \<bind> f) = set_spmf M \<bind> (set_spmf \<circ> f)"
-by(auto simp add: set_spmf_def bind_spmf_def bind_UNION split: option.splits)
+  by(auto simp: set_spmf_def bind_spmf_def bind_UNION split: option.splits)
 
 lemma bind_spmf_const_return_None [simp]: "bind_spmf p (\<lambda>_. return_pmf None) = return_pmf None"
-by(simp add: bind_spmf_def case_option_collapse)
+  by(simp add: bind_spmf_def case_option_collapse)
 
 lemma bind_commute_spmf:
   "bind_spmf p (\<lambda>x. bind_spmf q (f x)) = bind_spmf q (\<lambda>y. bind_spmf p (\<lambda>x. f x y))"
@@ -654,32 +659,24 @@
 subsection \<open>Relator\<close>
 
 abbreviation rel_spmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a spmf \<Rightarrow> 'b spmf \<Rightarrow> bool"
-where "rel_spmf R \<equiv> rel_pmf (rel_option R)"
-
-lemma rel_pmf_mono:
-  "\<lbrakk>rel_pmf A f g; \<And>x y. A x y \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> rel_pmf B f g"
-using pmf.rel_mono[of A B] by(simp add: le_fun_def)
+  where "rel_spmf R \<equiv> rel_pmf (rel_option R)"
 
 lemma rel_spmf_mono:
   "\<lbrakk>rel_spmf A f g; \<And>x y. A x y \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> rel_spmf B f g"
-apply(erule rel_pmf_mono)
-using option.rel_mono[of A B] by(simp add: le_fun_def)
+  by (metis option.rel_sel pmf.rel_mono_strong)
 
 lemma rel_spmf_mono_strong:
   "\<lbrakk> rel_spmf A f g; \<And>x y. \<lbrakk> A x y; x \<in> set_spmf f; y \<in> set_spmf g \<rbrakk> \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> rel_spmf B f g"
-apply(erule pmf.rel_mono_strong)
-apply(erule option.rel_mono_strong)
-apply(auto simp add: in_set_spmf)
-done
+  by (metis elem_set in_set_spmf option.rel_mono_strong pmf.rel_mono_strong)
 
 lemma rel_spmf_reflI: "(\<And>x. x \<in> set_spmf p \<Longrightarrow> P x x) \<Longrightarrow> rel_spmf P p p"
-by(rule rel_pmf_reflI)(auto simp add: set_spmf_def intro: rel_option_reflI)
+  by (metis (mono_tags, lifting) option.rel_eq pmf.rel_eq rel_spmf_mono_strong)
 
 lemma rel_spmfI [intro?]:
   "\<lbrakk> \<And>x y. (x, y) \<in> set_spmf pq \<Longrightarrow> P x y; map_spmf fst pq = p; map_spmf snd pq = q \<rbrakk>
   \<Longrightarrow> rel_spmf P p q"
 by(rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. case x of None \<Rightarrow> (None, None) | Some (a, b) \<Rightarrow> (Some a, Some b)) pq"])
-  (auto simp add: pmf.map_comp o_def in_set_spmf split: option.splits intro: pmf.map_cong)
+  (auto simp: pmf.map_comp o_def in_set_spmf split: option.splits intro: pmf.map_cong)
 
 lemma rel_spmfE [elim?, consumes 1, case_names rel_spmf]:
   assumes "rel_spmf P p q"
@@ -687,76 +684,76 @@
     "\<And>x y. (x, y) \<in> set_spmf pq \<Longrightarrow> P x y"
     "p = map_spmf fst pq"
     "q = map_spmf snd pq"
-using assms
+  using assms
 proof(cases rule: rel_pmf.cases[consumes 1, case_names rel_pmf])
   case (rel_pmf pq)
   let ?pq = "map_pmf (\<lambda>(a, b). case (a, b) of (Some x, Some y) \<Rightarrow> Some (x, y) | _ \<Rightarrow> None) pq"
   have "\<And>x y. (x, y) \<in> set_spmf ?pq \<Longrightarrow> P x y"
-    by(auto simp add: in_set_spmf split: option.split_asm dest: rel_pmf(1))
+    by(auto simp: in_set_spmf split: option.split_asm dest: rel_pmf(1))
   moreover
   have "\<And>x. (x, None) \<in> set_pmf pq \<Longrightarrow> x = None" by(auto dest!: rel_pmf(1))
   then have "p = map_spmf fst ?pq" using rel_pmf(2)
-    by(auto simp add: pmf.map_comp split_beta intro!: pmf.map_cong split: option.split)
+    by(auto simp: pmf.map_comp split_beta intro!: pmf.map_cong split: option.split)
   moreover
   have "\<And>y. (None, y) \<in> set_pmf pq \<Longrightarrow> y = None" by(auto dest!: rel_pmf(1))
   then have "q = map_spmf snd ?pq" using rel_pmf(3)
-    by(auto simp add: pmf.map_comp split_beta intro!: pmf.map_cong split: option.split)
+    by(auto simp: pmf.map_comp split_beta intro!: pmf.map_cong split: option.split)
   ultimately show thesis ..
 qed
 
 lemma rel_spmf_simps:
   "rel_spmf R p q \<longleftrightarrow> (\<exists>pq. (\<forall>(x, y)\<in>set_spmf pq. R x y) \<and> map_spmf fst pq = p \<and> map_spmf snd pq = q)"
-by(auto intro: rel_spmfI elim!: rel_spmfE)
+  by(auto intro: rel_spmfI elim!: rel_spmfE)
 
 lemma spmf_rel_map:
   shows spmf_rel_map1: "\<And>R f x. rel_spmf R (map_spmf f x) = rel_spmf (\<lambda>x. R (f x)) x"
-  and spmf_rel_map2: "\<And>R x g y. rel_spmf R x (map_spmf g y) = rel_spmf (\<lambda>x y. R x (g y)) x y"
-by(simp_all add: fun_eq_iff pmf.rel_map option.rel_map[abs_def])
+    and spmf_rel_map2: "\<And>R x g y. rel_spmf R x (map_spmf g y) = rel_spmf (\<lambda>x y. R x (g y)) x y"
+  by(simp_all add: fun_eq_iff pmf.rel_map option.rel_map[abs_def])
 
 lemma spmf_rel_conversep: "rel_spmf R\<inverse>\<inverse> = (rel_spmf R)\<inverse>\<inverse>"
-by(simp add: option.rel_conversep pmf.rel_conversep)
+  by(simp add: option.rel_conversep pmf.rel_conversep)
 
 lemma spmf_rel_eq: "rel_spmf (=) = (=)"
-by(simp add: pmf.rel_eq option.rel_eq)
+  by(simp add: pmf.rel_eq option.rel_eq)
 
 context includes lifting_syntax
 begin
 
 lemma bind_spmf_parametric [transfer_rule]:
   "(rel_spmf A ===> (A ===> rel_spmf B) ===> rel_spmf B) bind_spmf bind_spmf"
-unfolding bind_spmf_def[abs_def] by transfer_prover
+  unfolding bind_spmf_def[abs_def] by transfer_prover
 
 lemma return_spmf_parametric: "(A ===> rel_spmf A) return_spmf return_spmf"
-by transfer_prover
+  by transfer_prover
 
 lemma map_spmf_parametric: "((A ===> B) ===> rel_spmf A ===> rel_spmf B) map_spmf map_spmf"
-by transfer_prover
+  by transfer_prover
 
 lemma rel_spmf_parametric:
   "((A ===> B ===> (=)) ===> rel_spmf A ===> rel_spmf B ===> (=)) rel_spmf rel_spmf"
-by transfer_prover
+  by transfer_prover
 
 lemma set_spmf_parametric [transfer_rule]:
   "(rel_spmf A ===> rel_set A) set_spmf set_spmf"
-unfolding set_spmf_def[abs_def] by transfer_prover
+  unfolding set_spmf_def[abs_def] by transfer_prover
 
 lemma return_spmf_None_parametric:
   "(rel_spmf A) (return_pmf None) (return_pmf None)"
-by simp
+  by simp
 
 end
 
 lemma rel_spmf_bindI:
   "\<lbrakk> rel_spmf R p q; \<And>x y. R x y \<Longrightarrow> rel_spmf P (f x) (g y) \<rbrakk>
   \<Longrightarrow> rel_spmf P (p \<bind> f) (q \<bind> g)"
-by(fact bind_spmf_parametric[THEN rel_funD, THEN rel_funD, OF _ rel_funI])
+  by(fact bind_spmf_parametric[THEN rel_funD, THEN rel_funD, OF _ rel_funI])
 
 lemma rel_spmf_bind_reflI:
   "(\<And>x. x \<in> set_spmf p \<Longrightarrow> rel_spmf P (f x) (g x)) \<Longrightarrow> rel_spmf P (p \<bind> f) (p \<bind> g)"
-by(rule rel_spmf_bindI[where R="\<lambda>x y. x = y \<and> x \<in> set_spmf p"])(auto intro: rel_spmf_reflI)
+  by(rule rel_spmf_bindI[where R="\<lambda>x y. x = y \<and> x \<in> set_spmf p"])(auto intro: rel_spmf_reflI)
 
 lemma rel_pmf_return_pmfI: "P x y \<Longrightarrow> rel_pmf P (return_pmf x) (return_pmf y)"
-by(rule rel_pmf.intros[where pq="return_pmf (x, y)"])(simp_all)
+  by simp
 
 context includes lifting_syntax
 begin
@@ -770,97 +767,99 @@
   from this(1) obtain pq where A: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> A x y"
     and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto
   show "measure p X = measure q Y" unfolding p q measure_map_pmf
-    by(rule measure_pmf.finite_measure_eq_AE)(auto simp add: AE_measure_pmf_iff dest!: A rel_predD[OF \<open>rel_pred _ _ _\<close>])
+    by(rule measure_pmf.finite_measure_eq_AE)(auto simp: AE_measure_pmf_iff dest!: A rel_predD[OF \<open>rel_pred _ _ _\<close>])
 qed
 
 lemma measure_spmf_parametric:
   "(rel_spmf A ===> rel_pred A ===> (=)) (\<lambda>p. measure (measure_spmf p)) (\<lambda>q. measure (measure_spmf q))"
-unfolding measure_measure_spmf_conv_measure_pmf[abs_def]
-apply(rule rel_funI)+
-apply(erule measure_pmf_parametric[THEN rel_funD, THEN rel_funD])
-apply(auto simp add: rel_pred_def rel_fun_def elim: option.rel_cases)
-done
+proof -
+  have "\<And>x y xa ya. rel_pred A xa ya \<Longrightarrow> rel_pred (rel_option A) (Some ` xa) (Some ` ya)"
+    by(auto simp: rel_pred_def rel_fun_def elim: option.rel_cases)
+  then show ?thesis
+  unfolding measure_measure_spmf_conv_measure_pmf[abs_def]
+  by (intro rel_funI) (force elim!: measure_pmf_parametric[THEN rel_funD, THEN rel_funD])
+qed
 
 end
 
 subsection \<open>From \<^typ>\<open>'a pmf\<close> to \<^typ>\<open>'a spmf\<close>\<close>
 
 definition spmf_of_pmf :: "'a pmf \<Rightarrow> 'a spmf"
-where "spmf_of_pmf = map_pmf Some"
+  where "spmf_of_pmf = map_pmf Some"
 
 lemma set_spmf_spmf_of_pmf [simp]: "set_spmf (spmf_of_pmf p) = set_pmf p"
-by(auto simp add: spmf_of_pmf_def set_spmf_def bind_image o_def)
+  by(auto simp: spmf_of_pmf_def set_spmf_def bind_image o_def)
 
 lemma spmf_spmf_of_pmf [simp]: "spmf (spmf_of_pmf p) x = pmf p x"
-by(simp add: spmf_of_pmf_def)
+  by(simp add: spmf_of_pmf_def)
 
 lemma pmf_spmf_of_pmf_None [simp]: "pmf (spmf_of_pmf p) None = 0"
-using ennreal_pmf_map[of Some p None] by(simp add: spmf_of_pmf_def)
+  using ennreal_pmf_map[of Some p None] by(simp add: spmf_of_pmf_def)
 
 lemma emeasure_spmf_of_pmf [simp]: "emeasure (measure_spmf (spmf_of_pmf p)) A = emeasure (measure_pmf p) A"
-by(simp add: emeasure_measure_spmf_conv_measure_pmf spmf_of_pmf_def inj_vimage_image_eq)
+  by(simp add: emeasure_measure_spmf_conv_measure_pmf spmf_of_pmf_def inj_vimage_image_eq)
 
 lemma measure_spmf_spmf_of_pmf [simp]: "measure_spmf (spmf_of_pmf p) = measure_pmf p"
-by(rule measure_eqI) simp_all
+  by(rule measure_eqI) simp_all
 
 lemma map_spmf_of_pmf [simp]: "map_spmf f (spmf_of_pmf p) = spmf_of_pmf (map_pmf f p)"
-by(simp add: spmf_of_pmf_def pmf.map_comp o_def)
+  by(simp add: spmf_of_pmf_def pmf.map_comp o_def)
 
 lemma rel_spmf_spmf_of_pmf [simp]: "rel_spmf R (spmf_of_pmf p) (spmf_of_pmf q) = rel_pmf R p q"
-by(simp add: spmf_of_pmf_def pmf.rel_map)
+  by(simp add: spmf_of_pmf_def pmf.rel_map)
 
 lemma spmf_of_pmf_return_pmf [simp]: "spmf_of_pmf (return_pmf x) = return_spmf x"
-by(simp add: spmf_of_pmf_def)
+  by(simp add: spmf_of_pmf_def)
 
 lemma bind_spmf_of_pmf [simp]: "bind_spmf (spmf_of_pmf p) f = bind_pmf p f"
-by(simp add: spmf_of_pmf_def bind_spmf_def bind_map_pmf)
+  by(simp add: spmf_of_pmf_def bind_spmf_def bind_map_pmf)
 
 lemma set_spmf_bind_pmf: "set_spmf (bind_pmf p f) = Set.bind (set_pmf p) (set_spmf \<circ> f)"
-unfolding bind_spmf_of_pmf[symmetric] by(subst set_bind_spmf) simp
+  unfolding bind_spmf_of_pmf[symmetric] by(subst set_bind_spmf) simp
 
 lemma spmf_of_pmf_bind: "spmf_of_pmf (bind_pmf p f) = bind_pmf p (\<lambda>x. spmf_of_pmf (f x))"
-by(simp add: spmf_of_pmf_def map_bind_pmf)
+  by(simp add: spmf_of_pmf_def map_bind_pmf)
 
 lemma bind_pmf_return_spmf: "p \<bind> (\<lambda>x. return_spmf (f x)) = spmf_of_pmf (map_pmf f p)"
-by(simp add: map_pmf_def spmf_of_pmf_bind)
+  by(simp add: map_pmf_def spmf_of_pmf_bind)
 
 subsection \<open>Weight of a subprobability\<close>
 
 abbreviation weight_spmf :: "'a spmf \<Rightarrow> real"
-where "weight_spmf p \<equiv> measure (measure_spmf p) (space (measure_spmf p))"
+  where "weight_spmf p \<equiv> measure (measure_spmf p) (space (measure_spmf p))"
 
 lemma weight_spmf_def: "weight_spmf p = measure (measure_spmf p) UNIV"
-by(simp add: space_measure_spmf)
+  by(simp add: space_measure_spmf)
 
 lemma weight_spmf_le_1: "weight_spmf p \<le> 1"
-by(simp add: measure_spmf.subprob_measure_le_1)
+  by(rule measure_spmf.subprob_measure_le_1)
 
 lemma weight_return_spmf [simp]: "weight_spmf (return_spmf x) = 1"
-by(simp add: measure_spmf_return_spmf measure_return)
+  by(simp add: measure_spmf_return_spmf measure_return)
 
 lemma weight_return_pmf_None [simp]: "weight_spmf (return_pmf None) = 0"
-by(simp)
+  by(simp)
 
 lemma weight_map_spmf [simp]: "weight_spmf (map_spmf f p) = weight_spmf p"
-by(simp add: weight_spmf_def measure_map_spmf)
+  by(simp add: weight_spmf_def measure_map_spmf)
 
 lemma weight_spmf_of_pmf [simp]: "weight_spmf (spmf_of_pmf p) = 1"
-using measure_pmf.prob_space[of p] by(simp add: spmf_of_pmf_def weight_spmf_def)
+  by simp
 
 lemma weight_spmf_nonneg: "weight_spmf p \<ge> 0"
-by(fact measure_nonneg)
+  by(fact measure_nonneg)
 
 lemma (in finite_measure) integrable_weight_spmf [simp]:
   "(\<lambda>x. weight_spmf (f x)) \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. weight_spmf (f x))"
-by(rule integrable_const_bound[where B=1])(simp_all add: weight_spmf_nonneg weight_spmf_le_1)
+  by(rule integrable_const_bound[where B=1])(simp_all add: weight_spmf_nonneg weight_spmf_le_1)
 
 lemma weight_spmf_eq_nn_integral_spmf: "weight_spmf p = \<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV"
-by(simp add: measure_measure_spmf_conv_measure_pmf space_measure_spmf measure_pmf.emeasure_eq_measure[symmetric] nn_integral_pmf[symmetric] embed_measure_count_space[symmetric] inj_on_def nn_integral_embed_measure measurable_embed_measure1)
+  by (metis NO_MATCH_def measure_spmf.emeasure_eq_measure nn_integral_count_space_indicator nn_integral_indicator nn_integral_measure_spmf sets_UNIV sets_measure_spmf space_measure_spmf)
 
 lemma weight_spmf_eq_nn_integral_support:
   "weight_spmf p = \<integral>\<^sup>+ x. spmf p x \<partial>count_space (set_spmf p)"
-unfolding weight_spmf_eq_nn_integral_spmf
-by(auto simp add: nn_integral_count_space_indicator in_set_spmf_iff_spmf intro!: nn_integral_cong split: split_indicator)
+  unfolding weight_spmf_eq_nn_integral_spmf
+  by(auto simp: nn_integral_count_space_indicator in_set_spmf_iff_spmf intro!: nn_integral_cong split: split_indicator)
 
 lemma pmf_None_eq_weight_spmf: "pmf p None = 1 - weight_spmf p"
 proof -
@@ -876,30 +875,23 @@
 qed
 
 lemma weight_spmf_conv_pmf_None: "weight_spmf p = 1 - pmf p None"
-by(simp add: pmf_None_eq_weight_spmf)
-
-lemma weight_spmf_le_0: "weight_spmf p \<le> 0 \<longleftrightarrow> weight_spmf p = 0"
-by(rule measure_le_0_iff)
+  by(simp add: pmf_None_eq_weight_spmf)
 
 lemma weight_spmf_lt_0: "\<not> weight_spmf p < 0"
-by(simp add: not_less weight_spmf_nonneg)
+  by(simp add: not_less weight_spmf_nonneg)
 
 lemma spmf_le_weight: "spmf p x \<le> weight_spmf p"
-proof -
-  have "ennreal (spmf p x) \<le> weight_spmf p"
-    unfolding weight_spmf_eq_nn_integral_spmf by(rule nn_integral_ge_point) simp
-  then show ?thesis by simp
-qed
+  by (simp add: measure_spmf.bounded_measure spmf_conv_measure_spmf)
 
 lemma weight_spmf_eq_0: "weight_spmf p = 0 \<longleftrightarrow> p = return_pmf None"
-by(auto intro!: pmf_eqI simp add: pmf_None_eq_weight_spmf split: split_indicator)(metis not_Some_eq pmf_le_0_iff spmf_le_weight)
+  by (metis measure_le_0_iff measure_spmf.bounded_measure spmf_conv_measure_spmf spmf_eqI weight_return_pmf_None)
 
 lemma weight_bind_spmf: "weight_spmf (x \<bind> f) = lebesgue_integral (measure_spmf x) (weight_spmf \<circ> f)"
-unfolding weight_spmf_def
-by(simp add: measure_spmf_bind o_def measure_spmf.measure_bind[where N="count_space UNIV"])
+  unfolding weight_spmf_def
+  by(simp add: measure_spmf_bind o_def measure_spmf.measure_bind[where N="count_space UNIV"])
 
 lemma rel_spmf_weightD: "rel_spmf A p q \<Longrightarrow> weight_spmf p = weight_spmf q"
-by(erule rel_spmfE) simp
+  by(erule rel_spmfE) simp
 
 lemma rel_spmf_bij_betw:
   assumes f: "bij_betw f (set_spmf p) (set_spmf q)"
@@ -914,14 +906,22 @@
   then have "None \<in> set_pmf p \<longleftrightarrow> None \<in> set_pmf q"
     by(simp add: pmf_None_eq_weight_spmf set_pmf_iff)
   with f have "bij_betw (map_option f) (set_pmf p) (set_pmf q)"
-    apply(auto simp add: bij_betw_def in_set_spmf inj_on_def intro: option.expand)
+    apply(auto simp: bij_betw_def in_set_spmf inj_on_def intro: option.expand split: option.split)
     apply(rename_tac [!] x)
     apply(case_tac [!] x)
     apply(auto iff: in_set_spmf)
     done
   then have "rel_pmf (\<lambda>x y. ?f x = y) p q"
-    by(rule rel_pmf_bij_betw)(case_tac x, simp_all add: weq[simplified] eq in_set_spmf pmf_None_eq_weight_spmf)
-  thus ?thesis by(rule pmf.rel_mono_strong)(auto intro!: rel_optionI simp add: Option.is_none_def)
+  proof (rule rel_pmf_bij_betw)
+    show "pmf p x = pmf q (map_option f x)" if "x \<in> set_pmf p" for x
+    proof (cases x)
+      case None
+      then show ?thesis
+        by (metis ennreal_inj measure_nonneg option.map_disc_iff pmf_None_eq_weight_spmf weq)
+    qed (use eq in_set_spmf that in force)
+  qed
+  thus ?thesis
+    by (smt (verit, ccfv_SIG) None_eq_map_option_iff option.map_sel option.rel_sel pmf.rel_mono_strong) 
 qed
 
 subsection \<open>From density to spmfs\<close>
@@ -929,7 +929,7 @@
 context fixes f :: "'a \<Rightarrow> real" begin
 
 definition embed_spmf :: "'a spmf"
-where "embed_spmf = embed_pmf (\<lambda>x. case x of None \<Rightarrow> 1 - enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV) | Some x' \<Rightarrow> max 0 (f x'))"
+  where "embed_spmf = embed_pmf (\<lambda>x. case x of None \<Rightarrow> 1 - enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV) | Some x' \<Rightarrow> max 0 (f x'))"
 
 context
   assumes prob: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV) \<le> 1"
@@ -951,32 +951,25 @@
   also have "(\<integral>\<^sup>+ x. ?f x \<partial>\<dots>) = \<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV"
     by(subst nn_integral_embed_measure)(simp_all add: measurable_embed_measure1)
   also have "?None + \<dots> = 1" using prob
-    by(auto simp add: ennreal_minus[symmetric] ennreal_1[symmetric] ennreal_enn2real_if top_unique simp del: ennreal_1)(simp add: diff_add_self_ennreal)
+    by(auto simp: ennreal_minus[symmetric] ennreal_1[symmetric] ennreal_enn2real_if top_unique simp del: ennreal_1)(simp add: diff_add_self_ennreal)
   finally show ?thesis .
 qed
 
 lemma pmf_embed_spmf_None: "pmf embed_spmf None = 1 - enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV)"
 unfolding embed_spmf_def
-apply(subst pmf_embed_pmf)
-  subgoal using prob by(simp add: field_simps enn2real_leI split: option.split)
- apply(rule nn_integral_embed_spmf_eq_1)
-apply simp
-done
+  by (smt (verit, del_insts) enn2real_leI ennreal_1 nn_integral_cong nn_integral_embed_spmf_eq_1
+      option.case_eq_if pmf_embed_pmf prob)
 
 lemma spmf_embed_spmf [simp]: "spmf embed_spmf x = max 0 (f x)"
-unfolding embed_spmf_def
-apply(subst pmf_embed_pmf)
-  subgoal using prob by(simp add: field_simps enn2real_leI split: option.split)
- apply(rule nn_integral_embed_spmf_eq_1)
-apply simp
-done
+  unfolding embed_spmf_def
+  by (smt (verit, best) enn2real_leI ennreal_1 nn_integral_cong nn_integral_embed_spmf_eq_1 option.case_eq_if option.simps(5) pmf_embed_pmf prob)
 
 end
 
 end
 
-lemma embed_spmf_K_0[simp]: "embed_spmf (\<lambda>_. 0) = return_pmf None" (is "?lhs = ?rhs")
-by(rule spmf_eqI)(simp add: zero_ereal_def[symmetric])
+lemma embed_spmf_K_0[simp]: "embed_spmf (\<lambda>_. 0) = return_pmf None" 
+  by(rule spmf_eqI)(simp add: zero_ereal_def[symmetric])
 
 subsection \<open>Ordering on spmfs\<close>
 
@@ -996,20 +989,20 @@
 \<close>
 
 abbreviation ord_spmf :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a spmf \<Rightarrow> 'a spmf \<Rightarrow> bool"
-where "ord_spmf ord \<equiv> rel_pmf (ord_option ord)"
+  where "ord_spmf ord \<equiv> rel_pmf (ord_option ord)"
 
 locale ord_spmf_syntax begin
 notation ord_spmf (infix "\<sqsubseteq>\<index>" 60)
 end
 
 lemma ord_spmf_map_spmf1: "ord_spmf R (map_spmf f p) = ord_spmf (\<lambda>x. R (f x)) p"
-by(simp add: pmf.rel_map[abs_def] ord_option_map1[abs_def])
+  by(simp add: pmf.rel_map[abs_def] ord_option_map1[abs_def])
 
 lemma ord_spmf_map_spmf2: "ord_spmf R p (map_spmf f q) = ord_spmf (\<lambda>x y. R x (f y)) p q"
-by(simp add: pmf.rel_map ord_option_map2)
+  by(simp add: pmf.rel_map ord_option_map2)
 
 lemma ord_spmf_map_spmf12: "ord_spmf R (map_spmf f p) (map_spmf f q) = ord_spmf (\<lambda>x y. R (f x) (f y)) p q"
-by(simp add: pmf.rel_map ord_option_map1[abs_def] ord_option_map2)
+  by(simp add: pmf.rel_map ord_option_map1[abs_def] ord_option_map2)
 
 lemmas ord_spmf_map_spmf = ord_spmf_map_spmf1 ord_spmf_map_spmf2 ord_spmf_map_spmf12
 
@@ -1019,14 +1012,14 @@
 lemma ord_spmfI:
   "\<lbrakk> \<And>x y. (x, y) \<in> set_spmf pq \<Longrightarrow> ord x y; map_spmf fst pq = p; map_spmf snd pq = q \<rbrakk>
   \<Longrightarrow> p \<sqsubseteq> q"
-by(rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. case x of None \<Rightarrow> (None, None) | Some (a, b) \<Rightarrow> (Some a, Some b)) pq"])
-  (auto simp add: pmf.map_comp o_def in_set_spmf split: option.splits intro: pmf.map_cong)
+  by(rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. case x of None \<Rightarrow> (None, None) | Some (a, b) \<Rightarrow> (Some a, Some b)) pq"])
+    (auto simp: pmf.map_comp o_def in_set_spmf split: option.splits intro: pmf.map_cong)
 
 lemma ord_spmf_None [simp]: "return_pmf None \<sqsubseteq> x"
-by(rule rel_pmf.intros[where pq="map_pmf (Pair None) x"])(auto simp add: pmf.map_comp o_def)
+  by(rule rel_pmf.intros[where pq="map_pmf (Pair None) x"])(auto simp: pmf.map_comp o_def)
 
 lemma ord_spmf_reflI: "(\<And>x. x \<in> set_spmf p \<Longrightarrow> ord x x) \<Longrightarrow> p \<sqsubseteq> p"
-by(rule rel_pmf_reflI ord_option_reflI)+(auto simp add: in_set_spmf)
+  by (metis elem_set in_set_spmf ord_option_reflI pmf.rel_refl_strong)
 
 lemma rel_spmf_inf:
   assumes "p \<sqsubseteq> q"
@@ -1037,34 +1030,33 @@
 proof -
   from \<open>p \<sqsubseteq> q\<close> \<open>q \<sqsubseteq> p\<close>
   have "rel_pmf (inf (ord_option ord) (ord_option ord)\<inverse>\<inverse>) p q"
-    by(rule rel_pmf_inf)(blast intro: reflp_ord_option transp_ord_option refl trans)+
+    using local.refl local.trans reflp_ord_option rel_pmf_inf transp_ord_option by blast
   also have "inf (ord_option ord) (ord_option ord)\<inverse>\<inverse> = rel_option (inf ord ord\<inverse>\<inverse>)"
-    by(auto simp add: fun_eq_iff elim: ord_option.cases option.rel_cases)
+    by(auto simp: fun_eq_iff elim: ord_option.cases option.rel_cases)
   finally show ?thesis .
 qed
 
 end
 
 lemma ord_spmf_return_spmf2: "ord_spmf R p (return_spmf y) \<longleftrightarrow> (\<forall>x\<in>set_spmf p. R x y)"
-by(auto simp add: rel_pmf_return_pmf2 in_set_spmf ord_option.simps intro: ccontr)
+  by(auto simp: rel_pmf_return_pmf2 in_set_spmf ord_option.simps intro: ccontr)
 
 lemma ord_spmf_mono: "\<lbrakk> ord_spmf A p q; \<And>x y. A x y \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> ord_spmf B p q"
-by(erule rel_pmf_mono)(erule ord_option_mono)
+  by(erule pmf.rel_mono_strong)(erule ord_option_mono)
 
 lemma ord_spmf_compp: "ord_spmf (A OO B) = ord_spmf A OO ord_spmf B"
-by(simp add: ord_option_compp pmf.rel_compp)
+  by(simp add: ord_option_compp pmf.rel_compp)
 
 lemma ord_spmf_bindI:
   assumes pq: "ord_spmf R p q"
-  and fg: "\<And>x y. R x y \<Longrightarrow> ord_spmf P (f x) (g y)"
+    and fg: "\<And>x y. R x y \<Longrightarrow> ord_spmf P (f x) (g y)"
   shows "ord_spmf P (p \<bind> f) (q \<bind> g)"
-unfolding bind_spmf_def using pq
-by(rule rel_pmf_bindI)(auto split: option.split intro: fg)
+  unfolding bind_spmf_def using pq
+  by(rule rel_pmf_bindI)(auto split: option.split intro: fg)
 
 lemma ord_spmf_bind_reflI:
-  "(\<And>x. x \<in> set_spmf p \<Longrightarrow> ord_spmf R (f x) (g x))
-  \<Longrightarrow> ord_spmf R (p \<bind> f) (p \<bind> g)"
-by(rule ord_spmf_bindI[where R="\<lambda>x y. x = y \<and> x \<in> set_spmf p"])(auto intro: ord_spmf_reflI)
+  "(\<And>x. x \<in> set_spmf p \<Longrightarrow> ord_spmf R (f x) (g x)) \<Longrightarrow> ord_spmf R (p \<bind> f) (p \<bind> g)"
+  by(rule ord_spmf_bindI[where R="\<lambda>x y. x = y \<and> x \<in> set_spmf p"])(auto intro: ord_spmf_reflI)
 
 lemma ord_pmf_increaseI:
   assumes le: "\<And>x. spmf p x \<le> spmf q x"
@@ -1098,12 +1090,13 @@
     also have "?Some = \<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV"
       by(simp add: nn_integral_count_space_indicator[symmetric] embed_measure_count_space[symmetric] inj_on_def nn_integral_embed_measure measurable_embed_measure1)
     also have "pmf q None + (?Some2' - ?Some2'') + \<dots> = pmf q None + ?Some2'"
-      by(auto simp add: diff_add_self_ennreal le intro!: nn_integral_mono)
+      by(auto simp: diff_add_self_ennreal le intro!: nn_integral_mono)
     also have "\<dots> = \<integral>\<^sup>+ x. ennreal (pmf q x) * indicator {None} x + ennreal (pmf q x) * indicator (range Some) x \<partial>count_space UNIV"
       by(subst nn_integral_add)(simp_all add: nn_integral_count_space_indicator[symmetric] embed_measure_count_space[symmetric] nn_integral_embed_measure measurable_embed_measure1)
     also have "\<dots> = \<integral>\<^sup>+ x. pmf q x \<partial>count_space UNIV"
       by(rule nn_integral_cong)(auto split: split_indicator)
-    also have "\<dots> = 1" by(simp add: nn_integral_pmf)
+    also have "\<dots> = 1" 
+      by(simp add: nn_integral_pmf)
     finally show ?thesis .
   qed
   note f = nonneg integral
@@ -1115,18 +1108,17 @@
       by(simp add: spmf_eq_0_set_spmf refl split: option.split_asm if_split_asm) }
 
   have weight_le: "weight_spmf p \<le> weight_spmf q"
-    by(subst ennreal_le_iff[symmetric])(auto simp add: weight_spmf_eq_nn_integral_spmf intro!: nn_integral_mono le)
+    by(subst ennreal_le_iff[symmetric])(auto simp: weight_spmf_eq_nn_integral_spmf intro!: nn_integral_mono le)
 
   show "map_pmf fst pq = p"
   proof(rule pmf_eqI)
-    fix i
+    fix i :: "'a option"
+    have bi: "bij_betw (Pair i) UNIV (fst -` {i})"
+      by(auto simp: bij_betw_def inj_on_def)
     have "ennreal (pmf (map_pmf fst pq) i) = (\<integral>\<^sup>+ y. pmf pq (i, y) \<partial>count_space UNIV)"
       unfolding pq_def ennreal_pmf_map
-      apply(simp add: embed_pmf.rep_eq[OF f] o_def emeasure_density nn_integral_count_space_indicator[symmetric])
-      apply(subst pmf_embed_pmf[OF f])
-      apply(rule nn_integral_bij_count_space[symmetric])
-      apply(auto simp add: bij_betw_def inj_on_def)
-      done
+      apply (simp add: embed_pmf.rep_eq[OF f] o_def emeasure_density flip: nn_integral_count_space_indicator)
+      by (smt (verit, best) nn_integral_bij_count_space [OF bi] integral nn_integral_cong nonneg pmf_embed_pmf)
     also have "\<dots> = pmf p i"
     proof(cases i)
       case (Some x)
@@ -1150,21 +1142,20 @@
         by(simp add: pmf_None_eq_weight_spmf weight_spmf_eq_nn_integral_spmf[symmetric] ennreal_minus)
       also have "\<dots> = ennreal (pmf p None) - ennreal (pmf q None)" by(simp add: ennreal_minus)
       finally show ?thesis using None weight_le
-        by(auto simp add: diff_add_self_ennreal pmf_None_eq_weight_spmf intro: ennreal_leI)
+        by(auto simp: diff_add_self_ennreal pmf_None_eq_weight_spmf intro: ennreal_leI)
     qed
     finally show "pmf (map_pmf fst pq) i = pmf p i" by simp
   qed
 
   show "map_pmf snd pq = q"
   proof(rule pmf_eqI)
-    fix i
+    fix i :: "'a option"
+    have bi: "bij_betw (\<lambda>x. (x, i)) UNIV (snd -` {i})"
+      by (auto simp: bij_betw_def inj_on_def)
     have "ennreal (pmf (map_pmf snd pq) i) = (\<integral>\<^sup>+ x. pmf pq (x, i) \<partial>count_space UNIV)"
       unfolding pq_def ennreal_pmf_map
       apply(simp add: embed_pmf.rep_eq[OF f] o_def emeasure_density nn_integral_count_space_indicator[symmetric])
-      apply(subst pmf_embed_pmf[OF f])
-      apply(rule nn_integral_bij_count_space[symmetric])
-      apply(auto simp add: bij_betw_def inj_on_def)
-      done
+      by (smt (verit, best) nn_integral_bij_count_space [OF bi] integral nn_integral_cong nonneg pmf_embed_pmf)
     also have "\<dots> = ennreal (pmf q i)"
     proof(cases i)
       case None
@@ -1180,7 +1171,7 @@
       also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (pmf pq (x, Some y)) * indicator (range Some) x \<partial>count_space UNIV) + pmf pq (None, Some y)"
         by(subst nn_integral_add)(simp_all)
       also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (spmf p y) * indicator {Some y} x \<partial>count_space UNIV) + (spmf q y - spmf p y)"
-        by(auto simp add: pq_def pmf_embed_pmf[OF f] one_ereal_def[symmetric] simp del: nn_integral_indicator_singleton intro!: arg_cong2[where f="(+)"] nn_integral_cong split: option.split)
+        by(auto simp: pq_def pmf_embed_pmf[OF f] one_ereal_def[symmetric] simp del: nn_integral_indicator_singleton intro!: arg_cong2[where f="(+)"] nn_integral_cong split: option.split)
       also have "\<dots> = spmf q y" by(simp add: ennreal_minus[symmetric] le)
       finally show ?thesis using Some by simp
     qed
@@ -1203,7 +1194,7 @@
   have "ennreal (spmf p x) = integral\<^sup>N pq (indicator (fst -` {Some x}))"
     using p by(simp add: ennreal_pmf_map)
   also have "\<dots> = integral\<^sup>N pq (indicator {(Some x, Some x)})"
-    by(rule nn_integral_cong_AE)(auto simp add: AE_measure_pmf_iff split: split_indicator dest: pq)
+    by(rule nn_integral_cong_AE)(auto simp: AE_measure_pmf_iff split: split_indicator dest: pq)
   also have "\<dots> \<le> integral\<^sup>N pq (indicator (snd -` {Some x}))"
     by(rule nn_integral_mono) simp
   also have "\<dots> = ennreal (spmf q x)" using q by(simp add: ennreal_pmf_map)
@@ -1211,11 +1202,11 @@
 qed
 
 lemma ord_spmf_eqD_set_spmf: "ord_spmf (=) p q \<Longrightarrow> set_spmf p \<subseteq> set_spmf q"
-by(rule subsetI)(drule_tac x=x in ord_spmf_eq_leD, auto simp add: in_set_spmf_iff_spmf)
+  by (metis ord_spmf_eq_leD pmf_le_0_iff spmf_eq_0_set_spmf subsetI)
 
 lemma ord_spmf_eqD_emeasure:
   "ord_spmf (=) p q \<Longrightarrow> emeasure (measure_spmf p) A \<le> emeasure (measure_spmf q) A"
-by(auto intro!: nn_integral_mono split: split_indicator dest: ord_spmf_eq_leD simp add: nn_integral_measure_spmf nn_integral_indicator[symmetric])
+  by(auto intro!: nn_integral_mono split: split_indicator dest: ord_spmf_eq_leD simp add: nn_integral_measure_spmf nn_integral_indicator[symmetric])
 
 lemma ord_spmf_eqD_measure_spmf: "ord_spmf (=) p q \<Longrightarrow> measure_spmf p \<le> measure_spmf q"
   by (subst le_measure) (auto simp: ord_spmf_eqD_emeasure)
@@ -1229,9 +1220,10 @@
   \<comment> \<open>We go through \<^typ>\<open>ennreal\<close> to have a sensible definition even if \<^term>\<open>Y\<close> is empty.\<close>
 
 lemma lub_spmf_empty [simp]: "SPMF.lub_spmf {} = return_pmf None"
-by(simp add: SPMF.lub_spmf_def bot_ereal_def)
-
-context assumes chain: "Complete_Partial_Order.chain (ord_spmf (=)) Y" begin
+  by(simp add: SPMF.lub_spmf_def bot_ereal_def)
+
+context assumes chain: "Complete_Partial_Order.chain (ord_spmf (=)) Y" 
+begin
 
 lemma chain_ord_spmf_eqD: "Complete_Partial_Order.chain (\<le>) ((\<lambda>p x. ennreal (spmf p x)) ` Y)"
   (is "Complete_Partial_Order.chain _ (?f ` _)")
@@ -1241,17 +1233,7 @@
   then obtain p q where f: "f = ?f p" "p \<in> Y" and g: "g = ?f q" "q \<in> Y" by blast
   from chain \<open>p \<in> Y\<close> \<open>q \<in> Y\<close> have "ord_spmf (=) p q \<or> ord_spmf (=) q p" by(rule chainD)
   thus "f \<le> g \<or> g \<le> f"
-  proof
-    assume "ord_spmf (=) p q"
-    hence "\<And>x. spmf p x \<le> spmf q x" by(rule ord_spmf_eq_leD)
-    hence "f \<le> g" unfolding f g by(auto intro: le_funI)
-    thus ?thesis ..
-  next
-    assume "ord_spmf (=) q p"
-    hence "\<And>x. spmf q x \<le> spmf p x" by(rule ord_spmf_eq_leD)
-    hence "g \<le> f" unfolding f g by(auto intro: le_funI)
-    thus ?thesis ..
-  qed
+    by (metis ennreal_leI f(1) g(1) le_funI ord_spmf_eq_leD)
 qed
 
 lemma ord_spmf_eq_pmf_None_eq:
@@ -1275,11 +1257,11 @@
 lemma ord_spmf_eqD_pmf_None:
   assumes "ord_spmf (=) x y"
   shows "pmf x None \<ge> pmf y None"
-using assms
-apply cases
-apply(clarsimp simp only: ennreal_le_iff[symmetric, OF pmf_nonneg] ennreal_pmf_map)
-apply(fastforce simp add: AE_measure_pmf_iff intro!: nn_integral_mono_AE)
-done
+  using assms
+  apply cases
+  apply(clarsimp simp only: ennreal_le_iff[symmetric, OF pmf_nonneg] ennreal_pmf_map)
+  apply(fastforce simp: AE_measure_pmf_iff intro!: nn_integral_mono_AE)
+  done
 
 text \<open>
   Chains on \<^typ>\<open>'a spmf\<close> maintain countable support.
@@ -1302,7 +1284,7 @@
       using chainD[OF chain, of x y] ord_spmf_eqD_pmf_None[of x y] ord_spmf_eqD_pmf_None[of y x]
       by (auto simp: N_def)
     have N_eq_imp_eq: "\<lbrakk> x \<in> Y; y \<in> Y; N y = N x \<rbrakk> \<Longrightarrow> x = y" for x y
-      using chainD[OF chain, of x y] by(auto simp add: N_def dest: ord_spmf_eq_pmf_None_eq)
+      using chainD[OF chain, of x y] by(auto simp: N_def dest: ord_spmf_eq_pmf_None_eq)
 
     have NC: "N ` Y \<noteq> {}" "bdd_below (N ` Y)"
       using \<open>Y \<noteq> {}\<close> by(auto intro!: bdd_belowI[of _ 0] simp: N_def)
@@ -1311,7 +1293,7 @@
       assume **: "\<not> (\<exists>y\<in>N ` Y. y < N x)"
       { fix y
         assume "y \<in> Y"
-        with ** consider "N x < N y" | "N x = N y" by(auto simp add: not_less le_less)
+        with ** consider "N x < N y" | "N x = N y" by(auto simp: not_less le_less)
         hence "ord_spmf (=) y x" using \<open>y \<in> Y\<close> \<open>x \<in> Y\<close>
           by cases(auto dest: N_less_imp_le_spmf N_eq_imp_eq intro: ord_spmf_reflI) }
       with False \<open>x \<in> Y\<close> show False by blast
@@ -1387,7 +1369,7 @@
 qed
 
 lemma ennreal_spmf_lub_spmf: "Y \<noteq> {} \<Longrightarrow> ennreal (spmf lub_spmf x) = (SUP p\<in>Y. ennreal (spmf p x))"
-unfolding spmf_lub_spmf by(subst ennreal_SUP)(simp_all add: SUP_spmf_neq_top' del: SUP_eq_top_iff Sup_eq_top_iff)
+  by (metis SUP_spmf_neq_top' ennreal_SUP spmf_lub_spmf)
 
 lemma lub_spmf_upper:
   assumes p: "p \<in> Y"
@@ -1397,7 +1379,7 @@
   from p have [simp]: "Y \<noteq> {}" by auto
   from p have "ennreal (spmf p x) \<le> (SUP p\<in>Y. ennreal (spmf p x))" by(rule SUP_upper)
   also have "\<dots> = ennreal (spmf lub_spmf x)" using p
-    by(subst spmf_lub_spmf)(auto simp add: ennreal_SUP SUP_spmf_neq_top' simp del: SUP_eq_top_iff Sup_eq_top_iff)
+    by(subst spmf_lub_spmf)(auto simp: ennreal_SUP SUP_spmf_neq_top' simp del: SUP_eq_top_iff Sup_eq_top_iff)
   finally show "spmf p x \<le> spmf lub_spmf x" by simp
 qed simp
 
@@ -1411,7 +1393,7 @@
     fix x
     from nonempty obtain p where p: "p \<in> Y" by auto
     have "ennreal (spmf lub_spmf x) = (SUP p\<in>Y. ennreal (spmf p x))"
-      by(subst spmf_lub_spmf)(auto simp add: ennreal_SUP SUP_spmf_neq_top' nonempty simp del: SUP_eq_top_iff Sup_eq_top_iff)
+      by(subst spmf_lub_spmf)(auto simp: ennreal_SUP SUP_spmf_neq_top' nonempty simp del: SUP_eq_top_iff Sup_eq_top_iff)
     also have "\<dots> \<le> ennreal (spmf z x)" by(rule SUP_least)(simp add: ord_spmf_eq_leD z)
     finally show "spmf lub_spmf x \<le> spmf z x" by simp
   qed simp
@@ -1428,7 +1410,7 @@
     also have "\<dots> \<longleftrightarrow> (\<exists>p\<in>Y. ennreal (spmf p x) > 0)"
       by(simp add: ennreal_spmf_lub_spmf less_SUP_iff)
     also have "\<dots> \<longleftrightarrow> x \<in> ?rhs"
-      by(auto simp add: in_set_spmf_iff_spmf less_le)
+      by(auto simp: in_set_spmf_iff_spmf less_le)
     finally show "x \<in> ?lhs \<longleftrightarrow> x \<in> ?rhs" .
   qed
 qed simp
@@ -1440,7 +1422,7 @@
 proof -
   let ?M = "count_space (set_spmf lub_spmf)"
   have "?lhs = \<integral>\<^sup>+ x. ennreal (spmf lub_spmf x) * indicator A x \<partial>?M"
-    by(auto simp add: nn_integral_indicator[symmetric] nn_integral_measure_spmf')
+    by(auto simp: nn_integral_indicator[symmetric] nn_integral_measure_spmf')
   also have "\<dots> = \<integral>\<^sup>+ x. (SUP y\<in>Y. ennreal (spmf y x) * indicator A x) \<partial>?M"
     unfolding ennreal_indicator[symmetric]
     by(simp add: spmf_lub_spmf assms ennreal_SUP[OF SUP_spmf_neq_top'] SUP_mult_right_ennreal)
@@ -1449,13 +1431,13 @@
     have "(\<lambda>i x. ennreal (spmf i x) * indicator A x) ` Y = (\<lambda>f x. f x * indicator A x) ` (\<lambda>p x. ennreal (spmf p x)) ` Y"
       by(simp add: image_image)
     also have "Complete_Partial_Order.chain (\<le>) \<dots>" using chain_ord_spmf_eqD
-      by(rule chain_imageI)(auto simp add: le_fun_def split: split_indicator)
+      by(rule chain_imageI)(auto simp: le_fun_def split: split_indicator)
     finally show "Complete_Partial_Order.chain (\<le>) ((\<lambda>i x. ennreal (spmf i x) * indicator A x) ` Y)" .
   qed simp
   also have "\<dots> = (SUP y\<in>Y. \<integral>\<^sup>+ x. ennreal (spmf y x) * indicator A x \<partial>count_space UNIV)"
-    by(auto simp add: nn_integral_count_space_indicator set_lub_spmf spmf_eq_0_set_spmf split: split_indicator intro!: arg_cong [of _ _ Sup] image_cong nn_integral_cong)
+    by(auto simp: nn_integral_count_space_indicator set_lub_spmf spmf_eq_0_set_spmf split: split_indicator intro!: arg_cong [of _ _ Sup] image_cong nn_integral_cong)
   also have "\<dots> = ?rhs"
-    by(auto simp add: nn_integral_indicator[symmetric] nn_integral_measure_spmf)
+    by(auto simp: nn_integral_indicator[symmetric] nn_integral_measure_spmf)
   finally show ?thesis .
 qed
 
@@ -1474,7 +1456,7 @@
 lemma weight_lub_spmf:
   assumes Y: "Y \<noteq> {}"
   shows "weight_spmf lub_spmf = (SUP y\<in>Y. weight_spmf y)"
-unfolding weight_spmf_def by(rule measure_lub_spmf) fact
+  by (smt (verit, best) SUP_cong assms measure_lub_spmf space_measure_spmf)
 
 lemma measure_spmf_lub_spmf:
   assumes Y: "Y \<noteq> {}"
@@ -1519,11 +1501,11 @@
 qed
 
 lemma ccpo_spmf: "class.ccpo lub_spmf (ord_spmf (=)) (mk_less (ord_spmf (=)))"
-by(rule ccpo partial_function_definitions_spmf)+
+  by(metis ccpo partial_function_definitions_spmf)
 
 interpretation spmf: partial_function_definitions "ord_spmf (=)" "lub_spmf"
   rewrites "lub_spmf {} \<equiv> return_pmf None"
-by(rule partial_function_definitions_spmf) simp
+  by(rule partial_function_definitions_spmf) simp
 
 declaration \<open>Partial_Function.init "spmf" \<^term>\<open>spmf.fixp_fun\<close>
   \<^term>\<open>spmf.mono_body\<close> @{thm spmf.fixp_rule_uc} @{thm spmf.fixp_induct_uc}
@@ -1535,14 +1517,14 @@
 abbreviation "mono_spmf \<equiv> monotone (fun_ord (ord_spmf (=))) (ord_spmf (=))"
 
 lemma lub_spmf_const [simp]: "lub_spmf {p} = p"
-by(rule spmf_eqI)(simp add: spmf_lub_spmf[OF ccpo.chain_singleton[OF ccpo_spmf]])
+  by(rule spmf_eqI)(simp add: spmf_lub_spmf[OF ccpo.chain_singleton[OF ccpo_spmf]])
 
 lemma bind_spmf_mono':
   assumes fg: "ord_spmf (=) f g"
-  and hk: "\<And>x :: 'a. ord_spmf (=) (h x) (k x)"
+    and hk: "\<And>x :: 'a. ord_spmf (=) (h x) (k x)"
   shows "ord_spmf (=) (f \<bind> h) (g \<bind> k)"
-unfolding bind_spmf_def using assms(1)
-by(rule rel_pmf_bindI)(auto split: option.split simp add: hk)
+  unfolding bind_spmf_def using assms(1)
+  by(rule rel_pmf_bindI)(auto split: option.split simp add: hk)
 
 lemma bind_spmf_mono [partial_function_mono]:
   assumes mf: "mono_spmf B" and mg: "\<And>y. mono_spmf (\<lambda>f. C y f)"
@@ -1558,12 +1540,12 @@
 qed
 
 lemma monotone_bind_spmf1: "monotone (ord_spmf (=)) (ord_spmf (=)) (\<lambda>y. bind_spmf y g)"
-by(rule monotoneI)(simp add: bind_spmf_mono' ord_spmf_reflI)
+  by(rule monotoneI)(simp add: bind_spmf_mono' ord_spmf_reflI)
 
 lemma monotone_bind_spmf2:
   assumes g: "\<And>x. monotone ord (ord_spmf (=)) (\<lambda>y. g y x)"
   shows "monotone ord (ord_spmf (=)) (\<lambda>y. bind_spmf p (g y))"
-by(rule monotoneI)(auto intro: bind_spmf_mono' monotoneD[OF g] ord_spmf_reflI)
+  by(rule monotoneI)(auto intro: bind_spmf_mono' monotoneD[OF g] ord_spmf_reflI)
 
 lemma bind_lub_spmf:
   assumes chain: "Complete_Partial_Order.chain (ord_spmf (=)) Y"
@@ -1574,18 +1556,18 @@
   proof(rule spmf_eqI)
     fix i
     have chain': "Complete_Partial_Order.chain (\<le>) ((\<lambda>p x. ennreal (spmf p x * spmf (f x) i)) ` Y)"
-      using chain by(rule chain_imageI)(auto simp add: le_fun_def dest: ord_spmf_eq_leD intro: mult_right_mono)
+      using chain by(rule chain_imageI)(auto simp: le_fun_def dest: ord_spmf_eq_leD intro: mult_right_mono)
     have chain'': "Complete_Partial_Order.chain (ord_spmf (=)) ((\<lambda>p. p \<bind> f) ` Y)"
       using chain by(rule chain_imageI)(auto intro!: monotoneI bind_spmf_mono' ord_spmf_reflI)
     let ?M = "count_space (set_spmf (lub_spmf Y))"
     have "ennreal (spmf ?lhs i) = \<integral>\<^sup>+ x. ennreal (spmf (lub_spmf Y) x) * ennreal (spmf (f x) i) \<partial>?M"
-      by(auto simp add: ennreal_spmf_lub_spmf ennreal_spmf_bind nn_integral_measure_spmf')
+      by(auto simp: ennreal_spmf_lub_spmf ennreal_spmf_bind nn_integral_measure_spmf')
     also have "\<dots> = \<integral>\<^sup>+ x. (SUP p\<in>Y. ennreal (spmf p x * spmf (f x) i)) \<partial>?M"
       by(subst ennreal_spmf_lub_spmf[OF chain Y])(subst SUP_mult_right_ennreal, simp_all add: ennreal_mult Y)
     also have "\<dots> = (SUP p\<in>Y. \<integral>\<^sup>+ x. ennreal (spmf p x * spmf (f x) i) \<partial>?M)"
       using Y chain' by(rule nn_integral_monotone_convergence_SUP_countable) simp
     also have "\<dots> = (SUP p\<in>Y. ennreal (spmf (bind_spmf p f) i))"
-      by(auto simp add: ennreal_spmf_bind nn_integral_measure_spmf nn_integral_count_space_indicator set_lub_spmf[OF chain] in_set_spmf_iff_spmf ennreal_mult intro!: arg_cong [of _ _ Sup] image_cong nn_integral_cong split: split_indicator)
+      by(auto simp: ennreal_spmf_bind nn_integral_measure_spmf nn_integral_count_space_indicator set_lub_spmf[OF chain] in_set_spmf_iff_spmf ennreal_mult intro!: arg_cong [of _ _ Sup] image_cong nn_integral_cong split: split_indicator)
     also have "\<dots> = ennreal (spmf ?rhs i)" using chain'' by(simp add: ennreal_spmf_lub_spmf Y image_comp)
     finally show "spmf ?lhs i = spmf ?rhs i" by simp
   qed
@@ -1594,16 +1576,17 @@
 lemma map_lub_spmf:
   "Complete_Partial_Order.chain (ord_spmf (=)) Y
   \<Longrightarrow> map_spmf f (lub_spmf Y) = lub_spmf (map_spmf f ` Y)"
-unfolding map_spmf_conv_bind_spmf[abs_def] by(simp add: bind_lub_spmf o_def)
+  unfolding map_spmf_conv_bind_spmf[abs_def] by(simp add: bind_lub_spmf o_def)
 
 lemma mcont_bind_spmf1: "mcont lub_spmf (ord_spmf (=)) lub_spmf (ord_spmf (=)) (\<lambda>y. bind_spmf y f)"
-using monotone_bind_spmf1 by(rule mcontI)(rule contI, simp add: bind_lub_spmf)
+  using monotone_bind_spmf1
+  by(intro contI mcontI) (auto simp: bind_lub_spmf)
 
 lemma bind_lub_spmf2:
   assumes chain: "Complete_Partial_Order.chain ord Y"
-  and g: "\<And>y. monotone ord (ord_spmf (=)) (g y)"
+    and g: "\<And>y. monotone ord (ord_spmf (=)) (g y)"
   shows "bind_spmf x (\<lambda>y. lub_spmf (g y ` Y)) = lub_spmf ((\<lambda>p. bind_spmf x (\<lambda>y. g y p)) ` Y)"
-  (is "?lhs = ?rhs")
+    (is "?lhs = ?rhs")
 proof(cases "Y = {}")
   case Y: False
   show ?thesis
@@ -1612,7 +1595,7 @@
     have chain': "\<And>y. Complete_Partial_Order.chain (ord_spmf (=)) (g y ` Y)"
       using chain g[THEN monotoneD] by(rule chain_imageI)
     have chain'': "Complete_Partial_Order.chain (\<le>) ((\<lambda>p y. ennreal (spmf x y * spmf (g y p) i)) ` Y)"
-      using chain by(rule chain_imageI)(auto simp add: le_fun_def dest: ord_spmf_eq_leD monotoneD[OF g] intro!: mult_left_mono)
+      using chain by(rule chain_imageI)(auto simp: le_fun_def dest: ord_spmf_eq_leD monotoneD[OF g] intro!: mult_left_mono)
     have chain''': "Complete_Partial_Order.chain (ord_spmf (=)) ((\<lambda>p. bind_spmf x (\<lambda>y. g y p)) ` Y)"
       using chain by(rule chain_imageI)(rule monotone_bind_spmf2[OF g, THEN monotoneD])
 
@@ -1624,7 +1607,7 @@
     also have "\<dots> = (SUP p\<in>Y. ennreal (spmf (bind_spmf x (\<lambda>y. g y p)) i))"
       by(simp add: ennreal_spmf_bind nn_integral_measure_spmf' ennreal_mult)
     also have "\<dots> = ennreal (spmf ?rhs i)" using chain'''
-      by(auto simp add: ennreal_spmf_lub_spmf Y image_comp)
+      by(auto simp: ennreal_spmf_lub_spmf Y image_comp)
     finally show "spmf ?lhs i = spmf ?rhs i" by simp
   qed
 qed simp
@@ -1656,44 +1639,44 @@
 
 lemma bind_pmf_mono [partial_function_mono]:
   "(\<And>y. mono_spmf (\<lambda>f. C y f)) \<Longrightarrow> mono_spmf (\<lambda>f. bind_pmf p (\<lambda>x. C x f))"
-using bind_spmf_mono[of "\<lambda>_. spmf_of_pmf p" C] by simp
+  using bind_spmf_mono[of "\<lambda>_. spmf_of_pmf p" C] by simp
 
 lemma map_spmf_mono [partial_function_mono]: "mono_spmf B \<Longrightarrow> mono_spmf (\<lambda>g. map_spmf f (B g))"
-unfolding map_spmf_conv_bind_spmf by(rule bind_spmf_mono) simp_all
+  unfolding map_spmf_conv_bind_spmf by(rule bind_spmf_mono) simp_all
 
 lemma mcont_map_spmf [cont_intro]:
   "mcont luba orda lub_spmf (ord_spmf (=)) g
   \<Longrightarrow> mcont luba orda lub_spmf (ord_spmf (=)) (\<lambda>x. map_spmf f (g x))"
-unfolding map_spmf_conv_bind_spmf by(rule mcont_bind_spmf) simp_all
+  unfolding map_spmf_conv_bind_spmf by(rule mcont_bind_spmf) simp_all
 
 lemma monotone_set_spmf: "monotone (ord_spmf (=)) (\<subseteq>) set_spmf"
-by(rule monotoneI)(rule ord_spmf_eqD_set_spmf)
+  by(rule monotoneI)(rule ord_spmf_eqD_set_spmf)
 
 lemma cont_set_spmf: "cont lub_spmf (ord_spmf (=)) Union (\<subseteq>) set_spmf"
-by(rule contI)(subst set_lub_spmf; simp)
+  by(rule contI)(subst set_lub_spmf; simp)
 
 lemma mcont2mcont_set_spmf[THEN mcont2mcont, cont_intro]:
   shows mcont_set_spmf: "mcont lub_spmf (ord_spmf (=)) Union (\<subseteq>) set_spmf"
-by(rule mcontI monotone_set_spmf cont_set_spmf)+
+  by(rule mcontI monotone_set_spmf cont_set_spmf)+
 
 lemma monotone_spmf: "monotone (ord_spmf (=)) (\<le>) (\<lambda>p. spmf p x)"
-by(rule monotoneI)(simp add: ord_spmf_eq_leD)
+  by(rule monotoneI)(simp add: ord_spmf_eq_leD)
 
 lemma cont_spmf: "cont lub_spmf (ord_spmf (=)) Sup (\<le>) (\<lambda>p. spmf p x)"
-by(rule contI)(simp add: spmf_lub_spmf)
+  by(rule contI)(simp add: spmf_lub_spmf)
 
 lemma mcont_spmf: "mcont lub_spmf (ord_spmf (=)) Sup (\<le>) (\<lambda>p. spmf p x)"
-by(rule mcontI monotone_spmf cont_spmf)+
+  by(metis mcontI monotone_spmf cont_spmf)
 
 lemma cont_ennreal_spmf: "cont lub_spmf (ord_spmf (=)) Sup (\<le>) (\<lambda>p. ennreal (spmf p x))"
-by(rule contI)(simp add: ennreal_spmf_lub_spmf)
+  by(rule contI)(simp add: ennreal_spmf_lub_spmf)
 
 lemma mcont2mcont_ennreal_spmf [THEN mcont2mcont, cont_intro]:
   shows mcont_ennreal_spmf: "mcont lub_spmf (ord_spmf (=)) Sup (\<le>) (\<lambda>p. ennreal (spmf p x))"
-by(rule mcontI mono2mono_ennreal monotone_spmf cont_ennreal_spmf)+
+  by(metis mcontI mono2mono_ennreal monotone_spmf cont_ennreal_spmf)
 
 lemma nn_integral_map_spmf [simp]: "nn_integral (measure_spmf (map_spmf f p)) g = nn_integral (measure_spmf p) (g \<circ> f)"
-by(auto 4 3 simp add: measure_spmf_def nn_integral_distr nn_integral_restrict_space intro: nn_integral_cong split: split_indicator)
+  by(force simp: measure_spmf_def nn_integral_distr nn_integral_restrict_space intro: nn_integral_cong split: split_indicator)
 
 subsubsection \<open>Admissibility of \<^term>\<open>rel_spmf\<close>\<close>
 
@@ -1705,7 +1688,7 @@
   also have "\<dots> \<le> measure (measure_pmf q) {y. \<exists>x\<in>Some ` A. rel_option R x y}"
     using assms by(rule rel_pmf_measureD)
   also have "\<dots> = ?rhs" unfolding measure_measure_spmf_conv_measure_pmf
-    by(rule arg_cong2[where f=measure])(auto simp add: option_rel_Some1)
+    by(rule arg_cong2[where f=measure])(auto simp: option_rel_Some1)
   finally show ?thesis .
 qed
 
@@ -1728,7 +1711,7 @@
   define A' where "A' = the ` (A \<inter> range Some)"
   define A'' where "A'' = A \<inter> {None}"
   have A: "A = Some ` A' \<union> A''" "Some ` A' \<inter> A'' = {}"
-    unfolding A'_def A''_def by(auto 4 3 intro: rev_image_eqI)
+    unfolding A'_def A''_def by(auto simp: image_iff)
   have "measure (measure_pmf p) A = measure (measure_pmf p) (Some ` A') + measure (measure_pmf p) A''"
     by(simp add: A measure_pmf.finite_measure_Union)
   also have "measure (measure_pmf p) (Some ` A') = measure (measure_spmf p) A'"
@@ -1737,7 +1720,7 @@
   also (ord_eq_le_trans[OF _ add_right_mono])
   have "\<dots> = measure (measure_pmf q) {y. \<exists>x\<in>A'. rel_option R (Some x) y}"
     unfolding measure_measure_spmf_conv_measure_pmf
-    by(rule arg_cong2[where f=measure])(auto simp add: A'_def option_rel_Some1)
+    by(rule arg_cong2[where f=measure])(auto simp: A'_def option_rel_Some1)
   also
   { have "weight_spmf p \<le> measure (measure_spmf q) {y. \<exists>x. R x y}"
       using eq1[of UNIV] unfolding weight_spmf_def by simp
@@ -1769,7 +1752,7 @@
   have "rel_spmf R (lub_spmf (fst ` Y)) (lub_spmf (snd ` Y))"
   proof(rule rel_spmf_measureI)
     show "weight_spmf (lub_spmf (snd ` Y)) \<le> weight_spmf (lub_spmf (fst ` Y))"
-      by(auto simp add: weight_lub_spmf chain1 chain2 Y rel_spmf_weightD[OF R, symmetric] intro!: cSUP_least intro: cSUP_upper2[OF bdd_aboveI2[OF weight_spmf_le_1]])
+      by(auto simp: weight_lub_spmf chain1 chain2 Y rel_spmf_weightD[OF R, symmetric] intro!: cSUP_least intro: cSUP_upper2[OF bdd_aboveI2[OF weight_spmf_le_1]])
 
     fix A
     have "measure (measure_spmf (lub_spmf (fst ` Y))) A = (SUP y\<in>fst ` Y. measure (measure_spmf y) A)"
@@ -1786,17 +1769,17 @@
 lemma admissible_rel_spmf_mcont [cont_intro]:
   "\<lbrakk> mcont lub ord lub_spmf (ord_spmf (=)) f; mcont lub ord lub_spmf (ord_spmf (=)) g \<rbrakk>
   \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. rel_spmf R (f x) (g x))"
-by(rule admissible_subst[OF admissible_rel_spmf, where f="\<lambda>x. (f x, g x)", simplified])(rule mcont_Pair)
+  by(rule admissible_subst[OF admissible_rel_spmf, where f="\<lambda>x. (f x, g x)", simplified])(rule mcont_Pair)
 
 context includes lifting_syntax
 begin
 
 lemma fixp_spmf_parametric':
   assumes f: "\<And>x. monotone (ord_spmf (=)) (ord_spmf (=)) F"
-  and g: "\<And>x. monotone (ord_spmf (=)) (ord_spmf (=)) G"
-  and param: "(rel_spmf R ===> rel_spmf R) F G"
+    and g: "\<And>x. monotone (ord_spmf (=)) (ord_spmf (=)) G"
+    and param: "(rel_spmf R ===> rel_spmf R) F G"
   shows "(rel_spmf R) (ccpo.fixp lub_spmf (ord_spmf (=)) F) (ccpo.fixp lub_spmf (ord_spmf (=)) G)"
-by(rule parallel_fixp_induct[OF ccpo_spmf ccpo_spmf _ f g])(auto intro: param[THEN rel_funD])
+  by(rule parallel_fixp_induct[OF ccpo_spmf ccpo_spmf _ f g])(auto intro: param[THEN rel_funD])
 
 lemma fixp_spmf_parametric:
   assumes f: "\<And>x. mono_spmf (\<lambda>f. F f x)"
@@ -1807,13 +1790,9 @@
 proof(rule parallel_fixp_induct_1_1[OF partial_function_definitions_spmf partial_function_definitions_spmf _ _ reflexive reflexive, where P="(A ===> rel_spmf R)"])
   show "ccpo.admissible (prod_lub (fun_lub lub_spmf) (fun_lub lub_spmf)) (rel_prod (fun_ord (ord_spmf (=))) (fun_ord (ord_spmf (=)))) (\<lambda>x. (A ===> rel_spmf R) (fst x) (snd x))"
     unfolding rel_fun_def
-    apply(rule admissible_all admissible_imp admissible_rel_spmf_mcont)+
-    apply(rule spmf.mcont2mcont[OF mcont_call])
-     apply(rule mcont_fst)
-    apply(rule spmf.mcont2mcont[OF mcont_call])
-     apply(rule mcont_snd)
-    done
-  show "(A ===> rel_spmf R) (\<lambda>_. lub_spmf {}) (\<lambda>_. lub_spmf {})" by auto
+    by(fastforce intro: admissible_all admissible_imp admissible_rel_spmf_mcont)
+  show "(A ===> rel_spmf R) (\<lambda>_. lub_spmf {}) (\<lambda>_. lub_spmf {})" 
+    by auto
   show "(A ===> rel_spmf R) (F f) (G g)" if "(A ===> rel_spmf R) f g" for f g
     using that by(rule rel_funD[OF param])
 qed
@@ -1827,40 +1806,46 @@
 subsection \<open>Restrictions on spmfs\<close>
 
 definition restrict_spmf :: "'a spmf \<Rightarrow> 'a set \<Rightarrow> 'a spmf" (infixl "\<upharpoonleft>" 110)
-where "p \<upharpoonleft> A = map_pmf (\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) p"
+  where "p \<upharpoonleft> A = map_pmf (\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) p"
 
 lemma set_restrict_spmf [simp]: "set_spmf (p \<upharpoonleft> A) = set_spmf p \<inter> A"
-by(fastforce simp add: restrict_spmf_def set_spmf_def split: bind_splits if_split_asm)
+  by(fastforce simp: restrict_spmf_def set_spmf_def split: bind_splits if_split_asm)
 
 lemma restrict_map_spmf: "map_spmf f p \<upharpoonleft> A = map_spmf f (p \<upharpoonleft> (f -` A))"
-by(simp add: restrict_spmf_def pmf.map_comp o_def map_option_bind bind_map_option if_distrib cong del: if_weak_cong)
+  by(simp add: restrict_spmf_def pmf.map_comp o_def map_option_bind bind_map_option if_distrib cong del: if_weak_cong)
 
 lemma restrict_restrict_spmf [simp]: "p \<upharpoonleft> A \<upharpoonleft> B = p \<upharpoonleft> (A \<inter> B)"
-by(auto simp add: restrict_spmf_def pmf.map_comp o_def intro!: pmf.map_cong bind_option_cong)
+  by(auto simp: restrict_spmf_def pmf.map_comp o_def intro!: pmf.map_cong bind_option_cong)
 
 lemma restrict_spmf_empty [simp]: "p \<upharpoonleft> {} = return_pmf None"
-by(simp add: restrict_spmf_def)
+  by(simp add: restrict_spmf_def)
 
 lemma restrict_spmf_UNIV [simp]: "p \<upharpoonleft> UNIV = p"
-by(simp add: restrict_spmf_def)
+  by(simp add: restrict_spmf_def)
 
 lemma spmf_restrict_spmf_outside [simp]: "x \<notin> A \<Longrightarrow> spmf (p \<upharpoonleft> A) x = 0"
-by(simp add: spmf_eq_0_set_spmf)
-
-lemma emeasure_restrict_spmf [simp]:
-  "emeasure (measure_spmf (p \<upharpoonleft> A)) X = emeasure (measure_spmf p) (X \<inter> A)"
-by(auto simp add: restrict_spmf_def measure_spmf_def emeasure_distr measurable_restrict_space1 emeasure_restrict_space space_restrict_space intro: arg_cong2[where f=emeasure] split: bind_splits if_split_asm)
+  by(simp add: spmf_eq_0_set_spmf)
+
+lemma emeasure_restrict_spmf [simp]: "emeasure (measure_spmf (p \<upharpoonleft> A)) X = emeasure (measure_spmf p) (X \<inter> A)"
+proof -
+  have "(\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) -` the -` X \<inter>
+        (\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) -` range Some =
+        the -` X \<inter> the -` A \<inter> range Some"
+    by(auto split: bind_splits if_split_asm)
+  then show ?thesis
+    by (simp add: restrict_spmf_def measure_spmf_def emeasure_distr emeasure_restrict_space)
+qed
 
 lemma measure_restrict_spmf [simp]:
   "measure (measure_spmf (p \<upharpoonleft> A)) X = measure (measure_spmf p) (X \<inter> A)"
-using emeasure_restrict_spmf[of p A X]
-by(simp only: measure_spmf.emeasure_eq_measure ennreal_inj measure_nonneg)
+  using emeasure_restrict_spmf[of p A X]
+  by(simp only: measure_spmf.emeasure_eq_measure ennreal_inj measure_nonneg)
 
 lemma spmf_restrict_spmf: "spmf (p \<upharpoonleft> A) x = (if x \<in> A then spmf p x else 0)"
-by(simp add: spmf_conv_measure_spmf)
+  by(simp add: spmf_conv_measure_spmf)
 
 lemma spmf_restrict_spmf_inside [simp]: "x \<in> A \<Longrightarrow> spmf (p \<upharpoonleft> A) x = spmf p x"
-by(simp add: spmf_restrict_spmf)
+  by(simp add: spmf_restrict_spmf)
 
 lemma pmf_restrict_spmf_None: "pmf (p \<upharpoonleft> A) None = pmf p None + measure (measure_spmf p) (- A)"
 proof -
@@ -1873,37 +1858,37 @@
 qed
 
 lemma restrict_spmf_trivial: "(\<And>x. x \<in> set_spmf p \<Longrightarrow> x \<in> A) \<Longrightarrow> p \<upharpoonleft> A = p"
-by(rule spmf_eqI)(auto simp add: spmf_restrict_spmf spmf_eq_0_set_spmf)
+  by(rule spmf_eqI)(auto simp: spmf_restrict_spmf spmf_eq_0_set_spmf)
 
 lemma restrict_spmf_trivial': "set_spmf p \<subseteq> A \<Longrightarrow> p \<upharpoonleft> A = p"
-by(rule restrict_spmf_trivial) blast
+  by(rule restrict_spmf_trivial) blast
 
 lemma restrict_return_spmf: "return_spmf x \<upharpoonleft> A = (if x \<in> A then return_spmf x else return_pmf None)"
-by(simp add: restrict_spmf_def)
+  by(simp add: restrict_spmf_def)
 
 lemma restrict_return_spmf_inside [simp]: "x \<in> A \<Longrightarrow> return_spmf x \<upharpoonleft> A = return_spmf x"
-by(simp add: restrict_return_spmf)
+  by(simp add: restrict_return_spmf)
 
 lemma restrict_return_spmf_outside [simp]: "x \<notin> A \<Longrightarrow> return_spmf x \<upharpoonleft> A = return_pmf None"
-by(simp add: restrict_return_spmf)
+  by(simp add: restrict_return_spmf)
 
 lemma restrict_spmf_return_pmf_None [simp]: "return_pmf None \<upharpoonleft> A = return_pmf None"
-by(simp add: restrict_spmf_def)
+  by(simp add: restrict_spmf_def)
 
 lemma restrict_bind_pmf: "bind_pmf p g \<upharpoonleft> A = p \<bind> (\<lambda>x. g x \<upharpoonleft> A)"
-by(simp add: restrict_spmf_def map_bind_pmf o_def)
+  by(simp add: restrict_spmf_def map_bind_pmf o_def)
 
 lemma restrict_bind_spmf: "bind_spmf p g \<upharpoonleft> A = p \<bind> (\<lambda>x. g x \<upharpoonleft> A)"
-by(auto simp add: bind_spmf_def restrict_bind_pmf cong del: option.case_cong_weak cong: option.case_cong intro!: bind_pmf_cong split: option.split)
+  by(auto simp: bind_spmf_def restrict_bind_pmf cong del: option.case_cong_weak cong: option.case_cong intro!: bind_pmf_cong split: option.split)
 
 lemma bind_restrict_pmf: "bind_pmf (p \<upharpoonleft> A) g = p \<bind> (\<lambda>x. if x \<in> Some ` A then g x else g None)"
-by(auto simp add: restrict_spmf_def bind_map_pmf fun_eq_iff split: bind_split intro: arg_cong2[where f=bind_pmf])
+  by(auto simp: restrict_spmf_def bind_map_pmf fun_eq_iff split: bind_split intro: arg_cong2[where f=bind_pmf])
 
 lemma bind_restrict_spmf: "bind_spmf (p \<upharpoonleft> A) g = p \<bind> (\<lambda>x. if x \<in> A then g x else return_pmf None)"
-by(auto simp add: bind_spmf_def bind_restrict_pmf fun_eq_iff intro: arg_cong2[where f=bind_pmf] split: option.split)
+  by(auto simp: bind_spmf_def bind_restrict_pmf fun_eq_iff intro: arg_cong2[where f=bind_pmf] split: option.split)
 
 lemma spmf_map_restrict: "spmf (map_spmf fst (p \<upharpoonleft> (snd -` {y}))) x = spmf p (x, y)"
-by(subst spmf_map)(auto intro: arg_cong2[where f=measure] simp add: spmf_conv_measure_spmf)
+  by(subst spmf_map)(auto intro: arg_cong2[where f=measure] simp add: spmf_conv_measure_spmf)
 
 lemma measure_eqI_restrict_spmf:
   assumes "rel_spmf R (restrict_spmf p A) (restrict_spmf q B)"
@@ -1920,55 +1905,55 @@
   "spmf_of_set A = (if finite A \<and> A \<noteq> {} then spmf_of_pmf (pmf_of_set A) else return_pmf None)"
 
 lemma spmf_of_set: "spmf (spmf_of_set A) x = indicator A x / card A"
-by(auto simp add: spmf_of_set_def)
+  by(auto simp: spmf_of_set_def)
 
 lemma pmf_spmf_of_set_None [simp]: "pmf (spmf_of_set A) None = indicator {A. infinite A \<or> A = {}} A"
-by(simp add: spmf_of_set_def)
+  by(simp add: spmf_of_set_def)
 
 lemma set_spmf_of_set: "set_spmf (spmf_of_set A) = (if finite A then A else {})"
-by(simp add: spmf_of_set_def)
+  by(simp add: spmf_of_set_def)
 
 lemma set_spmf_of_set_finite [simp]: "finite A \<Longrightarrow> set_spmf (spmf_of_set A) = A"
-by(simp add: set_spmf_of_set)
+  by(simp add: set_spmf_of_set)
 
 lemma spmf_of_set_singleton: "spmf_of_set {x} = return_spmf x"
-by(simp add: spmf_of_set_def pmf_of_set_singleton)
+  by(simp add: spmf_of_set_def pmf_of_set_singleton)
 
 lemma map_spmf_of_set_inj_on [simp]:
   "inj_on f A \<Longrightarrow> map_spmf f (spmf_of_set A) = spmf_of_set (f ` A)"
-by(auto simp add: spmf_of_set_def map_pmf_of_set_inj dest: finite_imageD)
+  by(auto simp: spmf_of_set_def map_pmf_of_set_inj dest: finite_imageD)
 
 lemma spmf_of_pmf_pmf_of_set [simp]:
   "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> spmf_of_pmf (pmf_of_set A) = spmf_of_set A"
-by(simp add: spmf_of_set_def)
+  by(simp add: spmf_of_set_def)
 
 lemma weight_spmf_of_set:
   "weight_spmf (spmf_of_set A) = (if finite A \<and> A \<noteq> {} then 1 else 0)"
-by(auto simp only: spmf_of_set_def weight_spmf_of_pmf weight_return_pmf_None split: if_split)
+  by(auto simp only: spmf_of_set_def weight_spmf_of_pmf weight_return_pmf_None split: if_split)
 
 lemma weight_spmf_of_set_finite [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> weight_spmf (spmf_of_set A) = 1"
-by(simp add: weight_spmf_of_set)
+  by(simp add: weight_spmf_of_set)
 
 lemma weight_spmf_of_set_infinite [simp]: "infinite A \<Longrightarrow> weight_spmf (spmf_of_set A) = 0"
-by(simp add: weight_spmf_of_set)
+  by(simp add: weight_spmf_of_set)
 
 lemma measure_spmf_spmf_of_set:
   "measure_spmf (spmf_of_set A) = (if finite A \<and> A \<noteq> {} then measure_pmf (pmf_of_set A) else null_measure (count_space UNIV))"
-by(simp add: spmf_of_set_def del: spmf_of_pmf_pmf_of_set)
+  by(simp add: spmf_of_set_def del: spmf_of_pmf_pmf_of_set)
 
 lemma emeasure_spmf_of_set:
   "emeasure (measure_spmf (spmf_of_set S)) A = card (S \<inter> A) / card S"
-by(auto simp add: measure_spmf_spmf_of_set emeasure_pmf_of_set)
+  by(auto simp: measure_spmf_spmf_of_set emeasure_pmf_of_set)
 
 lemma measure_spmf_of_set:
   "measure (measure_spmf (spmf_of_set S)) A = card (S \<inter> A) / card S"
-by(auto simp add: measure_spmf_spmf_of_set measure_pmf_of_set)
+  by(auto simp: measure_spmf_spmf_of_set measure_pmf_of_set)
 
 lemma nn_integral_spmf_of_set: "nn_integral (measure_spmf (spmf_of_set A)) f = sum f A / card A"
-by(cases "finite A")(auto simp add: spmf_of_set_def nn_integral_pmf_of_set card_gt_0_iff simp del: spmf_of_pmf_pmf_of_set)
+  by(cases "finite A")(auto simp: spmf_of_set_def nn_integral_pmf_of_set card_gt_0_iff simp del: spmf_of_pmf_pmf_of_set)
 
 lemma integral_spmf_of_set: "integral\<^sup>L (measure_spmf (spmf_of_set A)) f = sum f A / card A"
-by(clarsimp simp add: spmf_of_set_def integral_pmf_of_set card_gt_0_iff simp del: spmf_of_pmf_pmf_of_set)
+  by (metis card.infinite div_0 division_ring_divide_zero integral_null_measure integral_pmf_of_set measure_spmf_spmf_of_set of_nat_0 sum.empty)
 
 notepad begin \<comment> \<open>\<^const>\<open>pmf_of_set\<close> is not fully parametric.\<close>
   define R :: "nat \<Rightarrow> nat \<Rightarrow> bool" where "R x y \<longleftrightarrow> (x \<noteq> 0 \<longrightarrow> y = 0)" for x y
@@ -1991,7 +1976,7 @@
     also have "\<dots> = emeasure (measure_pmf pq) (snd -` {2, 1})"
       unfolding 2[symmetric] measure_pmf.emeasure_eq_measure[symmetric] by(simp)
     also have "\<dots> = emeasure (measure_pmf pq) {(0, 2), (0, 1)}"
-      by(rule emeasure_eq_AE)(auto simp add: AE_measure_pmf_iff R_def dest!: pq)
+      by(rule emeasure_eq_AE)(auto simp: AE_measure_pmf_iff R_def dest!: pq)
     also have "\<dots> \<le> emeasure (measure_pmf pq) (fst -` {0})"
       by(rule emeasure_mono) auto
     also have "\<dots> = emeasure (measure_pmf (pmf_of_set A)) {0}"
@@ -2013,7 +1998,7 @@
   define AB where "AB = (\<lambda>x. (x, f x)) ` A"
   define R' where "R' x y \<longleftrightarrow> (x, y) \<in> AB" for x y
   have "(x, y) \<in> AB" if "(x, y) \<in> set_pmf (pmf_of_set AB)" for x y
-    using that by(auto simp add: AB_def A)
+    using that by(auto simp: AB_def A)
   moreover have "map_pmf fst (pmf_of_set AB) = pmf_of_set A"
     by(simp add: AB_def map_pmf_of_set_inj[symmetric] inj_on_def A pmf.map_comp o_def)
   moreover
@@ -2030,10 +2015,11 @@
   and R: "\<And>x. x \<in> A \<Longrightarrow> R x (f x)"
   shows "rel_spmf R (spmf_of_set A) (spmf_of_set B)"
 proof -
-  have "finite A \<longleftrightarrow> finite B" using f by(rule bij_betw_finite)
-  moreover have "A = {} \<longleftrightarrow> B = {}" using f by(auto dest: bij_betw_empty2 bij_betw_empty1)
-  ultimately show ?thesis using assms
-    by(auto simp add: spmf_of_set_def simp del: spmf_of_pmf_pmf_of_set intro: rel_pmf_of_set_bij)
+  obtain "finite A \<longleftrightarrow> finite B" "A = {} \<longleftrightarrow> B = {}"
+    using bij_betw_empty1 bij_betw_empty2 bij_betw_finite f by blast 
+  then show ?thesis 
+    using assms
+    by (metis rel_pmf_of_set_bij rel_spmf_spmf_of_pmf return_spmf_None_parametric spmf_of_set_def)
 qed
 
 context includes lifting_syntax
@@ -2047,7 +2033,8 @@
   assume R: "rel_set R A B"
   with assms obtain f where "bij_betw f A B" and f: "\<And>x. x \<in> A \<Longrightarrow> R x (f x)"
     by(auto dest: bi_unique_rel_set_bij_betw)
-  then show "rel_spmf R (spmf_of_set A) (spmf_of_set B)" by(rule rel_spmf_of_set_bij)
+  then show "rel_spmf R (spmf_of_set A) (spmf_of_set B)" 
+    by(rule rel_spmf_of_set_bij)
 qed
 
 end
@@ -2063,7 +2050,7 @@
   also have "\<dots> = (if i then card (B \<inter> A) / card B else card (B - A) / card B)"
     by(auto intro: arg_cong[where f=card])
   also have "\<dots> = (if i then card (B \<inter> A) / card B else (card B - card (B \<inter> A)) / card B)"
-    by(auto simp add: card_Diff_subset_Int assms)
+    by(auto simp: card_Diff_subset_Int assms)
   also have "\<dots> = ennreal (spmf ?rhs i)"
     by(simp add: assms card_gt_0_iff field_simps card_mono Int_commute of_nat_diff)
   finally show "spmf ?lhs i = spmf ?rhs i" by simp
@@ -2079,48 +2066,48 @@
 qed
 
 lemma bind_coin_spmf_eq_const: "coin_spmf \<bind> (\<lambda>x :: bool. return_spmf (b = x)) = coin_spmf"
-using map_eq_const_coin_spmf unfolding map_spmf_conv_bind_spmf by simp
+  using map_eq_const_coin_spmf unfolding map_spmf_conv_bind_spmf by simp
 
 lemma bind_coin_spmf_eq_const': "coin_spmf \<bind> (\<lambda>x :: bool. return_spmf (x = b)) = coin_spmf"
-by(rewrite in "_ = \<hole>" bind_coin_spmf_eq_const[symmetric, of b])(auto intro: bind_spmf_cong)
+  by(rewrite in "_ = \<hole>" bind_coin_spmf_eq_const[symmetric, of b])(auto intro: bind_spmf_cong)
 
 subsection \<open>Losslessness\<close>
 
 definition lossless_spmf :: "'a spmf \<Rightarrow> bool"
-where "lossless_spmf p \<longleftrightarrow> weight_spmf p = 1"
+  where "lossless_spmf p \<longleftrightarrow> weight_spmf p = 1"
 
 lemma lossless_iff_pmf_None: "lossless_spmf p \<longleftrightarrow> pmf p None = 0"
-by(simp add: lossless_spmf_def pmf_None_eq_weight_spmf)
+  by(simp add: lossless_spmf_def pmf_None_eq_weight_spmf)
 
 lemma lossless_return_spmf [iff]: "lossless_spmf (return_spmf x)"
-by(simp add: lossless_iff_pmf_None)
+  by(simp add: lossless_iff_pmf_None)
 
 lemma lossless_return_pmf_None [iff]: "\<not> lossless_spmf (return_pmf None)"
-by(simp add: lossless_iff_pmf_None)
+  by(simp add: lossless_iff_pmf_None)
 
 lemma lossless_map_spmf [simp]: "lossless_spmf (map_spmf f p) \<longleftrightarrow> lossless_spmf p"
-by(auto simp add: lossless_iff_pmf_None pmf_eq_0_set_pmf)
+  by(auto simp: lossless_iff_pmf_None pmf_eq_0_set_pmf)
 
 lemma lossless_bind_spmf [simp]:
   "lossless_spmf (p \<bind> f) \<longleftrightarrow> lossless_spmf p \<and> (\<forall>x\<in>set_spmf p. lossless_spmf (f x))"
-by(simp add: lossless_iff_pmf_None pmf_bind_spmf_None add_nonneg_eq_0_iff integral_nonneg_AE integral_nonneg_eq_0_iff_AE measure_spmf.integrable_const_bound[where B=1] pmf_le_1)
+  by(simp add: lossless_iff_pmf_None pmf_bind_spmf_None add_nonneg_eq_0_iff integral_nonneg_AE integral_nonneg_eq_0_iff_AE measure_spmf.integrable_const_bound[where B=1] pmf_le_1)
 
 lemma lossless_weight_spmfD: "lossless_spmf p \<Longrightarrow> weight_spmf p = 1"
-by(simp add: lossless_spmf_def)
+  by(simp add: lossless_spmf_def)
 
 lemma lossless_iff_set_pmf_None:
   "lossless_spmf p \<longleftrightarrow> None \<notin> set_pmf p"
-by (simp add: lossless_iff_pmf_None pmf_eq_0_set_pmf)
+  by (simp add: lossless_iff_pmf_None pmf_eq_0_set_pmf)
 
 lemma lossless_spmf_of_set [simp]: "lossless_spmf (spmf_of_set A) \<longleftrightarrow> finite A \<and> A \<noteq> {}"
-by(auto simp add: lossless_spmf_def weight_spmf_of_set)
+  by(auto simp: lossless_spmf_def weight_spmf_of_set)
 
 lemma lossless_spmf_spmf_of_spmf [simp]: "lossless_spmf (spmf_of_pmf p)"
-by(simp add: lossless_spmf_def)
+  by(simp add: lossless_spmf_def)
 
 lemma lossless_spmf_bind_pmf [simp]:
   "lossless_spmf (bind_pmf p f) \<longleftrightarrow> (\<forall>x\<in>set_pmf p. lossless_spmf (f x))"
-by(simp add: lossless_iff_pmf_None pmf_bind integral_nonneg_AE integral_nonneg_eq_0_iff_AE measure_pmf.integrable_const_bound[where B=1] AE_measure_pmf_iff pmf_le_1)
+  by(simp add: lossless_iff_pmf_None pmf_bind integral_nonneg_AE integral_nonneg_eq_0_iff_AE measure_pmf.integrable_const_bound[where B=1] AE_measure_pmf_iff pmf_le_1)
 
 lemma lossless_spmf_conv_spmf_of_pmf: "lossless_spmf p \<longleftrightarrow> (\<exists>p'. p = spmf_of_pmf p')"
 proof
@@ -2134,7 +2121,7 @@
     fix i
     have "ennreal (pmf (map_pmf the p) i) = \<integral>\<^sup>+ x. indicator (the -` {i}) x \<partial>p" by(simp add: ennreal_pmf_map)
     also have "\<dots> = \<integral>\<^sup>+ x. indicator {i} x \<partial>measure_spmf p" unfolding measure_spmf_def
-      by(subst nn_integral_distr)(auto simp add: nn_integral_restrict_space AE_measure_pmf_iff simp del: nn_integral_indicator intro!: nn_integral_cong_AE split: split_indicator dest!: * )
+      by(subst nn_integral_distr)(auto simp: nn_integral_restrict_space AE_measure_pmf_iff simp del: nn_integral_indicator intro!: nn_integral_cong_AE split: split_indicator dest!: * )
     also have "\<dots> = spmf p i" by(simp add: emeasure_spmf_single)
     finally show "spmf p i = spmf (spmf_of_pmf ?p) i" by simp
   qed
@@ -2142,22 +2129,24 @@
 qed auto
 
 lemma spmf_False_conv_True: "lossless_spmf p \<Longrightarrow> spmf p False = 1 - spmf p True"
-by(clarsimp simp add: lossless_spmf_conv_spmf_of_pmf pmf_False_conv_True)
+  by(clarsimp simp add: lossless_spmf_conv_spmf_of_pmf pmf_False_conv_True)
 
 lemma spmf_True_conv_False: "lossless_spmf p \<Longrightarrow> spmf p True = 1 - spmf p False"
-by(simp add: spmf_False_conv_True)
+  by(simp add: spmf_False_conv_True)
 
 lemma bind_eq_return_spmf:
   "bind_spmf p f = return_spmf x \<longleftrightarrow> (\<forall>y\<in>set_spmf p. f y = return_spmf x) \<and> lossless_spmf p"
-by(auto simp add: bind_spmf_def bind_eq_return_pmf in_set_spmf lossless_iff_pmf_None pmf_eq_0_set_pmf iff del: not_None_eq split: option.split)
+  apply (simp add: bind_spmf_def bind_eq_return_pmf split: option.split)
+  by (metis in_set_spmf lossless_iff_set_pmf_None not_None_eq)
 
 lemma rel_spmf_return_spmf2:
   "rel_spmf R p (return_spmf x) \<longleftrightarrow> lossless_spmf p \<and> (\<forall>a\<in>set_spmf p. R a x)"
-by(auto simp add: lossless_iff_set_pmf_None rel_pmf_return_pmf2 option_rel_Some2 in_set_spmf, metis in_set_spmf not_None_eq)
+  apply (simp add: lossless_iff_set_pmf_None rel_pmf_return_pmf2 option_rel_Some2 in_set_spmf)
+  by (metis in_set_spmf not_None_eq option.sel)
 
 lemma rel_spmf_return_spmf1:
   "rel_spmf R (return_spmf x) p \<longleftrightarrow> lossless_spmf p \<and> (\<forall>a\<in>set_spmf p. R x a)"
-using rel_spmf_return_spmf2[of "R\<inverse>\<inverse>"] by(simp add: spmf_rel_conversep)
+  using rel_spmf_return_spmf2[of "R\<inverse>\<inverse>"] by(simp add: spmf_rel_conversep)
 
 lemma rel_spmf_bindI1:
   assumes f: "\<And>x. x \<in> set_spmf p \<Longrightarrow> rel_spmf R (f x) q"
@@ -2173,7 +2162,7 @@
 lemma rel_spmf_bindI2:
   "\<lbrakk> \<And>x. x \<in> set_spmf q \<Longrightarrow> rel_spmf R p (f x); lossless_spmf q \<rbrakk>
   \<Longrightarrow> rel_spmf R p (bind_spmf q f)"
-using rel_spmf_bindI1[of q "conversep R" f p] by(simp add: spmf_rel_conversep)
+  using rel_spmf_bindI1[of q "conversep R" f p] by(simp add: spmf_rel_conversep)
 
 subsection \<open>Scaling\<close>
 
@@ -2192,71 +2181,72 @@
 qed
 
 lemma spmf_scale_spmf: "spmf (scale_spmf r p) x = max 0 (min (inverse (weight_spmf p)) r) * spmf p x" (is "?lhs = ?rhs")
-unfolding scale_spmf_def
-apply(subst spmf_embed_spmf[OF scale_spmf_le_1])
-apply(simp add: max_def min_def weight_spmf_le_0 field_simps weight_spmf_nonneg not_le order.strict_iff_order)
-apply(metis antisym_conv order_trans weight_spmf_nonneg zero_le_mult_iff zero_le_one)
-done
+  unfolding scale_spmf_def
+  apply(subst spmf_embed_spmf[OF scale_spmf_le_1])
+  apply(simp add: max_def min_def measure_le_0_iff field_simps weight_spmf_nonneg not_le order.strict_iff_order)
+  apply(metis antisym_conv order_trans weight_spmf_nonneg zero_le_mult_iff zero_le_one)
+  done
 
 lemma real_inverse_le_1_iff: fixes x :: real
   shows "\<lbrakk> 0 \<le> x; x \<le> 1 \<rbrakk> \<Longrightarrow> 1 / x \<le> 1 \<longleftrightarrow> x = 1 \<or> x = 0"
-by auto
+  by auto
 
 lemma spmf_scale_spmf': "r \<le> 1 \<Longrightarrow> spmf (scale_spmf r p) x = max 0 r * spmf p x"
-using real_inverse_le_1_iff[OF weight_spmf_nonneg weight_spmf_le_1, of p]
-by(auto simp add: spmf_scale_spmf max_def min_def field_simps)(metis pmf_le_0_iff spmf_le_weight)
+  using real_inverse_le_1_iff[OF weight_spmf_nonneg weight_spmf_le_1, of p]
+  by(auto simp: spmf_scale_spmf max_def min_def field_simps)(metis pmf_le_0_iff spmf_le_weight)
 
 lemma scale_spmf_neg: "r \<le> 0 \<Longrightarrow> scale_spmf r p = return_pmf None"
-by(rule spmf_eqI)(simp add: spmf_scale_spmf' max_def)
+  by(rule spmf_eqI)(simp add: spmf_scale_spmf' max_def)
 
 lemma scale_spmf_return_None [simp]: "scale_spmf r (return_pmf None) = return_pmf None"
-by(rule spmf_eqI)(simp add: spmf_scale_spmf)
+  by(rule spmf_eqI)(simp add: spmf_scale_spmf)
 
 lemma scale_spmf_conv_bind_bernoulli:
   assumes "r \<le> 1"
   shows "scale_spmf r p = bind_pmf (bernoulli_pmf r) (\<lambda>b. if b then p else return_pmf None)" (is "?lhs = ?rhs")
 proof(rule spmf_eqI)
   fix x
-  have "ennreal (spmf ?lhs x) = ennreal (spmf ?rhs x)" using assms
+  have "\<lbrakk>weight_spmf p = 0\<rbrakk> \<Longrightarrow> spmf p x = 0"
+    by (metis pmf_le_0_iff spmf_le_weight)
+  moreover have "\<lbrakk>weight_spmf p \<noteq> 0; 1 / weight_spmf p < 1\<rbrakk> \<Longrightarrow> weight_spmf p = 1"
+    by (smt (verit) divide_less_eq_1 measure_spmf.subprob_measure_le_1 weight_spmf_lt_0)
+  ultimately have "ennreal (spmf ?lhs x) = ennreal (spmf ?rhs x)" 
+    using assms
     unfolding spmf_scale_spmf ennreal_pmf_bind nn_integral_measure_pmf UNIV_bool bernoulli_pmf.rep_eq
-    apply(auto simp add: nn_integral_count_space_finite max_def min_def field_simps real_inverse_le_1_iff[OF weight_spmf_nonneg weight_spmf_le_1] weight_spmf_lt_0 not_le ennreal_mult[symmetric])
-    apply (metis pmf_le_0_iff spmf_le_weight)
-    apply (metis pmf_le_0_iff spmf_le_weight)
-    apply (meson le_divide_eq_1_pos measure_spmf.subprob_measure_le_1 not_less order_trans weight_spmf_le_0)
-    by (meson divide_le_0_1_iff less_imp_le order_trans weight_spmf_le_0)
+    by(auto simp: nn_integral_count_space_finite max_def min_def field_simps 
+          real_inverse_le_1_iff[OF weight_spmf_nonneg weight_spmf_le_1]  ennreal_mult[symmetric])
   thus "spmf ?lhs x = spmf ?rhs x" by simp
 qed
 
 lemma nn_integral_spmf: "(\<integral>\<^sup>+ x. spmf p x \<partial>count_space A) = emeasure (measure_spmf p) A"
-apply(simp add: measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space nn_integral_pmf[symmetric])
-apply(rule nn_integral_bij_count_space[where g=Some])
-apply(auto simp add: bij_betw_def)
-done
+proof -
+  have "bij_betw Some A (the -` A \<inter> range Some)"
+    by(auto simp: bij_betw_def)
+  then show ?thesis
+    by (metis bij_betw_def emeasure_measure_spmf_conv_measure_pmf nn_integral_pmf')
+qed
 
 lemma measure_spmf_scale_spmf: "measure_spmf (scale_spmf r p) = scale_measure (min (inverse (weight_spmf p)) r) (measure_spmf p)"
-apply(rule measure_eqI)
- apply simp
-apply(simp add: nn_integral_spmf[symmetric] spmf_scale_spmf)
-apply(subst nn_integral_cmult[symmetric])
-apply(auto simp add: max_def min_def ennreal_mult[symmetric] not_le ennreal_lt_0)
-done
+  by(rule measure_eqI; simp add: spmf_scale_spmf ennreal_mult' flip: nn_integral_spmf nn_integral_cmult)
 
 lemma measure_spmf_scale_spmf':
-  "r \<le> 1 \<Longrightarrow> measure_spmf (scale_spmf r p) = scale_measure r (measure_spmf p)"
-unfolding measure_spmf_scale_spmf
-apply(cases "weight_spmf p > 0")
- apply(simp add: min.absorb2 field_simps weight_spmf_le_1 mult_le_one)
-apply(clarsimp simp add: weight_spmf_le_0 min_def scale_spmf_neg weight_spmf_eq_0 not_less)
-done
+  assumes "r \<le> 1"
+  shows "measure_spmf (scale_spmf r p) = scale_measure r (measure_spmf p)"
+proof(cases "weight_spmf p > 0")
+  case True
+  with assms show ?thesis    
+    by(simp add: measure_spmf_scale_spmf field_simps weight_spmf_le_1 mult_le_one)
+next
+  case False
+  then show ?thesis
+    by (simp add: order_less_le weight_spmf_eq_0)
+qed
 
 lemma scale_spmf_1 [simp]: "scale_spmf 1 p = p"
-apply(rule spmf_eqI)
-apply(simp add: spmf_scale_spmf max_def min_def order.strict_iff_order field_simps weight_spmf_nonneg)
-apply(metis antisym_conv divide_le_eq_1 less_imp_le pmf_nonneg spmf_le_weight weight_spmf_nonneg weight_spmf_le_1)
-done
+  by (simp add: spmf_eqI spmf_scale_spmf')
 
 lemma scale_spmf_0 [simp]: "scale_spmf 0 p = return_pmf None"
-by(rule spmf_eqI)(simp add: spmf_scale_spmf min_def max_def weight_spmf_le_0)
+  by (simp add: scale_spmf_neg)
 
 lemma bind_scale_spmf:
   assumes r: "r \<le> 1"
@@ -2264,9 +2254,10 @@
   (is "?lhs = ?rhs")
 proof(rule spmf_eqI)
   fix x
-  have "ennreal (spmf ?lhs x) = ennreal (spmf ?rhs x)" using r
-    by(simp add: ennreal_spmf_bind measure_spmf_scale_spmf' nn_integral_scale_measure spmf_scale_spmf')
-      (simp add: ennreal_mult ennreal_lt_0 nn_integral_cmult max_def min_def)
+  have "ennreal (spmf ?lhs x) = ennreal (spmf ?rhs x)" 
+    using r
+    by(simp add: ennreal_spmf_bind measure_spmf_scale_spmf' nn_integral_scale_measure spmf_scale_spmf'
+        ennreal_mult nn_integral_cmult)
   thus "spmf ?lhs x = spmf ?rhs x" by simp
 qed
 
@@ -2295,45 +2286,50 @@
 proof(rule spmf_eqI)
   fix i
   show "spmf ?lhs i = spmf ?rhs i" unfolding spmf_scale_spmf
-    by(subst (1 2) spmf_map)(auto simp add: measure_spmf_scale_spmf max_def min_def ennreal_lt_0)
+    by(subst (1 2) spmf_map)(auto simp: measure_spmf_scale_spmf max_def min_def ennreal_lt_0)
 qed
 
 lemma set_scale_spmf: "set_spmf (scale_spmf r p) = (if r > 0 then set_spmf p else {})"
-apply(auto simp add: in_set_spmf_iff_spmf spmf_scale_spmf)
-apply(simp add: max_def min_def not_le weight_spmf_lt_0 weight_spmf_eq_0 split: if_split_asm)
-done
+  apply(auto simp: in_set_spmf_iff_spmf spmf_scale_spmf)
+  apply(simp add: min_def weight_spmf_eq_0 split: if_split_asm)
+  done
 
 lemma set_scale_spmf' [simp]: "0 < r \<Longrightarrow> set_spmf (scale_spmf r p) = set_spmf p"
-by(simp add: set_scale_spmf)
+  by(simp add: set_scale_spmf)
 
 lemma rel_spmf_scaleI:
   assumes "r > 0 \<Longrightarrow> rel_spmf A p q"
   shows "rel_spmf A (scale_spmf r p) (scale_spmf r q)"
 proof(cases "r > 0")
   case True
-  from assms[OF this] show ?thesis
-    by(rule rel_spmfE)(auto simp add: map_scale_spmf[symmetric] spmf_rel_map True intro: rel_spmf_reflI)
+  from assms[OF True] show ?thesis
+    by(rule rel_spmfE)(auto simp: map_scale_spmf[symmetric] spmf_rel_map True intro: rel_spmf_reflI)
 qed(simp add: not_less scale_spmf_neg)
 
 lemma weight_scale_spmf: "weight_spmf (scale_spmf r p) = min 1 (max 0 r * weight_spmf p)"
 proof -
+  have "\<lbrakk>1 / weight_spmf p \<le> r; ennreal r * ennreal (weight_spmf p) < 1\<rbrakk> \<Longrightarrow> weight_spmf p = 0"
+    by (smt (verit) ennreal_less_one_iff ennreal_mult'' measure_le_0_iff mult_imp_less_div_pos)
+  moreover
+  have "\<lbrakk>r < 1 / weight_spmf p; 1 \<le> ennreal r * ennreal (weight_spmf p)\<rbrakk> \<Longrightarrow> weight_spmf p = 0"
+    by (smt (verit, ccfv_threshold) ennreal_ge_1 ennreal_mult'' mult_imp_div_pos_le weight_spmf_lt_0)
+  ultimately
   have "ennreal (weight_spmf (scale_spmf r p)) = min 1 (max 0 r * ennreal (weight_spmf p))"
     unfolding weight_spmf_eq_nn_integral_spmf
     apply(simp add: spmf_scale_spmf ennreal_mult zero_ereal_def[symmetric] nn_integral_cmult)
-    apply(auto simp add: weight_spmf_eq_nn_integral_spmf[symmetric] field_simps min_def max_def not_le weight_spmf_lt_0 ennreal_mult[symmetric])
-    subgoal by(subst (asm) ennreal_mult[symmetric], meson divide_less_0_1_iff le_less_trans not_le weight_spmf_lt_0, simp+, meson not_le pos_divide_le_eq weight_spmf_le_0)
-    subgoal by(cases "r \<ge> 0")(simp_all add: ennreal_mult[symmetric] weight_spmf_nonneg ennreal_lt_0, meson le_less_trans not_le pos_divide_le_eq zero_less_divide_1_iff)
+    apply(auto simp: weight_spmf_eq_nn_integral_spmf[symmetric] field_simps min_def max_def not_le weight_spmf_lt_0 ennreal_mult[symmetric])
     done
-  thus ?thesis by(auto simp add: min_def max_def ennreal_mult[symmetric] split: if_split_asm)
+  thus ?thesis
+    by(auto simp: min_def max_def ennreal_mult[symmetric] split: if_split_asm)
 qed
 
 lemma weight_scale_spmf' [simp]:
   "\<lbrakk> 0 \<le> r; r \<le> 1 \<rbrakk> \<Longrightarrow> weight_spmf (scale_spmf r p) = r * weight_spmf p"
-by(simp add: weight_scale_spmf max_def min_def)(metis antisym_conv mult_left_le order_trans weight_spmf_le_1)
+  by(simp add: weight_scale_spmf max_def min_def)(metis antisym_conv mult_left_le order_trans weight_spmf_le_1)
 
 lemma pmf_scale_spmf_None:
   "pmf (scale_spmf k p) None = 1 - min 1 (max 0 k * (1 - pmf p None))"
-unfolding pmf_None_eq_weight_spmf by(simp add: weight_scale_spmf)
+  unfolding pmf_None_eq_weight_spmf by(simp add: weight_scale_spmf)
 
 lemma scale_scale_spmf:
   "scale_spmf r (scale_spmf r' p) = scale_spmf (r * max 0 (min (inverse (weight_spmf p)) r')) p"
@@ -2347,25 +2343,29 @@
   show ?thesis
   proof(rule spmf_eqI)
     fix i
-    have "max 0 (min (1 / weight_spmf p) r') * max 0 (min (1 / min 1 (weight_spmf p * max 0 r')) r) =
+    have *: "max 0 (min (1 / weight_spmf p) r') * max 0 (min (1 / min 1 (weight_spmf p * max 0 r')) r) =
           max 0 (min (1 / weight_spmf p) (r * max 0 (min (1 / weight_spmf p) r')))"
       using True
-      by(simp add: field_simps max_def min.absorb_iff2[symmetric])(auto simp add: min_def field_simps zero_le_mult_iff)
-    then show "spmf ?lhs i = spmf ?rhs i"
-      apply (subst spmf_scale_spmf)+  (*FOR SOME REASON we now get linarith_split_limit exceeded if simp is used*)
-      by (metis (no_types, opaque_lifting) inverse_eq_divide mult.commute mult.left_commute weight_scale_spmf)
+      by (simp add: max_def) (auto simp: min_def field_simps zero_le_mult_iff)
+    show "spmf ?lhs i = spmf ?rhs i"
+      by (simp add: spmf_scale_spmf) (metis * inverse_eq_divide mult.commute weight_scale_spmf)
   qed
 qed
 
 lemma scale_scale_spmf' [simp]:
-  "\<lbrakk> 0 \<le> r; r \<le> 1; 0 \<le> r'; r' \<le> 1 \<rbrakk>
-  \<Longrightarrow> scale_spmf r (scale_spmf r' p) = scale_spmf (r * r') p"
-apply(cases "weight_spmf p > 0")
-apply(auto simp add: scale_scale_spmf min_def max_def field_simps not_le weight_spmf_lt_0 weight_spmf_eq_0 not_less weight_spmf_le_0)
-apply(subgoal_tac "1 = r'")
- apply (metis (no_types) div_by_1 eq_iff measure_spmf.subprob_measure_le_1 mult.commute mult_cancel_right1)
-apply(meson eq_iff le_divide_eq_1_pos measure_spmf.subprob_measure_le_1 mult_imp_div_pos_le order.trans)
-done
+  assumes "0 \<le> r" "r \<le> 1" "0 \<le> r'" "r' \<le> 1"
+  shows "scale_spmf r (scale_spmf r' p) = scale_spmf (r * r') p"
+proof(cases "weight_spmf p > 0")
+  case True
+  with assms have "r' = 1" if "1 \<le> r' * weight_spmf p"
+    by (smt (verit, best) measure_spmf.subprob_measure_le_1 mult_eq_1 mult_le_one that)
+  with assms True show ?thesis
+    by (smt (verit, best) eq_divide_imp measure_le_0_iff mult.assoc mult_nonneg_nonneg scale_scale_spmf weight_scale_spmf')
+next
+  case False
+  with assms show ?thesis
+    by (simp add: weight_spmf_eq_0 zero_less_measure_iff)
+qed
 
 lemma scale_spmf_eq_same: "scale_spmf r p = p \<longleftrightarrow> weight_spmf p = 0 \<or> r = 1 \<or> r \<ge> 1 \<and> weight_spmf p = 1"
   (is "?lhs \<longleftrightarrow> ?rhs")
@@ -2373,53 +2373,59 @@
   assume ?lhs
   hence "weight_spmf (scale_spmf r p) = weight_spmf p" by simp
   hence *: "min 1 (max 0 r * weight_spmf p) = weight_spmf p" by(simp add: weight_scale_spmf)
-  hence **: "weight_spmf p = 0 \<or> r \<ge> 1" by(auto simp add: min_def max_def split: if_split_asm)
+  hence **: "weight_spmf p = 0 \<or> r \<ge> 1" by(auto simp: min_def max_def split: if_split_asm)
   show ?rhs
   proof(cases "weight_spmf p = 0")
     case False
-    with ** have "r \<ge> 1" by simp
-    with * False have "r = 1 \<or> weight_spmf p = 1" by(simp add: max_def min_def not_le split: if_split_asm)
-    with \<open>r \<ge> 1\<close> show ?thesis by simp
+    with ** have "r \<ge> 1" 
+      by simp
+    with * False have "r = 1 \<or> weight_spmf p = 1" 
+      by(simp add: max_def min_def not_le split: if_split_asm)
+    with \<open>r \<ge> 1\<close> show ?thesis 
+      by simp
   qed simp
-qed(auto intro!: spmf_eqI simp add: spmf_scale_spmf, metis pmf_le_0_iff spmf_le_weight)
+next
+  show "weight_spmf p = 0 \<or> r = 1 \<or> 1 \<le> r \<and> weight_spmf p = 1 \<Longrightarrow> scale_spmf r p = p"
+    by (smt (verit) div_by_1 inverse_eq_divide inverse_positive_iff_positive scale_scale_spmf scale_spmf_1)
+qed
 
 lemma map_const_spmf_of_set:
   "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> map_spmf (\<lambda>_. c) (spmf_of_set A) = return_spmf c"
-by(simp add: map_spmf_conv_bind_spmf bind_spmf_const)
+  by(simp add: map_spmf_conv_bind_spmf bind_spmf_const)
 
 subsection \<open>Conditional spmfs\<close>
 
 lemma set_pmf_Int_Some: "set_pmf p \<inter> Some ` A = {} \<longleftrightarrow> set_spmf p \<inter> A = {}"
-by(auto simp add: in_set_spmf)
+  by(auto simp: in_set_spmf)
 
 lemma measure_spmf_zero_iff: "measure (measure_spmf p) A = 0 \<longleftrightarrow> set_spmf p \<inter> A = {}"
-unfolding measure_measure_spmf_conv_measure_pmf by(simp add: measure_pmf_zero_iff set_pmf_Int_Some)
+  unfolding measure_measure_spmf_conv_measure_pmf by(simp add: measure_pmf_zero_iff set_pmf_Int_Some)
 
 definition cond_spmf :: "'a spmf \<Rightarrow> 'a set \<Rightarrow> 'a spmf"
-where "cond_spmf p A = (if set_spmf p \<inter> A = {} then return_pmf None else cond_pmf p (Some ` A))"
+  where "cond_spmf p A = (if set_spmf p \<inter> A = {} then return_pmf None else cond_pmf p (Some ` A))"
 
 lemma set_cond_spmf [simp]: "set_spmf (cond_spmf p A) = set_spmf p \<inter> A"
-by(auto 4 4 simp add: cond_spmf_def in_set_spmf iff: set_cond_pmf[THEN set_eq_iff[THEN iffD1], THEN spec, rotated])
+  by(auto 4 4 simp add: cond_spmf_def in_set_spmf iff: set_cond_pmf[THEN set_eq_iff[THEN iffD1], THEN spec, rotated])
 
 lemma cond_map_spmf [simp]: "cond_spmf (map_spmf f p) A = map_spmf f (cond_spmf p (f -` A))"
 proof -
   have "map_option f -` Some ` A = Some ` f -` A" by auto
   moreover have "set_pmf p \<inter> map_option f -` Some ` A \<noteq> {}" if "Some x \<in> set_pmf p" "f x \<in> A" for x
     using that by auto
-  ultimately show ?thesis by(auto simp add: cond_spmf_def in_set_spmf cond_map_pmf)
+  ultimately show ?thesis by(auto simp: cond_spmf_def in_set_spmf cond_map_pmf)
 qed
 
 lemma spmf_cond_spmf [simp]:
   "spmf (cond_spmf p A) x = (if x \<in> A then spmf p x / measure (measure_spmf p) A else 0)"
-by(auto simp add: cond_spmf_def pmf_cond set_pmf_Int_Some[symmetric] measure_measure_spmf_conv_measure_pmf measure_pmf_zero_iff)
+  by(auto simp: cond_spmf_def pmf_cond set_pmf_Int_Some[symmetric] measure_measure_spmf_conv_measure_pmf measure_pmf_zero_iff)
 
 lemma bind_eq_return_pmf_None:
   "bind_spmf p f = return_pmf None \<longleftrightarrow> (\<forall>x\<in>set_spmf p. f x = return_pmf None)"
-by(auto simp add: bind_spmf_def bind_eq_return_pmf in_set_spmf split: option.splits)
+  by(auto simp: bind_spmf_def bind_eq_return_pmf in_set_spmf split: option.splits)
 
 lemma return_pmf_None_eq_bind:
   "return_pmf None = bind_spmf p f \<longleftrightarrow> (\<forall>x\<in>set_spmf p. f x = return_pmf None)"
-using bind_eq_return_pmf_None[of p f] by auto
+  using bind_eq_return_pmf_None[of p f] by auto
 
 (* Conditional probabilities do not seem to interact nicely with bind. *)
 
@@ -2429,27 +2435,29 @@
 where "pair_spmf p q = bind_pmf (pair_pmf p q) (\<lambda>xy. case xy of (Some x, Some y) \<Rightarrow> return_spmf (x, y) | _ \<Rightarrow> return_pmf None)"
 
 lemma map_fst_pair_spmf [simp]: "map_spmf fst (pair_spmf p q) = scale_spmf (weight_spmf q) p"
-unfolding bind_spmf_const[symmetric]
-apply(simp add: pair_spmf_def map_bind_pmf pair_pmf_def bind_assoc_pmf option.case_distrib)
-apply(subst bind_commute_pmf)
-apply(auto intro!: bind_pmf_cong[OF refl] simp add: bind_return_pmf bind_spmf_def bind_return_pmf' case_option_collapse option.case_distrib[where h="map_spmf _"] option.case_distrib[symmetric] case_option_id split: option.split cong del: option.case_cong_weak)
-done
+  unfolding bind_spmf_const[symmetric]
+  apply(simp add: pair_spmf_def map_bind_pmf pair_pmf_def bind_assoc_pmf option.case_distrib)
+  apply(subst bind_commute_pmf)
+  apply(force intro!: bind_pmf_cong[OF refl] simp add: bind_return_pmf bind_spmf_def bind_return_pmf' case_option_collapse 
+      option.case_distrib[where h="map_spmf _"] option.case_distrib[symmetric] case_option_id split: option.split cong: option.case_cong)
+  done
 
 lemma map_snd_pair_spmf [simp]: "map_spmf snd (pair_spmf p q) = scale_spmf (weight_spmf p) q"
-unfolding bind_spmf_const[symmetric]
+  unfolding bind_spmf_const[symmetric]
   apply(simp add: pair_spmf_def map_bind_pmf pair_pmf_def bind_assoc_pmf option.case_distrib
-    cong del: option.case_cong_weak)
-apply(auto intro!: bind_pmf_cong[OF refl] simp add: bind_return_pmf bind_spmf_def bind_return_pmf' case_option_collapse option.case_distrib[where h="map_spmf _"] option.case_distrib[symmetric] case_option_id split: option.split cong del: option.case_cong_weak)
-done
+      cong del: option.case_cong_weak)
+  apply(auto intro!: bind_pmf_cong[OF refl] simp add: bind_return_pmf bind_spmf_def bind_return_pmf' case_option_collapse
+      option.case_distrib[where h="map_spmf _"] option.case_distrib[symmetric] case_option_id split: option.split cong del: option.case_cong_weak)
+  done
 
 lemma set_pair_spmf [simp]: "set_spmf (pair_spmf p q) = set_spmf p \<times> set_spmf q"
-by(auto 4 3 simp add: pair_spmf_def set_spmf_bind_pmf bind_UNION in_set_spmf intro: rev_bexI split: option.splits)
+  by(force simp add: pair_spmf_def set_spmf_bind_pmf bind_UNION in_set_spmf split: option.splits)
 
 lemma spmf_pair [simp]: "spmf (pair_spmf p q) (x, y) = spmf p x * spmf q y" (is "?lhs = ?rhs")
 proof -
   have "ennreal ?lhs = \<integral>\<^sup>+ a. \<integral>\<^sup>+ b. indicator {(x, y)} (a, b) \<partial>measure_spmf q \<partial>measure_spmf p"
     unfolding measure_spmf_def pair_spmf_def ennreal_pmf_bind nn_integral_pair_pmf'
-    by(auto simp add: zero_ereal_def[symmetric] nn_integral_distr nn_integral_restrict_space nn_integral_multc[symmetric] intro!: nn_integral_cong split: option.split split_indicator)
+    by(auto simp: zero_ereal_def[symmetric] nn_integral_distr nn_integral_restrict_space nn_integral_multc[symmetric] intro!: nn_integral_cong split: option.split split_indicator)
   also have "\<dots> = \<integral>\<^sup>+ a. (\<integral>\<^sup>+ b. indicator {y} b \<partial>measure_spmf q) * indicator {x} a \<partial>measure_spmf p"
     by(subst nn_integral_multc[symmetric])(auto intro!: nn_integral_cong split: split_indicator)
   also have "\<dots> = ennreal ?rhs" by(simp add: emeasure_spmf_single max_def ennreal_mult mult.commute)
@@ -2457,42 +2465,46 @@
 qed
 
 lemma pair_map_spmf2: "pair_spmf p (map_spmf f q) = map_spmf (apsnd f) (pair_spmf p q)"
-by(auto simp add: pair_spmf_def pair_map_pmf2 bind_map_pmf map_bind_pmf intro: bind_pmf_cong split: option.split)
+  unfolding pair_spmf_def pair_map_pmf2 bind_map_pmf map_bind_pmf
+  by (intro bind_pmf_cong refl) (auto split: option.split)
 
 lemma pair_map_spmf1: "pair_spmf (map_spmf f p) q = map_spmf (apfst f) (pair_spmf p q)"
-by(auto simp add: pair_spmf_def pair_map_pmf1 bind_map_pmf map_bind_pmf intro: bind_pmf_cong split: option.split)
+  unfolding pair_spmf_def pair_map_pmf1 bind_map_pmf map_bind_pmf
+  by (intro bind_pmf_cong refl) (auto split: option.split)
 
 lemma pair_map_spmf: "pair_spmf (map_spmf f p) (map_spmf g q) = map_spmf (map_prod f g) (pair_spmf p q)"
-unfolding pair_map_spmf2 pair_map_spmf1 spmf.map_comp by(simp add: apfst_def apsnd_def o_def prod.map_comp)
+  unfolding pair_map_spmf2 pair_map_spmf1 spmf.map_comp 
+  by(simp add: apfst_def apsnd_def o_def prod.map_comp)
 
 lemma pair_spmf_alt_def: "pair_spmf p q = bind_spmf p (\<lambda>x. bind_spmf q (\<lambda>y. return_spmf (x, y)))"
-by(auto simp add: pair_spmf_def pair_pmf_def bind_spmf_def bind_assoc_pmf bind_return_pmf split: option.split intro: bind_pmf_cong)
+  unfolding pair_spmf_def pair_pmf_def bind_spmf_def bind_assoc_pmf bind_return_pmf
+  by (intro bind_pmf_cong refl) (auto split: option.split)
 
 lemma weight_pair_spmf [simp]: "weight_spmf (pair_spmf p q) = weight_spmf p * weight_spmf q"
-unfolding pair_spmf_alt_def by(simp add: weight_bind_spmf o_def)
+  unfolding pair_spmf_alt_def by(simp add: weight_bind_spmf o_def)
 
 lemma pair_scale_spmf1: (* FIXME: generalise to arbitrary r *)
   "r \<le> 1 \<Longrightarrow> pair_spmf (scale_spmf r p) q = scale_spmf r (pair_spmf p q)"
-by(simp add: pair_spmf_alt_def scale_bind_spmf bind_scale_spmf)
+  by(simp add: pair_spmf_alt_def scale_bind_spmf bind_scale_spmf)
 
 lemma pair_scale_spmf2: (* FIXME: generalise to arbitrary r *)
   "r \<le> 1 \<Longrightarrow> pair_spmf p (scale_spmf r q) = scale_spmf r (pair_spmf p q)"
-by(simp add: pair_spmf_alt_def scale_bind_spmf bind_scale_spmf)
+  by(simp add: pair_spmf_alt_def scale_bind_spmf bind_scale_spmf)
 
 lemma pair_spmf_return_None1 [simp]: "pair_spmf (return_pmf None) p = return_pmf None"
-by(rule spmf_eqI)(clarsimp)
+  by(rule spmf_eqI)(clarsimp)
 
 lemma pair_spmf_return_None2 [simp]: "pair_spmf p (return_pmf None) = return_pmf None"
-by(rule spmf_eqI)(clarsimp)
+  by(rule spmf_eqI)(clarsimp)
 
 lemma pair_spmf_return_spmf1: "pair_spmf (return_spmf x) q = map_spmf (Pair x) q"
-by(rule spmf_eqI)(auto split: split_indicator simp add: spmf_map_inj' inj_on_def intro: spmf_map_outside)
+  by(rule spmf_eqI)(auto split: split_indicator simp add: spmf_map_inj' inj_on_def intro: spmf_map_outside)
 
 lemma pair_spmf_return_spmf2: "pair_spmf p (return_spmf y) = map_spmf (\<lambda>x. (x, y)) p"
-by(rule spmf_eqI)(auto split: split_indicator simp add: inj_on_def intro!: spmf_map_outside spmf_map_inj'[symmetric])
+  by(rule spmf_eqI)(auto split: split_indicator simp add: inj_on_def intro!: spmf_map_outside spmf_map_inj'[symmetric])
 
 lemma pair_spmf_return_spmf [simp]: "pair_spmf (return_spmf x) (return_spmf y) = return_spmf (x, y)"
-by(simp add: pair_spmf_return_spmf1)
+  by(simp add: pair_spmf_return_spmf1)
 
 lemma rel_pair_spmf_prod:
   "rel_spmf (rel_prod A B) (pair_spmf p q) (pair_spmf p' q') \<longleftrightarrow>
@@ -2532,7 +2544,7 @@
       by(simp add: pair_map_spmf[symmetric] p q map_scale_spmf spmf.map_comp)
     also have "\<dots> = pair_spmf p q" using full[of p q]
       by(simp add: pair_scale_spmf1 pair_scale_spmf2 weight_spmf_le_1 weight_spmf_nonneg)
-        (auto simp add: scale_scale_spmf max_def min_def field_simps weight_spmf_nonneg weight_spmf_eq_0)
+        (auto simp: scale_scale_spmf max_def min_def field_simps weight_spmf_nonneg weight_spmf_eq_0)
     finally show "map_spmf fst ?pq = \<dots>" .
 
     have [simp]: "snd \<circ> ?f = map_prod snd snd" by(simp add: fun_eq_iff)
@@ -2543,9 +2555,9 @@
       by(simp add: pair_map_spmf[symmetric] p' q' map_scale_spmf spmf.map_comp)
     also have "\<dots> = pair_spmf p' q'" using full[of p' q'] eq
       by(simp add: pair_scale_spmf1 pair_scale_spmf2 weight_spmf_le_1 weight_spmf_nonneg)
-        (auto simp add: scale_scale_spmf max_def min_def field_simps weight_spmf_nonneg weight_spmf_eq_0)
+        (auto simp: scale_scale_spmf max_def min_def field_simps weight_spmf_nonneg weight_spmf_eq_0)
     finally show "map_spmf snd ?pq = \<dots>" .
-  qed(auto simp add: set_scale_spmf split: if_split_asm dest: * ** )
+  qed(auto simp: set_scale_spmf split: if_split_asm dest: * ** )
 next
   assume ?lhs
   then obtain pq where pq: "map_spmf fst pq = pair_spmf p q"
@@ -2582,30 +2594,30 @@
 
 lemma pair_pair_spmf:
   "pair_spmf (pair_spmf p q) r = map_spmf (\<lambda>(x, (y, z)). ((x, y), z)) (pair_spmf p (pair_spmf q r))"
-by(simp add: pair_spmf_alt_def map_spmf_conv_bind_spmf)
+  by(simp add: pair_spmf_alt_def map_spmf_conv_bind_spmf)
 
 lemma pair_commute_spmf:
   "pair_spmf p q = map_spmf (\<lambda>(y, x). (x, y)) (pair_spmf q p)"
-unfolding pair_spmf_alt_def by(subst bind_commute_spmf)(simp add: map_spmf_conv_bind_spmf)
+  unfolding pair_spmf_alt_def by(subst bind_commute_spmf)(simp add: map_spmf_conv_bind_spmf)
 
 subsection \<open>Assertions\<close>
 
 definition assert_spmf :: "bool \<Rightarrow> unit spmf"
-where "assert_spmf b = (if b then return_spmf () else return_pmf None)"
+  where "assert_spmf b = (if b then return_spmf () else return_pmf None)"
 
 lemma assert_spmf_simps [simp]:
   "assert_spmf True = return_spmf ()"
   "assert_spmf False = return_pmf None"
-by(simp_all add: assert_spmf_def)
+  by(simp_all add: assert_spmf_def)
 
 lemma in_set_assert_spmf [simp]: "x \<in> set_spmf (assert_spmf p) \<longleftrightarrow> p"
-by(cases p) simp_all
+  by(cases p) simp_all
 
 lemma set_spmf_assert_spmf_eq_empty [simp]: "set_spmf (assert_spmf b) = {} \<longleftrightarrow> \<not> b"
-by(cases b) simp_all
+  by auto
 
 lemma lossless_assert_spmf [iff]: "lossless_spmf (assert_spmf b) \<longleftrightarrow> b"
-by(cases b) simp_all
+  by(cases b) simp_all
 
 subsection \<open>Try\<close>
 
@@ -2617,59 +2629,59 @@
   shows "TRY p ELSE q = p"
 proof -
   have "TRY p ELSE q = bind_pmf p return_pmf" unfolding try_spmf_def using assms
-    by(auto simp add: lossless_iff_set_pmf_None split: option.split intro: bind_pmf_cong)
+    by(auto simp: lossless_iff_set_pmf_None split: option.split intro: bind_pmf_cong)
   thus ?thesis by(simp add: bind_return_pmf')
 qed
 
 lemma try_spmf_return_spmf1: "TRY return_spmf x ELSE q = return_spmf x"
-by(simp add: try_spmf_def bind_return_pmf)
+  by simp
 
 lemma try_spmf_return_None [simp]: "TRY return_pmf None ELSE q = q"
-by(simp add: try_spmf_def bind_return_pmf)
+  by(simp add: try_spmf_def bind_return_pmf)
 
 lemma try_spmf_return_pmf_None2 [simp]: "TRY p ELSE return_pmf None = p"
-by(simp add: try_spmf_def option.case_distrib[symmetric] bind_return_pmf' case_option_id)
+  by(simp add: try_spmf_def option.case_distrib[symmetric] bind_return_pmf' case_option_id)
 
 lemma map_try_spmf: "map_spmf f (try_spmf p q) = try_spmf (map_spmf f p) (map_spmf f q)"
-by(simp add: try_spmf_def map_bind_pmf bind_map_pmf option.case_distrib[where h="map_spmf f"] o_def cong del: option.case_cong_weak)
+  by(simp add: try_spmf_def map_bind_pmf bind_map_pmf option.case_distrib[where h="map_spmf f"] o_def cong del: option.case_cong_weak)
 
 lemma try_spmf_bind_pmf: "TRY (bind_pmf p f) ELSE q = bind_pmf p (\<lambda>x. TRY (f x) ELSE q)"
-by(simp add: try_spmf_def bind_assoc_pmf)
+  by(simp add: try_spmf_def bind_assoc_pmf)
 
 lemma try_spmf_bind_spmf_lossless:
   "lossless_spmf p \<Longrightarrow> TRY (bind_spmf p f) ELSE q = bind_spmf p (\<lambda>x. TRY (f x) ELSE q)"
-by(auto simp add: try_spmf_def bind_spmf_def bind_assoc_pmf bind_return_pmf lossless_iff_set_pmf_None intro!: bind_pmf_cong split: option.split)
+  by (metis (mono_tags, lifting) bind_spmf_of_pmf lossless_spmf_conv_spmf_of_pmf try_spmf_bind_pmf)
 
 lemma try_spmf_bind_out:
   "lossless_spmf p \<Longrightarrow> bind_spmf p (\<lambda>x. TRY (f x) ELSE q) = TRY (bind_spmf p f) ELSE q"
-by(simp add: try_spmf_bind_spmf_lossless)
+  by(simp add: try_spmf_bind_spmf_lossless)
 
 lemma lossless_try_spmf [simp]:
   "lossless_spmf (TRY p ELSE q) \<longleftrightarrow> lossless_spmf p \<or> lossless_spmf q"
-by(auto simp add: try_spmf_def in_set_spmf lossless_iff_set_pmf_None split: option.splits)
+  by(auto simp: try_spmf_def in_set_spmf lossless_iff_set_pmf_None split: option.splits)
 
 context includes lifting_syntax
 begin
 
 lemma try_spmf_parametric [transfer_rule]:
   "(rel_spmf A ===> rel_spmf A ===> rel_spmf A) try_spmf try_spmf"
-unfolding try_spmf_def[abs_def] by transfer_prover
+  unfolding try_spmf_def[abs_def] by transfer_prover
 
 end
 
 lemma try_spmf_cong:
   "\<lbrakk> p = p'; \<not> lossless_spmf p' \<Longrightarrow> q = q' \<rbrakk> \<Longrightarrow> TRY p ELSE q = TRY p' ELSE q'"
-unfolding try_spmf_def
-by(rule bind_pmf_cong)(auto split: option.split simp add: lossless_iff_set_pmf_None)
+  unfolding try_spmf_def
+  by(rule bind_pmf_cong)(auto split: option.split simp add: lossless_iff_set_pmf_None)
 
 lemma rel_spmf_try_spmf:
   "\<lbrakk> rel_spmf R p p'; \<not> lossless_spmf p' \<Longrightarrow> rel_spmf R q q' \<rbrakk>
   \<Longrightarrow> rel_spmf R (TRY p ELSE q) (TRY p' ELSE q')"
-unfolding try_spmf_def
-apply(rule rel_pmf_bindI[where R="\<lambda>x y. rel_option R x y \<and> x \<in> set_pmf p \<and> y \<in> set_pmf p'"])
- apply(erule pmf.rel_mono_strong; simp)
-apply(auto split: option.split simp add: lossless_iff_set_pmf_None)
-done
+  unfolding try_spmf_def
+  apply(rule rel_pmf_bindI[where R="\<lambda>x y. rel_option R x y \<and> x \<in> set_pmf p \<and> y \<in> set_pmf p'"])
+  apply (simp add: pmf.rel_mono_strong)
+  apply(auto split: option.split simp add: lossless_iff_set_pmf_None)
+  done
 
 lemma spmf_try_spmf:
   "spmf (TRY p ELSE q) x = spmf p x + pmf p None * spmf q x"
@@ -2679,11 +2691,11 @@
   also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (spmf q x) * indicator {None} y \<partial>measure_pmf p) + \<integral>\<^sup>+ y. indicator {Some x} y \<partial>measure_pmf p"
     by(simp add: nn_integral_add)
   also have "\<dots> = ennreal (spmf q x) * pmf p None + spmf p x" by(simp add: emeasure_pmf_single)
-  finally show ?thesis by(simp add: ennreal_mult[symmetric] ennreal_plus[symmetric] del: ennreal_plus)
+  finally show ?thesis by(simp flip: ennreal_plus ennreal_mult)
 qed
 
 lemma try_scale_spmf_same [simp]: "lossless_spmf p \<Longrightarrow> TRY scale_spmf k p ELSE p = p"
-by(rule spmf_eqI)(auto simp add: spmf_try_spmf spmf_scale_spmf pmf_scale_spmf_None lossless_iff_pmf_None weight_spmf_conv_pmf_None min_def max_def field_simps)
+  by(rule spmf_eqI)(auto simp: spmf_try_spmf spmf_scale_spmf pmf_scale_spmf_None lossless_iff_pmf_None weight_spmf_conv_pmf_None min_def max_def field_simps)
 
 lemma pmf_try_spmf_None [simp]: "pmf (TRY p ELSE q) None = pmf p None * pmf q None" (is "?lhs = ?rhs")
 proof -
@@ -2714,7 +2726,7 @@
   and fundamental_lemma: "\<bar>measure (measure_spmf p) {x. A x} - measure (measure_spmf q) {y. B y}\<bar> \<le>
     measure (measure_spmf p) {x. bad1 x}" (is ?fundamental)
 proof -
-  have good: "rel_fun ?A (=) (\<lambda>x. A x \<and> \<not> bad1 x) (\<lambda>y. B y \<and> \<not> bad2 y)" by(auto simp add: rel_fun_def)
+  have good: "rel_fun ?A (=) (\<lambda>x. A x \<and> \<not> bad1 x) (\<lambda>y. B y \<and> \<not> bad2 y)" by(auto simp: rel_fun_def)
   from assms have 1: "measure (measure_spmf p) {x. A x \<and> \<not> bad1 x} = measure (measure_spmf q) {y. B y \<and> \<not> bad2 y}"
     by(rule measure_spmf_parametric[THEN rel_funD, THEN rel_funD])(rule Collect_parametric[THEN rel_funD, OF good])
 
@@ -2731,7 +2743,7 @@
     by(subst (1 2) measure_Union)(auto)
   also have "\<dots> = \<bar>?\<mu>p {x. A x \<and> bad1 x} - ?\<mu>q {y. B y \<and> bad2 y}\<bar>" using 1 by simp
   also have "\<dots> \<le> max (?\<mu>p {x. A x \<and> bad1 x}) (?\<mu>q {y. B y \<and> bad2 y})"
-    by(rule abs_leI)(auto simp add: max_def not_le, simp_all only: add_increasing measure_nonneg mult_2)
+    by(rule abs_leI)(auto simp: max_def not_le, simp_all only: add_increasing measure_nonneg mult_2)
   also have "\<dots> \<le> max (?\<mu>p {x. bad1 x}) (?\<mu>q {y. bad2 y})"
     by(rule max.mono; rule measure_spmf.finite_measure_mono; auto)
   also note 2[symmetric]
--- a/src/HOL/Tools/ATP/system_on_tptp.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Tools/ATP/system_on_tptp.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -70,7 +70,7 @@
   object Run_System extends Scala.Fun_Strings("SystemOnTPTP.run_system", thread = true) {
     val here = Scala_Project.here
     def apply(args: List[String]): List[String] = {
-      val List(url, system, problem_path, extra, Value.Int(timeout)) = args
+      val List(url, system, problem_path, extra, Value.Int(timeout)) = args : @unchecked
       val problem = File.read(Path.explode(problem_path))
 
       val res = run_system(Url(url), system, problem, extra = extra, timeout = Time.ms(timeout))
--- a/src/HOL/Topological_Spaces.thy	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/HOL/Topological_Spaces.thy	Sun Sep 03 19:28:59 2023 +0200
@@ -1220,6 +1220,11 @@
   "f \<longlonglongrightarrow> f0 \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> f0 \<in> S \<longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> S))"
   unfolding tendsto_def eventually_sequentially by auto
 
+lemma closed_sequentially:
+  assumes "closed S" and "\<And>n. f n \<in> S" and "f \<longlonglongrightarrow> l"
+  shows "l \<in> S"
+  by (metis Lim_in_closed_set assms eventually_sequentially trivial_limit_sequentially)
+
 
 subsection \<open>Monotone sequences and subsequences\<close>
 
--- a/src/Pure/Admin/afp.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Admin/afp.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -21,8 +21,6 @@
 
   val chapter: String = "AFP"
 
-  val force_partition1: List[String] = List("Category3", "HOL-ODE")
-
   val BASE: Path = Path.explode("$AFP_BASE")
 
   def main_dir(base_dir: Path = BASE): Path = base_dir + Path.explode("thys")
@@ -207,19 +205,4 @@
   }
  }"""
     }).mkString("[", ", ", "\n]\n")
-
-
-  /* partition sessions */
-
-  def partition(n: Int): List[String] =
-    n match {
-      case 0 => Nil
-      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
-        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_history.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Admin/build_history.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -112,12 +112,11 @@
     root: Path,
     progress: Progress = new Progress,
     afp: Boolean = false,
-    afp_partition: Int = 0,
     isabelle_identifier: String = default_isabelle_identifier,
     ml_statistics_step: Int = 1,
     component_repository: String = Components.static_component_repository,
     components_base: String = Components.dynamic_components_base,
-    clean_platforms: Option[List[Platform.Family.Value]] = None,
+    clean_platforms: Option[List[Platform.Family]] = None,
     clean_archives: Boolean = false,
     fresh: Boolean = false,
     hostname: String = "",
@@ -167,22 +166,9 @@
     }
 
     val isabelle_directory = directory(root)
-    val afp_directory = if (afp) Some(directory(root + Path.explode("AFP"))) else None
-
-    val (afp_build_args, afp_sessions) =
-      if (afp_directory.isEmpty) (Nil, Nil)
-      else {
-        val (opt, sessions) = {
-          if (afp_partition == 0) ("-d", Nil)
-          else {
-            try {
-              val afp_info = AFP.init(options, base_dir = afp_directory.get.root)
-              ("-d", afp_info.partition(afp_partition))
-            } catch { case ERROR(_) => ("-D", Nil) }
-          }
-        }
-        (List(opt, "~~/AFP/thys"), sessions)
-      }
+    val (afp_directory, afp_build_args) =
+      if (afp) (Some(directory(root + Path.explode("AFP"))), List("-d", "~~/AFP/thys"))
+      else (None, Nil)
 
 
     /* main */
@@ -266,7 +252,7 @@
       val build_result =
         Other_Isabelle(root, isabelle_identifier = isabelle_identifier,
           progress = build_out_progress)
-        .bash("bin/isabelle build " + Bash.strings(build_args1 ::: afp_sessions),
+        .bash("bin/isabelle build " + Bash.strings(build_args1),
           redirect = true, echo = true, strict = false)
 
       val build_end = Date.now()
@@ -423,8 +409,7 @@
       var max_heap: Option[Int] = None
       var multicore_list = List(default_multicore)
       var isabelle_identifier = default_isabelle_identifier
-      var clean_platforms: Option[List[Platform.Family.Value]] = None
-      var afp_partition = 0
+      var clean_platforms: Option[List[Platform.Family]] = None
       var clean_archives = false
       var component_repository = Components.static_component_repository
       var arch_apple = false
@@ -453,7 +438,6 @@
     -N NAME      alternative ISABELLE_IDENTIFIER (default: """ + default_isabelle_identifier + """)
     -O PLATFORMS clean resolved components, retaining only the given list
                  platform families (separated by commas; default: do nothing)
-    -P NUMBER    AFP partition number (0, 1, 2, default: 0=unrestricted)
     -Q           clean archives of downloaded components
     -R URL       remote repository for Isabelle components (default: """ +
       Components.static_component_repository + """)
@@ -485,7 +469,6 @@
         "M:" -> (arg => multicore_list = space_explode(',', arg).map(Multicore.parse)),
         "N:" -> (arg => isabelle_identifier = arg),
         "O:" -> (arg => clean_platforms = Some(space_explode(',',arg).map(Platform.Family.parse))),
-        "P:" -> (arg => afp_partition = Value.Int.parse(arg)),
         "Q" -> (_ => clean_archives = true),
         "R:" -> (arg => component_repository = arg),
         "U:" -> (arg => max_heap = Some(Value.Int.parse(arg))),
@@ -516,8 +499,7 @@
       val progress = new Console_Progress(stderr = true)
 
       val results =
-        local_build(Options.init(), root, progress = progress,
-          afp = afp, afp_partition = afp_partition,
+        local_build(Options.init(), root, progress = progress, afp = afp,
           isabelle_identifier = isabelle_identifier, ml_statistics_step = ml_statistics_step,
           component_repository = component_repository, components_base = components_base,
           clean_platforms = clean_platforms, clean_archives = clean_archives,
--- a/src/Pure/Admin/build_log.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Admin/build_log.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -201,7 +201,7 @@
     /* inlined text */
 
     def filter(Marker: Protocol_Message.Marker): List[String] =
-      for (Marker(text) <- lines) yield text
+      for (case Marker(text) <- lines) yield text
 
     def find(Marker: Protocol_Message.Marker): Option[String] =
       lines.collectFirst({ case Marker(text) => text })
@@ -395,9 +395,7 @@
 
   val SESSION_NAME = "session_name"
 
-  object Session_Status extends Enumeration {
-    val existing, finished, failed, cancelled = Value
-  }
+  enum Session_Status { case existing, finished, failed, cancelled }
 
   sealed case class Session_Entry(
     chapter: String = "",
@@ -407,7 +405,7 @@
     ml_timing: Timing = Timing.zero,
     sources: Option[String] = None,
     heap_size: Option[Space] = None,
-    status: Option[Session_Status.Value] = None,
+    status: Option[Session_Status] = None,
     errors: List[String] = Nil,
     theory_timings: Map[String, Timing] = Map.empty,
     ml_statistics: List[Properties.T] = Nil
@@ -445,7 +443,7 @@
     val Session_Timing =
       new Regex("""^Timing (\S+) \((\d+) threads, (\d+\.\d+)s elapsed time, (\d+\.\d+)s cpu time, (\d+\.\d+)s GC time.*$""")
     val Session_Started1 = new Regex("""^(?:Running|Building) (\S+) \.\.\.$""")
-    val Session_Started2 = new Regex("""^(?:Running|Building) (\S+) on \S+ \.\.\.$""")
+    val Session_Started2 = new Regex("""^(?:Running|Building) (\S+) \(?on \S+\)? \.\.\.$""")
     val Sources = new Regex("""^Sources (\S+) (\S{""" + SHA1.digest_length + """})$""")
     val Heap = new Regex("""^Heap (\S+) \((\d+) bytes\)$""")
 
@@ -1040,13 +1038,13 @@
           Par_List.map[JFile, Exn.Result[Log_File]](
             file => Exn.capture { Log_File(file) }, file_group)
         db.transaction {
-          for (Exn.Res(log_file) <- log_files) {
+          for (case Exn.Res(log_file) <- log_files) {
             progress.echo("Log " + quote(log_file.name), verbose = true)
             try { status.foreach(_.update(log_file)) }
             catch { case exn: Throwable => add_error(log_file.name, exn) }
           }
         }
-        for ((file, Exn.Exn(exn)) <- file_group.zip(log_files)) {
+        for (case (file, Exn.Exn(exn)) <- file_group.zip(log_files)) {
           add_error(Log_File.plain_name(file), exn)
         }
       }
@@ -1071,8 +1069,8 @@
                else
                 res.get_string(c)))
           val n = Prop.all_props.length
-          val props = for ((x, Some(y)) <- results.take(n)) yield (x, y)
-          val settings = for ((x, Some(y)) <- results.drop(n)) yield (x, y)
+          val props = for (case (x, Some(y)) <- results.take(n)) yield (x, y)
+          val settings = for (case (x, Some(y)) <- results.drop(n)) yield (x, y)
           Meta_Info(props, settings)
         }
       )
@@ -1129,7 +1127,7 @@
                     Data.ml_timing_gc),
                 sources = res.get_string(Data.sources),
                 heap_size = res.get_long(Data.heap_size).map(Space.bytes),
-                status = res.get_string(Data.status).map(Session_Status.withName),
+                status = res.get_string(Data.status).map(Session_Status.valueOf),
                 errors = uncompress_errors(res.bytes(Data.errors), cache = cache),
                 ml_statistics =
                   if (ml_statistics) {
--- a/src/Pure/Admin/build_release.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Admin/build_release.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -71,7 +71,7 @@
 """)
     }
 
-    def bundle_info(platform: Platform.Family.Value): Bundle_Info =
+    def bundle_info(platform: Platform.Family): Bundle_Info =
       platform match {
         case Platform.Family.linux_arm =>
           Bundle_Info(platform, "Linux (ARM)", dist_name + "_linux_arm.tar.gz")
@@ -82,7 +82,7 @@
   }
 
   sealed case class Bundle_Info(
-    platform: Platform.Family.Value,
+    platform: Platform.Family,
     platform_description: String,
     name: String
   ) {
@@ -150,7 +150,7 @@
 
   /* bundled components */
 
-  class Bundled(platform: Option[Platform.Family.Value] = None) {
+  class Bundled(platform: Option[Platform.Family] = None) {
     def detect(s: String): Boolean =
       s.startsWith("#bundled") && !s.startsWith("#bundled ")
 
@@ -186,17 +186,16 @@
         } yield bundled(line)).toList))
   }
 
-  def get_bundled_components(dir: Path, platform: Platform.Family.Value): (List[String], String) = {
+  def get_bundled_components(dir: Path, platform: Platform.Family): (List[String], String) = {
     val Bundled = new Bundled(platform = Some(platform))
     val components =
-      for { Bundled(name) <- Components.Directory(dir).read_components() } yield name
+      for { case Bundled(name) <- Components.Directory(dir).read_components() } yield name
     val jdk_component =
       components.find(_.startsWith("jdk")) getOrElse error("Missing jdk component")
     (components, jdk_component)
   }
 
-  def activate_components(
-    dir: Path, platform: Platform.Family.Value, more_names: List[String]): Unit = {
+  def activate_components(dir: Path, platform: Platform.Family, more_names: List[String]): Unit = {
     def contrib_name(name: String): String =
       Components.contrib(name = name).implode
 
@@ -219,7 +218,7 @@
 
   private def build_heaps(
     options: Options,
-    platform: Platform.Family.Value,
+    platform: Platform.Family,
     build_sessions: List[String],
     local_dir: Path,
     progress: Progress = new Progress,
@@ -275,7 +274,7 @@
   }
 
   def make_isabelle_app(
-    platform: Platform.Family.Value,
+    platform: Platform.Family,
     isabelle_target: Path,
     isabelle_name: String,
     jdk_component: String,
@@ -491,13 +490,13 @@
     }
   }
 
-  def default_platform_families: List[Platform.Family.Value] = Platform.Family.list0
+  def default_platform_families: List[Platform.Family] = Platform.Family.list0
 
   def build_release(
     options: Options,
     context: Release_Context,
     afp_rev: String = "",
-    platform_families: List[Platform.Family.Value] = default_platform_families,
+    platform_families: List[Platform.Family] = default_platform_families,
     more_components: List[Path] = Nil,
     website: Option[Path] = None,
     build_sessions: List[String] = Nil,
--- a/src/Pure/Admin/build_status.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Admin/build_status.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -196,7 +196,7 @@
     maximum_heap: Space,
     average_heap: Space,
     stored_heap: Space,
-    status: Build_Log.Session_Status.Value,
+    status: Build_Log.Session_Status,
     errors: List[String]
   ) {
     val date: Long = (afp_pull_date getOrElse pull_date).unix_epoch
@@ -323,7 +323,7 @@
                   maximum_heap = Space.B(ml_stats.maximum(ML_Statistics.HEAP_SIZE)),
                   average_heap = Space.B(ml_stats.average(ML_Statistics.HEAP_SIZE)),
                   stored_heap = Space.bytes(res.long(Build_Log.Data.heap_size)),
-                  status = Build_Log.Session_Status.withName(res.string(Build_Log.Data.status)),
+                  status = Build_Log.Session_Status.valueOf(res.string(Build_Log.Data.status)),
                   errors =
                     Build_Log.uncompress_errors(
                       res.bytes(Build_Log.Data.errors), cache = store.cache))
--- a/src/Pure/Admin/isabelle_cronjob.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Admin/isabelle_cronjob.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -307,14 +307,6 @@
         }
       }
 
-  val remote_build_mini3 =
-    Remote_Build("macOS 13 Ventura (ARM64)", "mini3",
-      history_base = "8e590adaac5e",
-      options = "-a -m32 -B -M1x4,2x2,4 -p pide_session=false" +
-        " -e ISABELLE_MLTON=/opt/homebrew/bin/mlton -e ISABELLE_MLTON_OPTIONS=" +
-        " -e ISABELLE_SWIPL=/opt/homebrew/bin/swipl",
-      args = "-a -d '~~/src/Benchmarks'")
-
   val remote_builds1: List[List[Remote_Build]] = {
     List(
       List(Remote_Build("Linux A", "augsburg1",
@@ -355,7 +347,13 @@
         Remote_Build("macOS, skip_proofs", "mini2",
           options = "-m32 -M4 -t skip_proofs -p pide_session=false", args = "-a -o skip_proofs",
           detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("skip_proofs"))),
-      List(remote_build_mini3, remote_build_mini3, remote_build_mini3),
+      List(
+        Remote_Build("macOS 13 Ventura (ARM64)", "mini3",
+          history_base = "8e590adaac5e",
+          options = "-a -m32 -B -M1x4,2x2,4 -p pide_session=false" +
+            " -e ISABELLE_MLTON=/opt/homebrew/bin/mlton -e ISABELLE_MLTON_OPTIONS=" +
+            " -e ISABELLE_SWIPL=/opt/homebrew/bin/swipl",
+          args = "-a -d '~~/src/Benchmarks'")),
       List(
         Remote_Build("macOS 12 Monterey", "monterey", user = "makarius",
           options = "-m32 -M4 -e ISABELLE_GHC_SETUP=true -p pide_session=false",
--- a/src/Pure/Concurrent/consumer_thread.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Concurrent/consumer_thread.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -89,7 +89,7 @@
 
     val (results, cont) = consume(reqs.map(_.arg))
     for {
-      (Some(req), Some(res)) <- reqs.map(Some(_)).zipAll(results.map(Some(_)), None, None)
+      case (Some(req), Some(res)) <- reqs.map(Some(_)).zipAll(results.map(Some(_)), None, None)
     } {
       (req.ack, res) match {
         case (Some(a), _) => a.change(_ => Some(res))
--- a/src/Pure/Concurrent/future.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Concurrent/future.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -65,30 +65,32 @@
 /* task future via thread pool */
 
 private class Task_Future[A](body: => A) extends Future[A] {
-  private sealed abstract class Status
-  private case object Ready extends Status
-  private case class Running(thread: Thread) extends Status
-  private case object Terminated extends Status
-  private case class Finished(result: Exn.Result[A]) extends Status
+  private enum Status {
+    case Ready extends Status
+    case Running(thread: Thread) extends Status
+    case Terminated extends Status
+    case Finished(result: Exn.Result[A]) extends Status
+  }
 
-  private val status = Synchronized[Status](Ready)
+  private val status = Synchronized[Status](Status.Ready)
 
   def peek: Option[Exn.Result[A]] =
     status.value match {
-      case Finished(result) => Some(result)
+      case Status.Finished(result) => Some(result)
       case _ => None
     }
 
   private def try_run(): Unit = {
     val do_run =
       status.change_result {
-        case Ready => (true, Running(Thread.currentThread))
+        case Status.Ready => (true, Status.Running(Thread.currentThread))
         case st => (false, st)
       }
     if (do_run) {
       val result = Exn.capture(body)
-      status.change(_ => Terminated)
-      status.change(_ => Finished(if (Thread.interrupted()) Exn.Exn(Exn.Interrupt()) else result))
+      status.change(_ => Status.Terminated)
+      status.change(_ =>
+        Status.Finished(if (Thread.interrupted()) Exn.Exn(Exn.Interrupt()) else result))
     }
   }
   private val task = Isabelle_Thread.pool.submit(new Callable[Unit] { def call = try_run() })
@@ -96,15 +98,15 @@
   def join_result: Exn.Result[A] = {
     try_run()
     status.guarded_access {
-      case st @ Finished(result) => Some((result, st))
+      case st @ Status.Finished(result) => Some((result, st))
       case _ => None
     }
   }
 
   def cancel(): Unit = {
     status.change {
-      case Ready => task.cancel(false); Finished(Exn.Exn(Exn.Interrupt()))
-      case st @ Running(thread) => thread.interrupt(); st
+      case Status.Ready => task.cancel(false); Status.Finished(Exn.Exn(Exn.Interrupt()))
+      case st @ Status.Running(thread) => thread.interrupt(); st
       case st => st
     }
   }
--- a/src/Pure/General/json.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/General/json.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -41,11 +41,9 @@
 
   /* lexer */
 
-  object Kind extends Enumeration {
-    val KEYWORD, STRING, NUMBER, ERROR = this.Value
-  }
+  enum Kind { case KEYWORD, STRING, NUMBER, ERROR }
 
-  sealed case class Token(kind: Kind.Value, text: String) {
+  sealed case class Token(kind: Kind, text: String) {
     def is_keyword: Boolean = kind == Kind.KEYWORD
     def is_keyword(name: String): Boolean = kind == Kind.KEYWORD && text == name
     def is_string: Boolean = kind == Kind.STRING
--- a/src/Pure/General/json_api.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/General/json_api.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -69,7 +69,7 @@
   sealed case class Links(json: JSON.T) {
     def get_next: Option[Service] =
       for {
-        JSON.Value.String(next) <- JSON.value(json, "next")
+        case JSON.Value.String(next) <- JSON.value(json, "next")
         if Url.is_wellformed(next)
       } yield new Service(Url(next))
 
--- a/src/Pure/General/mercurial.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/General/mercurial.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -100,7 +100,7 @@
 
   sealed case class Archive_Info(lines: List[String]) {
     def id: Option[String] = lines.collectFirst({ case Archive_Node(a) => a })
-    def tags: List[String] = for (Archive_Tag(tag) <- lines if tag != "tip") yield tag
+    def tags: List[String] = for (case Archive_Tag(tag) <- lines if tag != "tip") yield tag
   }
 
   def archive_info(root: Path): Option[Archive_Info] = {
--- a/src/Pure/General/sql.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/General/sql.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -95,26 +95,26 @@
 
   /* types */
 
-  object Type extends Enumeration {
-    val Boolean = Value("BOOLEAN")
-    val Int = Value("INTEGER")
-    val Long = Value("BIGINT")
-    val Double = Value("DOUBLE PRECISION")
-    val String = Value("TEXT")
-    val Bytes = Value("BLOB")
-    val Date = Value("TIMESTAMP WITH TIME ZONE")
-  }
+  enum Type { case Boolean, Int, Long, Double, String, Bytes, Date }
 
-  def sql_type_default(T: Type.Value): Source = T.toString
+  val sql_type_postgresql: Type => Source =
+    {
+      case Type.Boolean => "BOOLEAN"
+      case Type.Int => "INTEGER"
+      case Type.Long => "BIGINT"
+      case Type.Double => "DOUBLE PRECISION"
+      case Type.String => "TEXT"
+      case Type.Bytes => "BYTEA"
+      case Type.Date => "TIMESTAMP WITH TIME ZONE"
+    }
 
-  def sql_type_sqlite(T: Type.Value): Source =
-    if (T == Type.Boolean) "INTEGER"
-    else if (T == Type.Date) "TEXT"
-    else sql_type_default(T)
-
-  def sql_type_postgresql(T: Type.Value): Source =
-    if (T == Type.Bytes) "BYTEA"
-    else sql_type_default(T)
+  val sql_type_sqlite: Type => Source =
+    {
+      case Type.Boolean => "INTEGER"
+      case Type.Bytes => "BLOB"
+      case Type.Date => "TEXT"
+      case t => sql_type_postgresql(t)
+    }
 
 
   /* columns */
@@ -138,7 +138,7 @@
 
   sealed case class Column(
     name: String,
-    T: Type.Value,
+    T: Type,
     strict: Boolean = false,
     primary_key: Boolean = false,
     expr: SQL.Source = ""
@@ -152,7 +152,7 @@
       if (expr == "") SQL.ident(name)
       else enclose(expr) + " AS " + SQL.ident(name)
 
-    def decl(sql_type: Type.Value => Source): Source =
+    def decl(sql_type: Type => Source): Source =
       ident + " " + sql_type(T) + (if (strict || primary_key) " NOT NULL" else "")
 
     def defined: String = ident + " IS NOT NULL"
@@ -194,7 +194,7 @@
 
     def query_named: Source = query + " AS " + SQL.ident(name)
 
-    def create(sql_type: Type.Value => Source): Source = {
+    def create(sql_type: Type => Source): Source = {
       val primary_key =
         columns.filter(_.primary_key).map(_.name) match {
           case Nil => Nil
@@ -407,7 +407,7 @@
 
     /* types */
 
-    def sql_type(T: Type.Value): Source
+    def sql_type(T: Type): Source
 
 
     /* connection */
@@ -627,7 +627,7 @@
 
     override def now(): Date = Date.now()
 
-    def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_sqlite(T)
+    def sql_type(T: SQL.Type): SQL.Source = SQL.sql_type_sqlite(T)
 
     def update_date(stmt: SQL.Statement, i: Int, date: Date): Unit =
       if (date == null) stmt.string(i) = (null: String)
@@ -751,7 +751,7 @@
         .getOrElse(error("Failed to get current date/time from database server " + toString))
     }
 
-    def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_postgresql(T)
+    def sql_type(T: SQL.Type): SQL.Source = SQL.sql_type_postgresql(T)
 
     // see https://jdbc.postgresql.org/documentation/head/8-date-time.html
     def update_date(stmt: SQL.Statement, i: Int, date: Date): Unit =
--- a/src/Pure/General/ssh.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/General/ssh.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -493,7 +493,7 @@
 
     def isabelle_platform: Isabelle_Platform = Isabelle_Platform()
 
-    def isabelle_platform_family: Platform.Family.Value =
+    def isabelle_platform_family: Platform.Family =
       Platform.Family.parse(isabelle_platform.ISABELLE_PLATFORM_FAMILY)
   }
 
--- a/src/Pure/General/symbol.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/General/symbol.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -286,13 +286,13 @@
 
   /** defined symbols **/
 
-  object Argument extends Enumeration {
-    val none, cartouche, space_cartouche = Value
+  object Argument {
+    def unapply(s: String): Option[Argument] =
+      try { Some(valueOf(s)) }
+      catch { case _: IllegalArgumentException => None}
+  }
 
-    def unapply(s: String): Option[Value] =
-      try { Some(withName(s)) }
-      catch { case _: NoSuchElementException => None}
-  }
+  enum Argument { case none, cartouche, space_cartouche }
 
   object Entry {
     private val Name = new Regex("""\\<\^?([A-Za-z][A-Za-z0-9_']*)>""")
@@ -327,9 +327,10 @@
           case _ => None
         }
 
-      val groups = proper_list(for ((Group.name, a) <- props) yield a).getOrElse(List("unsorted"))
+      val groups =
+        proper_list(for (case (Group.name, a) <- props) yield a).getOrElse(List("unsorted"))
 
-      val abbrevs = for ((Abbrev.name, a) <- props) yield a
+      val abbrevs = for (case (Abbrev.name, a) <- props) yield a
 
       new Entry(symbol, name, argument, code, Font.unapply(props), groups, abbrevs)
     }
@@ -338,7 +339,7 @@
   class Entry private(
     val symbol: Symbol,
     val name: String,
-    val argument: Symbol.Argument.Value,
+    val argument: Symbol.Argument,
     val code: Option[Int],
     val font: Option[String],
     val groups: List[String],
--- a/src/Pure/General/toml.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/General/toml.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -133,11 +133,9 @@
 
   /* lexer */
 
-  object Kind extends Enumeration {
-    val KEYWORD, VALUE, STRING, MULTILINE_STRING, LINE_SEP, ERROR = Value
-  }
+  enum Kind { case KEYWORD, VALUE, STRING, MULTILINE_STRING, LINE_SEP, ERROR }
 
-  sealed case class Token(kind: Kind.Value, text: Str) {
+  sealed case class Token(kind: Kind, text: Str) {
     def is_keyword(name: Str): Bool = kind == Kind.KEYWORD && text == name
     def is_value: Bool = kind == Kind.VALUE
     def is_string: Bool = kind == Kind.STRING
--- a/src/Pure/General/word.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/General/word.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -25,44 +25,41 @@
   def lowercase(str: String): String = str.toLowerCase(Locale.ROOT)
   def uppercase(str: String): String = str.toUpperCase(Locale.ROOT)
 
-  def capitalize(str: String): String =
+  def capitalized(str: String): String =
     if (str.length == 0) str
     else {
       val n = Character.charCount(str.codePointAt(0))
       uppercase(str.substring(0, n)) + lowercase(str.substring(n))
     }
 
-  def perhaps_capitalize(str: String): String =
+  def perhaps_capitalized(str: String): String =
     if (Codepoint.iterator(str).forall(c => Character.isLowerCase(c) || Character.isDigit(c)))
-      capitalize(str)
+      capitalized(str)
     else str
 
-  sealed abstract class Case
-  case object Lowercase extends Case
-  case object Uppercase extends Case
-  case object Capitalized extends Case
-
   object Case {
     def apply(c: Case, str: String): String =
       c match {
-        case Lowercase => lowercase(str)
-        case Uppercase => uppercase(str)
-        case Capitalized => capitalize(str)
+        case Case.lowercase => Word.lowercase(str)
+        case Case.uppercase => Word.uppercase(str)
+        case Case.capitalized => Word.capitalized(str)
       }
     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(Case.lowercase)
+        else if (Codepoint.iterator(str).forall(Character.isUpperCase)) Some(Case.uppercase)
         else {
           val it = Codepoint.iterator(str)
           if (Character.isUpperCase(it.next()) && it.forall(Character.isLowerCase))
-            Some(Capitalized)
+            Some(Case.capitalized)
           else None
         }
       }
       else None
   }
 
+  enum Case { case lowercase, uppercase, capitalized }
+
 
   /* sequence of words */
 
--- a/src/Pure/PIDE/command.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/PIDE/command.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -147,13 +147,13 @@
       else other.rep.iterator.foldLeft(this)(_ + _)
 
     def redirection_iterator: Iterator[Document_ID.Generic] =
-      for (Markup_Index(_, Symbol.Text_Chunk.Id(id)) <- rep.keysIterator)
+      for (case Markup_Index(_, Symbol.Text_Chunk.Id(id)) <- rep.keysIterator)
         yield id
 
     def redirect(other_id: Document_ID.Generic): Markups = {
       val rep1 =
         (for {
-          (Markup_Index(status, Symbol.Text_Chunk.Id(id)), markup) <- rep.iterator
+          case (Markup_Index(status, Symbol.Text_Chunk.Id(id)), markup) <- rep.iterator
           if other_id == id
         } yield (Markup_Index(status, Symbol.Text_Chunk.Default), markup)).toMap
       if (rep1.isEmpty) Markups.empty else new Markups(rep1)
@@ -500,13 +500,13 @@
   def blobs_ok: Boolean = blobs.forall(Exn.the_res.isDefinedAt)
 
   def blobs_names: List[Document.Node.Name] =
-    for (Exn.Res(blob) <- blobs) yield blob.name
+    for (case Exn.Res(blob) <- blobs) yield blob.name
 
   def blobs_undefined: List[Document.Node.Name] =
-    for (Exn.Res(blob) <- blobs if blob.content.isEmpty) yield blob.name
+    for (case Exn.Res(blob) <- blobs if blob.content.isEmpty) yield blob.name
 
   def blobs_defined: List[(Document.Node.Name, SHA1.Digest)] =
-    for (Exn.Res(blob) <- blobs; (digest, _) <- blob.content) yield (blob.name, digest)
+    for (case Exn.Res(blob) <- blobs; (digest, _) <- blob.content) yield (blob.name, digest)
 
   def blobs_changed(doc_blobs: Document.Blobs): Boolean =
     blobs.exists({ case Exn.Res(blob) => doc_blobs.changed(blob.name) case _ => false })
@@ -518,7 +518,7 @@
 
   val chunks: Map[Symbol.Text_Chunk.Name, Symbol.Text_Chunk] =
     ((Symbol.Text_Chunk.Default -> chunk) ::
-      (for (Exn.Res(blob) <- blobs; (_, file) <- blob.content)
+      (for (case Exn.Res(blob) <- blobs; (_, file) <- blob.content)
         yield blob.chunk_file -> file)).toMap
 
   def length: Int = source.length
--- a/src/Pure/PIDE/document.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/PIDE/document.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -785,7 +785,7 @@
             case some => Some(some)
           }
       }
-      for (Text.Info(r, Some(x)) <- cumulate(range, None, elements, result1, status))
+      for (case Text.Info(r, Some(x)) <- cumulate(range, None, elements, result1, status))
         yield Text.Info(r, x)
     }
   }
@@ -1262,7 +1262,7 @@
       val pending_edits1 =
         (for {
           change <- history.undo_list.takeWhile(_ != stable)
-          (name, Node.Edits(es)) <- change.rev_edits
+          case (name, Node.Edits(es)) <- change.rev_edits
         } yield (name -> es)).foldLeft(pending_edits)(_ + _)
 
       new Snapshot(this, version, node_name, pending_edits1, snippet_command)
--- a/src/Pure/PIDE/document_status.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/PIDE/document_status.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -230,9 +230,7 @@
 
   /* nodes status */
 
-  object Overall_Node_Status extends Enumeration {
-    val ok, failed, pending = Value
-  }
+  enum Overall_Node_Status { case ok, failed, pending }
 
   object Nodes_Status {
     val empty: Nodes_Status = new Nodes_Status(Map.empty, Document.Nodes.empty)
@@ -259,7 +257,7 @@
         case None => false
       }
 
-    def overall_node_status(name: Document.Node.Name): Overall_Node_Status.Value =
+    def overall_node_status(name: Document.Node.Name): Overall_Node_Status =
       rep.get(name) match {
         case Some(st) if st.consolidated =>
           if (st.ok) Overall_Node_Status.ok else Overall_Node_Status.failed
--- a/src/Pure/PIDE/query_operation.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/PIDE/query_operation.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -9,11 +9,7 @@
 
 
 object Query_Operation {
-  object Status extends Enumeration {
-    val WAITING = Value("waiting")
-    val RUNNING = Value("running")
-    val FINISHED = Value("finished")
-  }
+  enum Status { case waiting, running, finished }
 
   object State {
     val empty: State = State()
@@ -22,7 +18,7 @@
       State(instance = Document_ID.make().toString,
         location = Some(command),
         query = query,
-        status = Status.WAITING)
+        status = Status.waiting)
   }
 
   sealed case class State(
@@ -31,7 +27,7 @@
     query: List[String] = Nil,
     update_pending: Boolean = false,
     output: List[XML.Tree] = Nil,
-    status: Status.Value = Status.FINISHED,
+    status: Status = Status.finished,
     exec_id: Document_ID.Exec = Document_ID.none)
 }
 
@@ -39,7 +35,7 @@
   editor: Editor[Editor_Context],
   editor_context: Editor_Context,
   operation_name: String,
-  consume_status: Query_Operation.Status.Value => Unit,
+  consume_status: Query_Operation.Status => Unit,
   consume_output: (Document.Snapshot, Command.Results, XML.Body) => Unit
 ) {
   private val print_function = operation_name + "_query"
@@ -76,7 +72,7 @@
           val command_results = snapshot.command_results(cmd)
           val results =
             (for {
-              (_, elem @ XML.Elem(Markup(Markup.RESULT, props), _)) <- command_results.iterator
+              case (_, elem @ XML.Elem(Markup(Markup.RESULT, props), _)) <- command_results.iterator
               if props.contains((Markup.INSTANCE, state0.instance))
             } yield elem).toList
           val removed = !snapshot.get_node(cmd.node_name).commands.contains(cmd)
@@ -116,7 +112,7 @@
 
     val new_output =
       for {
-        XML.Elem(_, List(XML.Elem(markup, body))) <- results
+        case XML.Elem(_, List(XML.Elem(markup, body))) <- results
         if Markup.messages.contains(markup.name)
         body1 = resolve_sendback(body)
       } yield Protocol.make_message(body1, markup.name, props = markup.properties)
@@ -124,21 +120,20 @@
 
     /* status */
 
-    def get_status(name: String, status: Query_Operation.Status.Value)
-        : Option[Query_Operation.Status.Value] =
+    def get_status(name: String, status: Query_Operation.Status): Option[Query_Operation.Status] =
       results.collectFirst({ case XML.Elem(_, List(elem: XML.Elem)) if elem.name == name => status })
 
     val new_status =
-      if (removed) Query_Operation.Status.FINISHED
+      if (removed) Query_Operation.Status.finished
       else
-        get_status(Markup.FINISHED, Query_Operation.Status.FINISHED) orElse
-        get_status(Markup.RUNNING, Query_Operation.Status.RUNNING) getOrElse
-        Query_Operation.Status.WAITING
+        get_status(Markup.FINISHED, Query_Operation.Status.finished) orElse
+        get_status(Markup.RUNNING, Query_Operation.Status.running) getOrElse
+        Query_Operation.Status.waiting
 
 
     /* state update */
 
-    if (new_status == Query_Operation.Status.RUNNING)
+    if (new_status == Query_Operation.Status.running)
       results.collectFirst(
       {
         case XML.Elem(Markup(_, Position.Id(id)), List(elem: XML.Elem))
@@ -157,7 +152,7 @@
         if (state0.status != new_status) {
           current_state.change(_.copy(status = new_status))
           consume_status(new_status)
-          if (new_status == Query_Operation.Status.FINISHED)
+          if (new_status == Query_Operation.Status.finished)
             remove_overlay()
         }
       }
@@ -214,7 +209,7 @@
         state.location match {
           case Some(command)
           if state.update_pending ||
-            (state.status != Query_Operation.Status.FINISHED &&
+            (state.status != Query_Operation.Status.finished &&
               changed.commands.contains(command)) =>
             editor.send_dispatcher { content_update() }
           case _ =>
@@ -230,6 +225,6 @@
     remove_overlay()
     current_state.change(_ => Query_Operation.State.empty)
     consume_output(Document.Snapshot.init, Command.Results.empty, Nil)
-    consume_status(Query_Operation.Status.FINISHED)
+    consume_status(Query_Operation.Status.finished)
   }
 }
--- a/src/Pure/PIDE/rendering.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/PIDE/rendering.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -405,7 +405,7 @@
                 Some(snapshot.convert(info.range))
               else None)
         })
-    for (Text.Info(range, Some(range1)) <- result)
+    for (case Text.Info(range, Some(range1)) <- result)
       yield Text.Info(range, range1)
   }
 
--- a/src/Pure/ROOT.ML	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/ROOT.ML	Sun Sep 03 19:28:59 2023 +0200
@@ -370,3 +370,4 @@
 ML_file "Tools/jedit.ML";
 ML_file "Tools/ghc.ML";
 ML_file "Tools/generated_files.ML";
+
--- a/src/Pure/System/components.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/System/components.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -35,7 +35,7 @@
 
   /* platforms */
 
-  private val family_platforms: Map[Platform.Family.Value, List[String]] =
+  private val family_platforms: Map[Platform.Family, List[String]] =
     Map(
       Platform.Family.linux_arm -> List("arm64-linux", "arm64_32-linux"),
       Platform.Family.linux -> List("x86_64-linux", "x86_64_32-linux"),
@@ -47,7 +47,7 @@
   private val platform_names: Set[String] =
     Set("x86-linux", "x86-cygwin") ++ family_platforms.iterator.flatMap(_._2)
 
-  def platform_purge(platforms: List[Platform.Family.Value]): String => Boolean = {
+  def platform_purge(platforms: List[Platform.Family]): String => Boolean = {
     val preserve =
       (for {
         family <- platforms.iterator
@@ -91,7 +91,7 @@
 
   def clean(
     component_dir: Path,
-    platforms: List[Platform.Family.Value] = Platform.Family.list,
+    platforms: List[Platform.Family] = Platform.Family.list,
     ssh: SSH.System = SSH.Local,
     progress: Progress = new Progress
   ): Unit = {
@@ -108,7 +108,7 @@
 
   def clean_base(
     base_dir: Path,
-    platforms: List[Platform.Family.Value] = Platform.Family.list,
+    platforms: List[Platform.Family] = Platform.Family.list,
     ssh: SSH.System = SSH.Local,
     progress: Progress = new Progress
   ): Unit = {
@@ -124,7 +124,7 @@
     name: String,
     target_dir: Option[Path] = None,
     copy_dir: Option[Path] = None,
-    clean_platforms: Option[List[Platform.Family.Value]] = None,
+    clean_platforms: Option[List[Platform.Family]] = None,
     clean_archives: Boolean = false,
     component_repository: String = Components.static_component_repository,
     ssh: SSH.System = SSH.Local,
--- a/src/Pure/System/executable.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/System/executable.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -31,7 +31,7 @@
       if (Platform.is_macos) {
         val Pattern = """^\s*(/.+)\s+\(.*\)$""".r
         for {
-          Pattern(lib) <- ldd_lines
+          case Pattern(lib) <- ldd_lines
           if !lib.startsWith("@executable_path/") && filter(lib_name(lib))
         } yield lib
       }
@@ -42,7 +42,7 @@
             case None => ""
             case Some(path) => path.absolute.implode
           }
-        for { Pattern(lib) <- ldd_lines if filter(lib_name(lib)) }
+        for { case Pattern(lib) <- ldd_lines if filter(lib_name(lib)) }
           yield prefix + lib
       }
 
--- a/src/Pure/System/isabelle_system.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/System/isabelle_system.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -481,7 +481,7 @@
             (entry, result)
           }
         for {
-          (entry, Some(res)) <- items
+          case (entry, Some(res)) <- items
           if !entry.isDirectory
           t <- Option(entry.getLastModifiedTime)
         } Files.setLastModifiedTime(res, t)
--- a/src/Pure/System/options.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/System/options.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -117,7 +117,7 @@
           case word :: rest if word == strip => rest
           case _ => words
         }
-      Word.implode(words1.map(Word.perhaps_capitalize))
+      Word.implode(words1.map(Word.perhaps_capitalized))
     }
     def title_jedit: String = title("jedit")
 
--- a/src/Pure/System/other_isabelle.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/System/other_isabelle.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -89,7 +89,7 @@
 
   def resolve_components(
     echo: Boolean = false,
-    clean_platforms: Option[List[Platform.Family.Value]] = None,
+    clean_platforms: Option[List[Platform.Family]] = None,
     clean_archives: Boolean = false,
     component_repository: String = Components.static_component_repository
   ): Unit = {
@@ -152,7 +152,7 @@
     other_settings: List[String] = init_components(),
     fresh: Boolean = false,
     echo: Boolean = false,
-    clean_platforms: Option[List[Platform.Family.Value]] = None,
+    clean_platforms: Option[List[Platform.Family]] = None,
     clean_archives: Boolean = false,
     component_repository: String = Components.static_component_repository
   ): Unit = {
--- a/src/Pure/System/platform.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/System/platform.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -17,38 +17,42 @@
 
   def is_arm: Boolean = cpu_arch.startsWith("arm")
 
-  def family: Family.Value =
+  def family: Family =
     if (is_linux && is_arm) Family.linux_arm
     else if (is_linux) Family.linux
     else if (is_macos) Family.macos
     else if (is_windows) Family.windows
     else error("Failed to determine current platform family")
 
-  object Family extends Enumeration {
-    val linux_arm, linux, macos, windows = Value
-    val list0: List[Value] = List(linux, windows, macos)
-    val list: List[Value] = List(linux, linux_arm, windows, macos)
+  object Family {
+    val list0: List[Family] = List(Family.linux, Family.windows, Family.macos)
+    val list: List[Family] = List(Family.linux, Family.linux_arm, Family.windows, Family.macos)
 
-    def unapply(name: String): Option[Value] =
-      try { Some(withName(name)) }
-      catch { case _: NoSuchElementException => None }
+    def unapply(name: String): Option[Family] =
+      try { Some(Family.valueOf(name)) }
+      catch { case _: IllegalArgumentException => None }
 
-    def parse(name: String): Value =
+    def parse(name: String): Family =
       unapply(name) getOrElse error("Bad platform family: " + quote(name))
 
-    def standard(platform: Value): String =
-      if (platform == linux_arm) "arm64-linux"
-      else if (platform == linux) "x86_64-linux"
-      else if (platform == macos) "x86_64-darwin"
-      else if (platform == windows) "x86_64-cygwin"
-      else error("Bad platform family: " + quote(platform.toString))
+    val standard: Family => String =
+      {
+        case Family.linux_arm => "arm64-linux"
+        case Family.linux => "x86_64-linux"
+        case Family.macos => "x86_64-darwin"
+        case Family.windows => "x86_64-cygwin"
+      }
 
-    def native(platform: Value): String =
-      if (platform == macos) "arm64-darwin"
-      else if (platform == windows) "x86_64-windows"
-      else standard(platform)
+    val native: Family => String =
+      {
+        case Family.macos => "arm64-darwin"
+        case Family.windows => "x86_64-windows"
+        case platform => standard(platform)
+      }
   }
 
+  enum Family { case linux_arm, linux, macos, windows }
+
 
   /* platform identifiers */
 
--- a/src/Pure/System/progress.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/System/progress.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -18,9 +18,9 @@
 
   sealed abstract class Output { def message: Message }
 
-  object Kind extends Enumeration { val writeln, warning, error_message = Value }
+  enum Kind { case writeln, warning, error_message }
   sealed case class Message(
-    kind: Kind.Value,
+    kind: Kind,
     text: String,
     verbose: Boolean = false
   ) extends Output {
@@ -157,7 +157,7 @@
         SortedMap.from[Long, Message],
         { res =>
           val serial = res.long(Messages.serial)
-          val kind = Kind(res.int(Messages.kind))
+          val kind = Kind.fromOrdinal(res.int(Messages.kind))
           val text = res.string(Messages.text)
           val verbose = res.bool(Messages.verbose)
           serial -> Message(kind, text, verbose = verbose)
@@ -173,7 +173,7 @@
         for ((serial, message) <- messages) yield { (stmt: SQL.Statement) =>
           stmt.long(1) = context
           stmt.long(2) = serial
-          stmt.int(3) = message.kind.id
+          stmt.int(3) = message.kind.ordinal
           stmt.string(4) = message.text
           stmt.bool(5) = message.verbose
         })
--- a/src/Pure/System/scala.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/System/scala.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -102,11 +102,10 @@
 
   object Compiler {
     object Message {
-      object Kind extends Enumeration {
-        val error, warning, info, other = Value
-      }
+      enum Kind { case error, warning, info, other }
+
       private val Header = """^--.* (Error|Warning|Info): .*$""".r
-      val header_kind: String => Kind.Value =
+      val header_kind: String => Kind =
         {
           case "Error" => Kind.error
           case "Warning" => Kind.warning
@@ -139,7 +138,7 @@
       }
     }
 
-    sealed case class Message(kind: Message.Kind.Value, text: String)
+    sealed case class Message(kind: Message.Kind, text: String)
     {
       def is_error: Boolean = kind == Message.Kind.error
       override def toString: String = text
@@ -275,9 +274,7 @@
 
   /* invoke function */
 
-  object Tag extends Enumeration {
-    val NULL, OK, ERROR, FAIL, INTERRUPT = Value
-  }
+  enum Tag { case NULL, OK, ERROR, FAIL, INTERRUPT }
 
   def function_thread(name: String): Boolean =
     functions.find(fun => fun.name == name) match {
@@ -285,7 +282,7 @@
       case None => false
     }
 
-  def function_body(session: Session, name: String, args: List[Bytes]): (Tag.Value, List[Bytes]) =
+  def function_body(session: Session, name: String, args: List[Bytes]): (Tag, List[Bytes]) =
     functions.find(fun => fun.name == name) match {
       case Some(fun) =>
         Exn.capture { fun.invoke(session, args) } match {
@@ -312,10 +309,11 @@
       futures = Map.empty
     }
 
-    private def result(id: String, tag: Scala.Tag.Value, res: List[Bytes]): Unit =
+    private def result(id: String, tag: Scala.Tag, res: List[Bytes]): Unit =
       synchronized {
         if (futures.isDefinedAt(id)) {
-          session.protocol_command_raw("Scala.result", Bytes(id) :: Bytes(tag.id.toString) :: res)
+          session.protocol_command_raw(
+            "Scala.result", Bytes(id) :: Bytes(tag.ordinal.toString) :: res)
           futures -= id
         }
       }
--- a/src/Pure/Thy/bibtex.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Thy/bibtex.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -132,11 +132,11 @@
               case (Error(msg, Value.Int(l)), _) =>
                 Some((true, (msg, get_line_pos(l))))
               case (Warning_in_Chunk(msg, name), _) if chunk_pos.isDefinedAt(name) =>
-                Some((false, (Word.capitalize(msg) + " in entry " + quote(name), chunk_pos(name))))
+                Some((false, (Word.capitalized(msg) + " in entry " + quote(name), chunk_pos(name))))
               case (Warning(msg), Warning_Line(Value.Int(l))) =>
-                Some((false, (Word.capitalize(msg), get_line_pos(l))))
+                Some((false, (Word.capitalized(msg), get_line_pos(l))))
               case (Warning(msg), _) =>
-                Some((false, (Word.capitalize(msg), Position.none)))
+                Some((false, (Word.capitalized(msg), Position.none)))
               case _ => None
             }
           ).partition(_._1)
--- a/src/Pure/Thy/export.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Thy/export.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -379,10 +379,10 @@
                 yield name -> store.try_open_database(name, server_mode = false)
             attempts.collectFirst({ case (name, None) => name }) match {
               case Some(bad) =>
-                for ((_, Some(db)) <- attempts) db.close()
+                for (case (_, Some(db)) <- attempts) db.close()
                 store.error_database(bad)
               case None =>
-                for ((name, Some(db)) <- attempts) yield {
+                for (case (name, Some(db)) <- attempts) yield {
                   new Session_Database(name, db) { override def close(): Unit = this.db.close() }
                 }
             }
--- a/src/Pure/Thy/export_theory.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Thy/export_theory.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -239,14 +239,12 @@
 
   /* approximative syntax */
 
-  object Assoc extends Enumeration {
-    val NO_ASSOC, LEFT_ASSOC, RIGHT_ASSOC = Value
-  }
+  enum Assoc { case NO_ASSOC, LEFT_ASSOC, RIGHT_ASSOC }
 
   sealed abstract class Syntax
   case object No_Syntax extends Syntax
   case class Prefix(delim: String) extends Syntax
-  case class Infix(assoc: Assoc.Value, delim: String, pri: Int) extends Syntax
+  case class Infix(assoc: Assoc, delim: String, pri: Int) extends Syntax
 
   def decode_syntax: XML.Decode.T[Syntax] =
     XML.Decode.variant(List(
@@ -255,7 +253,7 @@
       { case (Nil, body) =>
           import XML.Decode._
           val (ass, delim, pri) = triple(int, string, int)(body)
-          Infix(Assoc(ass), delim, pri) }))
+          Infix(Assoc.fromOrdinal(ass), delim, pri) }))
 
 
   /* types */
--- a/src/Pure/Thy/html.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Thy/html.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -151,7 +151,7 @@
   def check_control_blocks(body: XML.Body): Boolean = {
     var ok = true
     var open = List.empty[Symbol.Symbol]
-    for { XML.Text(text) <- body; sym <- Symbol.iterator(text) } {
+    for { case XML.Text(text) <- body; sym <- Symbol.iterator(text) } {
       if (is_control_block_begin(sym)) open ::= sym
       else if (is_control_block_end(sym)) {
         open match {
@@ -355,15 +355,13 @@
   /* GUI layout */
 
   object Wrap_Panel {
-    object Alignment extends Enumeration {
-      val left, right, center = Value
-    }
+    enum Alignment { case left, right, center }
 
     def apply(
       contents: List[XML.Elem],
       name: String = "",
       action: String = "",
-      alignment: Alignment.Value = Alignment.right
+      alignment: Alignment = Alignment.right
     ): XML.Elem = {
       val body = Library.separate(XML.Text(" "), contents)
       GUI.form(List(div(body) + ("style" -> ("text-align: " + alignment))),
--- a/src/Pure/Thy/latex.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Thy/latex.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -93,9 +93,7 @@
   /* tags */
 
   object Tags {
-    object Op extends Enumeration {
-      val fold, drop, keep = Value
-    }
+    enum Op { case fold, drop, keep }
 
     val standard = "document,theory,proof,ML,visible,-invisible,important,unimportant"
 
@@ -103,7 +101,7 @@
 
     def apply(spec: String): Tags =
       new Tags(spec,
-        (explode(standard) ::: explode(spec)).foldLeft(TreeMap.empty[String, Op.Value]) {
+        (explode(standard) ::: explode(spec)).foldLeft(TreeMap.empty[String, Op]) {
           case (m, tag) =>
             tag.toList match {
               case '/' :: cs => m + (cs.mkString -> Op.fold)
@@ -116,10 +114,10 @@
     val empty: Tags = apply("")
   }
 
-  class Tags private(spec: String, map: TreeMap[String, Tags.Op.Value]) {
+  class Tags private(spec: String, map: TreeMap[String, Tags.Op]) {
     override def toString: String = spec
 
-    def get(name: String): Option[Tags.Op.Value] = map.get(name)
+    def get(name: String): Option[Tags.Op] = map.get(name)
 
     def sty(comment_latex: Boolean): File.Content = {
       val path = Path.explode("isabelletags.sty")
--- a/src/Pure/Tools/build.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/build.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -753,7 +753,7 @@
 
       val results =
         Command.Results.make(
-          for (elem @ XML.Elem(Markup(_, Markup.Serial(i)), _) <- read_xml(Export.MESSAGES))
+          for (case elem@XML.Elem(Markup(_, Markup.Serial(i)), _) <- read_xml(Export.MESSAGES))
             yield i -> elem)
 
       val command =
--- a/src/Pure/Tools/build_cluster.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/build_cluster.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -263,7 +263,7 @@
       _sessions
     }
     else {
-      for (Exn.Res(session) <- attempts) session.close()
+      for (case Exn.Res(session) <- attempts) session.close()
       error("Failed to connect build cluster")
     }
   }
--- a/src/Pure/Tools/check_keywords.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/check_keywords.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -22,7 +22,7 @@
 
       val result =
         parse_all(rep(item), Token.reader(Token.explode(keywords, input), start)) match {
-          case Success(res, _) => for (Some(x) <- res) yield x
+          case Success(res, _) => for (case Some(x) <- res) yield x
           case bad => error(bad.toString)
         }
     }
--- a/src/Pure/Tools/dotnet_setup.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/dotnet_setup.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -11,7 +11,7 @@
   /* platforms */
 
   sealed case class Platform_Info(
-    family: Platform.Family.Value,
+    family: Platform.Family,
     name: String,
     os: String = "",
     arch: String = "x64",
--- a/src/Pure/Tools/dump.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/dump.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -152,12 +152,6 @@
       val session_graph = deps.sessions_structure.build_graph
       val all_sessions = session_graph.topological_order
 
-      val afp_sessions =
-        (for (name <- all_sessions if session_info(name).is_afp) yield name).toSet
-
-      val afp_bulky_sessions =
-        (for (name <- all_sessions if session_info(name).is_afp_bulky) yield name).toList
-
       val base_sessions =
         session_graph.all_preds_rev(List(logic).filter(session_graph.defined))
 
@@ -187,25 +181,12 @@
       val main =
         make_session(
           session_graph.topological_order.filterNot(name =>
-            afp_sessions.contains(name) ||
             base_sessions.contains(name) ||
             proof_sessions.contains(name)))
 
       val proofs = make_session(proof_sessions, session_logic = PURE, record_proofs = true)
 
-      val afp =
-        if (afp_sessions.isEmpty) Nil
-        else {
-          val (part1, part2) = {
-            val graph = session_graph.restrict(afp_sessions -- afp_bulky_sessions)
-            val force_partition1 = AFP.force_partition1.filter(graph.defined)
-            val force_part1 = graph.all_preds(graph.all_succs(force_partition1)).toSet
-            graph.keys.partition(a => force_part1(a) || graph.is_isolated(a))
-          }
-          List(part1, part2, afp_bulky_sessions).flatMap(make_session(_))
-        }
-
-      proofs ::: base ::: main ::: afp
+      proofs ::: base ::: main
     }
 
 
--- a/src/Pure/Tools/phabricator.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/phabricator.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -841,7 +841,7 @@
     /* repository information */
 
     sealed case class Repository(
-      vcs: VCS.Value,
+      vcs: VCS,
       id: Long,
       phid: String,
       name: String,
@@ -853,12 +853,11 @@
       def is_hg: Boolean = vcs == VCS.hg
     }
 
-    object VCS extends Enumeration {
-      val hg, git, svn = Value
-      def read(s: String): Value =
-        try { withName(s) }
-        catch { case _: java.util.NoSuchElementException => error("Unknown vcs type " + quote(s)) }
-    }
+    enum VCS { case hg, git, svn }
+
+    def read_vcs(s: String): VCS =
+      try { VCS.valueOf(s) }
+      catch { case _: IllegalArgumentException => error("Unknown vcs type " + quote(s)) }
 
     def edits(typ: String, value: JSON.T): List[JSON.Object.T] =
       List(JSON.Object("type" -> typ, "value" -> value))
@@ -999,7 +998,7 @@
                 importing <- JSON.bool(fields, "isImporting")
               }
               yield {
-                val vcs = API.VCS.read(vcs_name)
+                val vcs = API.read_vcs(vcs_name)
                 val url_path =
                   if (short_name.isEmpty) "/diffusion/" + id else "/source/" + short_name
                 val ssh_url =
@@ -1024,7 +1023,7 @@
       short_name: String = "",  // unique name
       description: String = "",
       public: Boolean = false,
-      vcs: API.VCS.Value = API.VCS.hg
+      vcs: API.VCS = API.VCS.hg
     ): API.Repository = {
       require(name.nonEmpty, "bad repository name")
 
--- a/src/Pure/Tools/prismjs.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/prismjs.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -27,7 +27,7 @@
     val components_json = JSON.parse(File.read(components_path))
     JSON.value(components_json, "languages") match {
       case Some(JSON.Object(langs)) =>
-        (for ((name, JSON.Object(info)) <- langs.iterator if name != "meta") yield {
+        (for (case (name, JSON.Object(info)) <- langs.iterator if name != "meta") yield {
           val alias =
             JSON.Value.List.unapply(info.getOrElse("alias", Nil), JSON.Value.String.unapply)
               .getOrElse(Nil)
--- a/src/Pure/Tools/profiling_report.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/profiling_report.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -30,7 +30,7 @@
               thy <- used_theories.iterator
               if theories.isEmpty || theories.contains(thy)
               snapshot <- Build.read_theory(session_context.theory(thy)).iterator
-              (Protocol.ML_Profiling(report), _) <- snapshot.messages.iterator
+              case (Protocol.ML_Profiling(report), _) <- snapshot.messages.iterator
             } yield if (clean_name) report.clean_name else report).toList
 
           for (report <- ML_Profiling.account(reports)) progress.echo(report.print)
--- a/src/Pure/Tools/server.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/server.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -91,29 +91,29 @@
       case _ => JSON.Object.empty
     }
 
-  object Reply extends Enumeration {
-    val OK, ERROR, FINISHED, FAILED, NOTE = Value
-
+  object Reply {
     def message(msg: String, kind: String = ""): JSON.Object.T =
       JSON.Object(Markup.KIND -> proper_string(kind).getOrElse(Markup.WRITELN), "message" -> msg)
 
     def error_message(msg: String): JSON.Object.T =
       message(msg, kind = Markup.ERROR)
 
-    def unapply(msg: String): Option[(Reply.Value, Any)] = {
+    def unapply(msg: String): Option[(Reply, Any)] = {
       if (msg == "") None
       else {
         val (name, argument) = Argument.split(msg)
         for {
           reply <-
-            try { Some(withName(name)) }
-            catch { case _: NoSuchElementException => None }
+            try { Some(Reply.valueOf(name)) }
+            catch { case _: IllegalArgumentException => None }
           arg <- Argument.unapply(argument)
         } yield (reply, arg)
       }
     }
   }
 
+  enum Reply { case OK, ERROR, FINISHED, FAILED, NOTE }
+
 
   /* handler: port, password, thread */
 
@@ -194,7 +194,7 @@
     def write_byte_message(chunks: List[Bytes]): Unit =
       out_lock.synchronized { Byte_Message.write_message(out, chunks) }
 
-    def reply(r: Reply.Value, arg: Any): Unit = {
+    def reply(r: Reply, arg: Any): Unit = {
       val argument = Argument.print(arg)
       write_line_message(if (argument == "") r.toString else r.toString + " " + argument)
     }
@@ -216,7 +216,7 @@
 
     def command_list: List[String] = command_table.keys.toList.sorted
 
-    def reply(r: Reply.Value, arg: Any): Unit = connection.reply(r, arg)
+    def reply(r: Reply, arg: Any): Unit = connection.reply(r, arg)
     def notify(arg: Any): Unit = connection.notify(arg)
     def message(kind: String, msg: String, more: JSON.Object.Entry*): Unit =
       notify(Reply.message(msg, kind = kind) ++ more)
--- a/src/Pure/Tools/simplifier_trace.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/simplifier_trace.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -208,7 +208,7 @@
                       def purge(queue: Vector[Long]): Unit =
                         queue match {
                           case s +: rest =>
-                            for (Item(Markup.SIMP_TRACE_STEP, data) <- results.get(s))
+                            for (case Item(Markup.SIMP_TRACE_STEP, data) <- results.get(s))
                               memory -= Index.of_data(data)
                             val children = memory_children.getOrElse(s, Set.empty)
                             memory_children -= s
--- a/src/Pure/Tools/spell_checker.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/spell_checker.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -129,7 +129,7 @@
     val permanent_updates =
       if (dictionary.user_path.is_file)
         for {
-          Spell_Checker.Decl(word, include) <- split_lines(File.read(dictionary.user_path))
+          case Spell_Checker.Decl(word, include) <- split_lines(File.read(dictionary.user_path))
         } yield (word, Spell_Checker.Update(include, true))
       else Nil
 
@@ -201,7 +201,7 @@
 
   def check(word: String): Boolean =
     word match {
-      case Word.Case(c) if c != Word.Lowercase =>
+      case Word.Case(c) if c != Word.Case.lowercase =>
         contains(word) || contains(Word.lowercase(word))
       case _ =>
         contains(word)
@@ -231,7 +231,7 @@
         }
       val result =
         word_case match {
-          case Some(c) if c != Word.Lowercase =>
+          case Some(c) if c != Word.Case.lowercase =>
             suggestions(word) orElse suggestions(Word.lowercase(word))
           case _ =>
             suggestions(word)
--- a/src/Pure/Tools/task_statistics.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Pure/Tools/task_statistics.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -29,7 +29,7 @@
 
   def chart(bins: Int = 100): JFreeChart = {
     val values = new Array[Double](task_statistics.length)
-    for ((Run(x), i) <- task_statistics.iterator.zipWithIndex)
+    for (case (Run(x), i) <- task_statistics.iterator.zipWithIndex)
       values(i) = java.lang.Math.log10((x max 1).toDouble / 1000000)
 
     val data = new HistogramDataset
--- a/src/Tools/Code/code_scala.ML	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/Code/code_scala.ML	Sun Sep 03 19:28:59 2023 +0200
@@ -70,7 +70,7 @@
       | print_tupled_typ tyvars (tys, ty) =
           concat [enum "," "(" ")" (map (print_typ tyvars NOBR) tys),
             str "=>", print_typ tyvars NOBR ty];
-    fun constraint p1 p2 = Pretty.block [p1, str ":", Pretty.brk 1, p2];
+    fun constraint p1 p2 = Pretty.block [p1, str " : ", p2];
     fun print_var vars NONE = str "_"
       | print_var vars (SOME v) = (str o lookup_var vars) v;
     fun applify_dict tyvars (Dict (_, d)) = applify_plain_dict tyvars d
--- a/src/Tools/Graphview/layout.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/Graphview/layout.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -400,13 +400,13 @@
     output_graph.get_node(Layout.Node(node))
 
   def nodes_iterator: Iterator[Layout.Info] =
-    for ((_: Layout.Node, (info, _)) <- output_graph.iterator) yield info
+    for (case (_: Layout.Node, (info, _)) <- output_graph.iterator) yield info
 
 
   /* dummies */
 
   def dummies_iterator: Iterator[Layout.Info] =
-    for ((_: Layout.Dummy, (info, _)) <- output_graph.iterator) yield info
+    for (case (_: Layout.Dummy, (info, _)) <- output_graph.iterator) yield info
 
   def dummies_iterator(edge: Graph_Display.Edge): Iterator[Layout.Info] =
     new Iterator[Layout.Info] {
--- a/src/Tools/Graphview/model.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/Graphview/model.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -20,12 +20,12 @@
   def apply(): List[Mutator.Info] = _mutators
   def apply(mutators: List[Mutator.Info]): Unit = {
     _mutators = mutators
-    events.event(Mutator_Event.New_List(mutators))
+    events.event(Mutator_Event.Message.New_List(mutators))
   }
 
   def add(mutator: Mutator.Info): Unit = {
     _mutators = _mutators ::: List(mutator)
-    events.event(Mutator_Event.Add(mutator))
+    events.event(Mutator_Event.Message.Add(mutator))
   }
 }
 
--- a/src/Tools/Graphview/mutator_dialog.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/Graphview/mutator_dialog.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -37,8 +37,8 @@
 
   container.events +=
     {
-      case Mutator_Event.Add(m) => add_panel(new Mutator_Panel(m))
-      case Mutator_Event.New_List(ms) => panels = get_panels(ms)
+      case Mutator_Event.Message.Add(m) => add_panel(new Mutator_Panel(m))
+      case Mutator_Event.Message.New_List(ms) => panels = get_panels(ms)
     }
 
   override def open(): Unit = {
--- a/src/Tools/Graphview/mutator_event.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/Graphview/mutator_event.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -12,11 +12,12 @@
 
 
 object Mutator_Event {
-  sealed abstract class Message
-  case class Add(m: Mutator.Info) extends Message
-  case class New_List(m: List[Mutator.Info]) extends Message
+  enum Message {
+    case Add(m: Mutator.Info) extends Message
+    case New_List(m: List[Mutator.Info]) extends Message
+  }
 
-  type Receiver = PartialFunction[Message, Unit]
+  type Receiver = Message => Unit
 
   class Bus {
     private val receivers = Synchronized[List[Receiver]](Nil)
@@ -25,4 +26,4 @@
     def -= (r: Receiver): Unit = receivers.change(Library.remove(r))
     def event(x: Message): Unit = receivers.value.reverse.foreach(r => r(x))
   }
-}
\ No newline at end of file
+}
--- a/src/Tools/VSCode/src/component_vscodium.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/VSCode/src/component_vscodium.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -58,7 +58,7 @@
   /* platform info */
 
   sealed case class Platform_Info(
-    platform: Platform.Family.Value,
+    platform: Platform.Family,
     download_template: String,
     build_name: String,
     env: List[String]
@@ -233,7 +233,7 @@
       .text.replaceAll("=", "")
   }
 
-  private val platform_infos: Map[Platform.Family.Value, Platform_Info] =
+  private val platform_infos: Map[Platform.Family, Platform_Info] =
     Iterator(
       Platform_Info(Platform.Family.linux, "linux-x64-{VERSION}.tar.gz", "VSCode-linux-x64",
         List("OS_NAME=linux", "SKIP_LINUX_PACKAGES=True")),
@@ -250,7 +250,7 @@
           "SHOULD_BUILD_MSI_NOUP=no")))
       .map(info => info.platform -> info).toMap
 
-  def the_platform_info(platform: Platform.Family.Value): Platform_Info =
+  def the_platform_info(platform: Platform.Family): Platform_Info =
     platform_infos.getOrElse(platform, error("No platform info for " + quote(platform.toString)))
 
   def linux_platform_info: Platform_Info =
@@ -259,7 +259,7 @@
 
   /* check system */
 
-  def check_system(platforms: List[Platform.Family.Value]): Unit = {
+  def check_system(platforms: List[Platform.Family]): Unit = {
     if (Platform.family != Platform.Family.linux) error("Not a Linux/x86_64 system")
 
     Isabelle_System.require_command("git")
@@ -301,11 +301,11 @@
 
   /* build vscodium */
 
-  def default_platforms: List[Platform.Family.Value] = Platform.Family.list
+  def default_platforms: List[Platform.Family] = Platform.Family.list
 
   def component_vscodium(
     target_dir: Path = Path.current,
-    platforms: List[Platform.Family.Value] = default_platforms,
+    platforms: List[Platform.Family] = default_platforms,
     progress: Progress = new Progress
   ): Unit = {
     check_system(platforms)
--- a/src/Tools/jEdit/jedit_base/plugin.props	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/jedit_base/plugin.props	Sun Sep 03 19:28:59 2023 +0200
@@ -13,5 +13,5 @@
 plugin.isabelle.jedit_base.Plugin.usePluginHome=false
 
 #dependencies
-plugin.isabelle.jedit_base.Plugin.depend.0=jdk 11
+plugin.isabelle.jedit_base.Plugin.depend.0=jdk 17
 plugin.isabelle.jedit_base.Plugin.depend.1=jedit 05.05.00.00
--- a/src/Tools/jEdit/jedit_main/plugin.props	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/jedit_main/plugin.props	Sun Sep 03 19:28:59 2023 +0200
@@ -13,7 +13,7 @@
 plugin.isabelle.jedit_main.Plugin.usePluginHome=false
 
 #dependencies
-plugin.isabelle.jedit_main.Plugin.depend.0=jdk 11
+plugin.isabelle.jedit_main.Plugin.depend.0=jdk 17
 plugin.isabelle.jedit_main.Plugin.depend.1=jedit 05.06.00.00
 plugin.isabelle.jedit_main.Plugin.depend.2=plugin console.ConsolePlugin 5.1.4
 plugin.isabelle.jedit_main.Plugin.depend.3=plugin errorlist.ErrorListPlugin 2.4.0
--- a/src/Tools/jEdit/src/document_dockable.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/src/document_dockable.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -278,7 +278,7 @@
 
   private val auto_build_button =
     new JEdit_Options.Bool_GUI(document_auto, "Auto") {
-      tooltip = Word.capitalize(document_auto.description)
+      tooltip = Word.capitalized(document_auto.description)
       override def clicked(state: Boolean): Unit = {
         super.clicked(state)
 
--- a/src/Tools/jEdit/src/document_model.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/src/document_model.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -255,12 +255,12 @@
             val open_nodes =
               (for ((_, model) <- st.buffer_models.iterator) yield model.node_name).toList
             val touched_nodes = model_edits.map(_._1)
-            val pending_nodes = for ((node_name, None) <- purged) yield node_name
+            val pending_nodes = for (case (node_name, None) <- purged) yield node_name
             (open_nodes ::: touched_nodes ::: pending_nodes).map((_, Position.none))
           }
           val retain = PIDE.resources.dependencies(imports).theories.toSet
 
-          for ((node_name, Some(edits)) <- purged if !retain(node_name); edit <- edits)
+          for (case (node_name, Some(edits)) <- purged if !retain(node_name); edit <- edits)
             yield edit
         }
         else Nil
--- a/src/Tools/jEdit/src/jedit_sessions.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/src/jedit_sessions.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -91,7 +91,7 @@
   extends GUI.Selector[String](batches: _*) with JEdit_Options.Entry {
     name = option_name
     tooltip =
-      if (standalone) Word.capitalize(options.value.description(option_name))
+      if (standalone) Word.capitalized(options.value.description(option_name))
       else options.value.check_name(option_name).print_default
 
     override val title: String =
--- a/src/Tools/jEdit/src/query_dockable.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/src/query_dockable.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -55,14 +55,14 @@
 
   def consume_status(
     process_indicator: Process_Indicator,
-    status: Query_Operation.Status.Value
+    status: Query_Operation.Status
   ): Unit = {
     status match {
-      case Query_Operation.Status.WAITING =>
+      case Query_Operation.Status.waiting =>
         process_indicator.update("Waiting for evaluation of context ...", 5)
-      case Query_Operation.Status.RUNNING =>
+      case Query_Operation.Status.running =>
         process_indicator.update("Running find operation ...", 15)
-      case Query_Operation.Status.FINISHED =>
+      case Query_Operation.Status.finished =>
         process_indicator.update(null, 0)
     }
   }
--- a/src/Tools/jEdit/src/sledgehammer_dockable.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/src/sledgehammer_dockable.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -34,13 +34,13 @@
 
   private val process_indicator = new Process_Indicator
 
-  private def consume_status(status: Query_Operation.Status.Value): Unit = {
+  private def consume_status(status: Query_Operation.Status): Unit = {
     status match {
-      case Query_Operation.Status.WAITING =>
+      case Query_Operation.Status.waiting =>
         process_indicator.update("Waiting for evaluation of context ...", 5)
-      case Query_Operation.Status.RUNNING =>
+      case Query_Operation.Status.running =>
         process_indicator.update("Sledgehammering ...", 15)
-      case Query_Operation.Status.FINISHED =>
+      case Query_Operation.Status.finished =>
         process_indicator.update(null, 0)
     }
   }
--- a/src/Tools/jEdit/src/symbols_dockable.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/src/symbols_dockable.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -169,7 +169,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_capitalized))
   }
   set_content(group_tabs)
 
--- a/src/Tools/jEdit/src/theories_status.scala	Sun Sep 03 18:09:00 2023 +0200
+++ b/src/Tools/jEdit/src/theories_status.scala	Sun Sep 03 19:28:59 2023 +0200
@@ -30,9 +30,7 @@
   private def is_loaded_theory(name: Document.Node.Name): Boolean =
     PIDE.resources.session_base.loaded_theory(name)
 
-  private def overall_node_status(
-    name: Document.Node.Name
-  ): Document_Status.Overall_Node_Status.Value = {
+  private def overall_node_status(name: Document.Node.Name): Document_Status.Overall_Node_Status = {
     if (is_loaded_theory(name)) Document_Status.Overall_Node_Status.ok
     else nodes_status.overall_node_status(name)
   }