merged
authorpaulson
Thu, 03 Aug 2023 19:10:43 +0200
changeset 78476 032a4344903e
parent 78473 ba2afdd29e1d (current diff)
parent 78475 a5f6d2fc1b1f (diff)
child 78477 37abfe400ae6
merged
--- a/src/HOL/Analysis/Bochner_Integration.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Bochner_Integration.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -99,22 +99,19 @@
           unfolding A_def by simp
       qed
       ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)"
-        by (auto simp add: F_def L_def) }
+        by (auto simp: F_def L_def) }
     note * = this
 
     fix x assume "x \<in> space M"
     show "(\<lambda>i. F i x) \<longlonglongrightarrow> f x"
-    proof cases
-      assume "f x = z"
+    proof (cases "f x = z")
+      case True
       then have "\<And>i n. x \<notin> A i n"
         unfolding A_def by auto
-      then have "\<And>i. F i x = z"
-        by (auto simp: F_def)
       then show ?thesis
-        using \<open>f x = z\<close> by auto
+        by (metis B_imp_A F_def LIMSEQ_def True dist_self)
     next
-      assume "f x \<noteq> z"
-
+      case False
       show ?thesis
       proof (rule tendstoI)
         fix e :: real assume "0 < e"
@@ -148,10 +145,7 @@
       finally (xtrans) have "dist (e m) z \<le> 2 * dist (f x) z" . }
     then show "dist (F i x) z \<le> 2 * dist (f x) z"
       unfolding F_def
-      apply auto
-      apply (rule LeastI2)
-      apply auto
-      done
+      by (smt (verit, ccfv_threshold) LeastI2 B_imp_A dist_eq_0_iff zero_le_dist)
   qed
 qed
 
@@ -184,10 +178,7 @@
   show "P u"
   proof (rule seq)
     show U': "U' i \<in> borel_measurable M" "\<And>x. 0 \<le> U' i x" for i
-      using U by (auto
-          intro: borel_measurable_simple_function
-          intro!: borel_measurable_enn2real borel_measurable_times
-          simp: U'_def zero_le_mult_iff)
+      using U'_sf by (auto simp: U'_def borel_measurable_simple_function)
     show "incseq U'"
       using U(2,3)
       by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono)
@@ -238,7 +229,7 @@
 lemma scaleR_cong_right:
   fixes x :: "'a :: real_vector"
   shows "(x \<noteq> 0 \<Longrightarrow> r = p) \<Longrightarrow> r *\<^sub>R x = p *\<^sub>R x"
-  by (cases "x = 0") auto
+  by auto
 
 inductive simple_bochner_integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> bool" for M f where
   "simple_function M f \<Longrightarrow> emeasure M {y\<in>space M. f y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow>
@@ -309,11 +300,7 @@
   shows "simple_bochner_integrable M f"
 proof
   have "emeasure M {y \<in> space M. ennreal (norm (f y)) \<noteq> 0} \<noteq> \<infinity>"
-  proof (rule simple_function_finite_support)
-    show "simple_function M (\<lambda>x. ennreal (norm (f x)))"
-      using f by (rule simple_function_compose1)
-    show "(\<integral>\<^sup>+ y. ennreal (norm (f y)) \<partial>M) < \<infinity>" by fact
-  qed simp
+    using simple_function_finite_support simple_function_compose1 f fin by force
   then show "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" by simp
 qed fact
 
@@ -453,18 +440,17 @@
 lemma simple_bochner_integral_nonneg[simp]:
   fixes f :: "'a \<Rightarrow> real"
   shows "(\<And>x. 0 \<le> f x) \<Longrightarrow> 0 \<le> simple_bochner_integral M f"
-  by (force simp add: simple_bochner_integral_def intro: sum_nonneg)
+  by (force simp: simple_bochner_integral_def intro: sum_nonneg)
 
 lemma simple_bochner_integral_eq_nn_integral:
   assumes f: "simple_bochner_integrable M f" "\<And>x. 0 \<le> f x"
   shows "simple_bochner_integral M f = (\<integral>\<^sup>+x. f x \<partial>M)"
 proof -
-  { fix x y z have "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ennreal x * y = ennreal x * z"
-      by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) }
-  note ennreal_cong_mult = this
+  have ennreal_cong_mult: "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ennreal x * y = ennreal x * z" for x y z
+      by fastforce
 
   have [measurable]: "f \<in> borel_measurable M"
-    using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
+    by (meson borel_measurable_simple_function f(1) simple_bochner_integrable.cases)
 
   { fix y assume y: "y \<in> space M" "f y \<noteq> 0"
     have "ennreal (measure M {x \<in> space M. f x = f y}) = emeasure M {x \<in> space M. f x = f y}"
@@ -487,8 +473,7 @@
              simp flip: sum_ennreal)
   also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
     using f
-    by (intro nn_integral_eq_simple_integral[symmetric])
-       (auto simp: simple_function_compose1 simple_bochner_integrable.simps)
+    by (metis nn_integral_eq_simple_integral simple_bochner_integrable.cases simple_function_compose1)
   finally show ?thesis .
 qed
 
@@ -512,9 +497,13 @@
     using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t
     by (auto intro!: simple_bochner_integral_eq_nn_integral)
   also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)"
-    by (auto intro!: nn_integral_mono simp flip: ennreal_plus)
-       (metis (erased, opaque_lifting) add_diff_cancel_left add_diff_eq diff_add_eq order_trans
-              norm_minus_commute norm_triangle_ineq4 order_refl)
+  proof -
+    have "\<And>x. x \<in> space M \<Longrightarrow>
+         norm (s x - t x) \<le> norm (f x - s x) + norm (f x - t x)"
+      by (metis dual_order.refl norm_diff_triangle_le norm_minus_commute)
+    then show ?thesis
+      by (metis (mono_tags, lifting) ennreal_leI ennreal_plus nn_integral_mono norm_ge_zero)
+  qed
   also have "\<dots> = ?S + ?T"
    by (rule nn_integral_add) auto
   finally show ?thesis .
@@ -555,7 +544,7 @@
   by (rule has_bochner_integral.intros[where s="\<lambda>_. f"])
      (auto intro: borel_measurable_simple_function
            elim: simple_bochner_integrable.cases
-           simp: zero_ennreal_def[symmetric])
+           simp flip: zero_ennreal_def)
 
 lemma has_bochner_integral_real_indicator:
   assumes [measurable]: "A \<in> sets M" and A: "emeasure M A < \<infinity>"
@@ -665,21 +654,21 @@
 
 lemma has_bochner_integral_scaleR_left[intro]:
   "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x *\<^sub>R c) (x *\<^sub>R c)"
-  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left])
+  by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left])
 
 lemma has_bochner_integral_scaleR_right[intro]:
   "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c *\<^sub>R f x) (c *\<^sub>R x)"
-  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right])
+  by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right])
 
 lemma has_bochner_integral_mult_left[intro]:
   fixes c :: "_::{real_normed_algebra,second_countable_topology}"
   shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x * c) (x * c)"
-  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left])
+  by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left])
 
 lemma has_bochner_integral_mult_right[intro]:
   fixes c :: "_::{real_normed_algebra,second_countable_topology}"
   shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c * f x) (c * x)"
-  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right])
+  by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right])
 
 lemmas has_bochner_integral_divide =
   has_bochner_integral_bounded_linear[OF bounded_linear_divide]
@@ -691,11 +680,11 @@
 
 lemma has_bochner_integral_inner_left[intro]:
   "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x \<bullet> c) (x \<bullet> c)"
-  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left])
+  by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left])
 
 lemma has_bochner_integral_inner_right[intro]:
   "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c \<bullet> f x) (c \<bullet> x)"
-  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right])
+  by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right])
 
 lemmas has_bochner_integral_minus =
   has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]
@@ -857,16 +846,10 @@
 qed (auto intro: g)
 
 lemma has_bochner_integral_eq_AE:
-  assumes f: "has_bochner_integral M f x"
-    and g: "has_bochner_integral M g y"
-    and ae: "AE x in M. f x = g x"
+  assumes "has_bochner_integral M f x" and "has_bochner_integral M g y"
+    and "AE x in M. f x = g x"
   shows "x = y"
-proof -
-  from assms have "has_bochner_integral M g x"
-    by (auto intro: has_bochner_integralI_AE)
-  from this g show "x = y"
-    by (rule has_bochner_integral_eq)
-qed
+  by (meson assms has_bochner_integral.simps has_bochner_integral_cong_AE has_bochner_integral_eq)
 
 lemma simple_bochner_integrable_restrict_space:
   fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
@@ -1082,12 +1065,7 @@
   next
     assume not: "\<not> integrable M f"
     moreover have "\<not> integrable M (\<lambda>x. T (f x))"
-    proof
-      assume "integrable M (\<lambda>x. T (f x))"
-      from integrable_bounded_linear[OF T' this] not *[OF **]
-      show False
-        by auto
-    qed
+      by (metis (full_types) "*" "**" T' integrable_bounded_linear integrable_cong not)
     ultimately show ?thesis
       using T by (simp add: not_integrable_integral_eq linear_simps)
   qed
@@ -1229,7 +1207,7 @@
       show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
         by (simp add: tendsto_norm_zero_iff flip: ennreal_0)
     qed
-  qed (insert bnd w_nonneg, auto)
+  qed (use bnd w_nonneg in auto)
   then show ?thesis by simp
 qed
 
@@ -1342,14 +1320,8 @@
   shows "integrable M f \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. norm (g x) \<le> norm (f x)) \<Longrightarrow>
     integrable M g"
   unfolding integrable_iff_bounded
-proof safe
-  assume "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  assume "AE x in M. norm (g x) \<le> norm (f x)"
-  then have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
-    by  (intro nn_integral_mono_AE) auto
-  also assume "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
-  finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" .
-qed
+  by (smt (verit) AE_cong ennreal_le_iff nn_integral_mono_AE norm_ge_zero order_less_subst2 linorder_not_le order_less_le_trans)
+
 
 lemma integrable_mult_indicator:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -1419,7 +1391,7 @@
 proof cases
   assume *: "A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
   have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \<inter> space M))"
-    by (intro integral_cong) (auto split: split_indicator)
+    by (metis (no_types, lifting) Int_iff indicator_simps integral_cong)
   also have "\<dots> = measure M (A \<inter> space M)"
     using * by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator) auto
   finally show ?thesis .
@@ -1434,6 +1406,26 @@
   finally show ?thesis .
 qed
 
+lemma integrable_discrete_difference_aux:
+  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+  assumes f: "integrable M f" and X: "countable X"
+  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
+  assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+  assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
+  shows "integrable M g"
+  unfolding integrable_iff_bounded
+proof
+  have fmeas: "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
+    using f integrable_iff_bounded by auto
+  then show "g \<in> borel_measurable M"
+    using measurable_discrete_difference[where X=X]
+    by (smt (verit) UNIV_I X eq sets space_borel)
+  have "AE x in M. x \<notin> X"
+    using AE_discrete_difference X null sets by blast
+  with fmeas show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>"
+    by (metis (mono_tags, lifting) AE_I2 AE_mp eq nn_integral_cong_AE)
+qed
+
 lemma integrable_discrete_difference:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes X: "countable X"
@@ -1441,22 +1433,7 @@
   assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
   assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
   shows "integrable M f \<longleftrightarrow> integrable M g"
-  unfolding integrable_iff_bounded
-proof (rule conj_cong)
-  { assume "f \<in> borel_measurable M" then have "g \<in> borel_measurable M"
-      by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) }
-  moreover
-  { assume "g \<in> borel_measurable M" then have "f \<in> borel_measurable M"
-      by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) }
-  ultimately show "f \<in> borel_measurable M \<longleftrightarrow> g \<in> borel_measurable M" ..
-next
-  have "AE x in M. x \<notin> X"
-    by (rule AE_discrete_difference) fact+
-  then have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. norm (g x) \<partial>M)"
-    by (intro nn_integral_cong_AE) (auto simp: eq)
-  then show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity> \<longleftrightarrow> (\<integral>\<^sup>+ x. norm (g x) \<partial>M) < \<infinity>"
-    by simp
-qed
+  by (metis X eq integrable_discrete_difference_aux null sets)
 
 lemma integral_discrete_difference:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -1474,7 +1451,6 @@
   proof (rule integral_cong_AE)
     show "f \<in> borel_measurable M" "g \<in> borel_measurable M"
       using f eq by (auto intro: borel_measurable_integrable)
-
     have "AE x in M. x \<notin> X"
       by (rule AE_discrete_difference) fact+
     with AE_space show "AE x in M. f x = g x"
@@ -1484,14 +1460,12 @@
 
 lemma has_bochner_integral_discrete_difference:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
-  assumes X: "countable X"
-  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
-  assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
-  assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
+  assumes "countable X"
+  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
+  assumes "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
   shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral M g x"
-  using integrable_discrete_difference[of X M f g, OF assms]
-  using integral_discrete_difference[of X M f g, OF assms]
-  by (metis has_bochner_integral_iff)
+  by (metis assms has_bochner_integral_iff integrable_discrete_difference integral_discrete_difference)
 
 lemma
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real"
@@ -1547,7 +1521,7 @@
       finally show "?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)" .
     qed
     show "(\<lambda>n. \<integral>\<^sup>+x. norm (s n x - f x) \<partial>M) \<longlonglongrightarrow> ennreal 0"
-      unfolding ennreal_0
+      unfolding ennreal_0 
       apply (subst norm_minus_commute)
     proof (rule nn_integral_dominated_convergence_norm[where w=w])
       show "\<And>n. s n \<in> borel_measurable M"
@@ -1687,7 +1661,7 @@
         by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure)
     next
       case (mult f c) then show ?case
-        by (auto simp add: nn_integral_cmult ennreal_mult integral_nonneg_AE)
+        by (auto simp: nn_integral_cmult ennreal_mult integral_nonneg_AE)
     next
       case (add g f)
       then have "integrable M f" "integrable M g"
@@ -1723,14 +1697,9 @@
 qed
 
 lemma nn_integral_eq_integrable:
-  assumes f: "f \<in> M \<rightarrow>\<^sub>M borel" "AE x in M. 0 \<le> f x" and "0 \<le> x"
+  assumes "f \<in> M \<rightarrow>\<^sub>M borel" "AE x in M. 0 \<le> f x" and "0 \<le> x"
   shows "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x \<longleftrightarrow> (integrable M f \<and> integral\<^sup>L M f = x)"
-proof (safe intro!: nn_integral_eq_integral assms)
-  assume *: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
-  with integrableI_nn_integral_finite[OF f this] nn_integral_eq_integral[of M f, OF _ f(2)]
-  show "integrable M f" "integral\<^sup>L M f = x"
-    by (simp_all add: * assms integral_nonneg_AE)
-qed
+  by (metis assms ennreal_inj integrableI_nn_integral_finite integral_nonneg_AE nn_integral_eq_integral)
 
 lemma
   fixes f :: "_ \<Rightarrow> _ \<Rightarrow> 'a :: {banach, second_countable_topology}"
@@ -1744,12 +1713,11 @@
 proof -
   have 1: "integrable M (\<lambda>x. \<Sum>i. norm (f i x))"
   proof (rule integrableI_bounded)
-    have "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ennreal (norm (f i x))) \<partial>M)"
-      apply (intro nn_integral_cong_AE)
-      using summable
-      apply eventually_elim
-      apply (simp add: suminf_nonneg ennreal_suminf_neq_top)
-      done
+    have "\<And>x. summable (\<lambda>i. norm (f i x)) \<Longrightarrow>
+         ennreal (norm (\<Sum>i. norm (f i x))) = (\<Sum>i. ennreal (norm (f i x)))"
+      by (simp add: suminf_nonneg ennreal_suminf_neq_top)
+  then have "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ennreal (norm (f i x))) \<partial>M)"
+      using local.summable by (intro nn_integral_cong_AE) auto
     also have "\<dots> = (\<Sum>i. \<integral>\<^sup>+ x. norm (f i x) \<partial>M)"
       by (intro nn_integral_suminf) auto
     also have "\<dots> = (\<Sum>i. ennreal (\<integral>x. norm (f i x) \<partial>M))"
@@ -1802,63 +1770,28 @@
 using integral_norm_bound[of M f] by auto
 
 lemma integral_eq_nn_integral:
-  assumes [measurable]: "f \<in> borel_measurable M"
-  assumes nonneg: "AE x in M. 0 \<le> f x"
+  assumes "f \<in> borel_measurable M"
+  assumes "AE x in M. 0 \<le> f x"
   shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)"
-proof cases
-  assume *: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = \<infinity>"
-  also have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
-    using nonneg by (intro nn_integral_cong_AE) auto
-  finally have "\<not> integrable M f"
-    by (auto simp: integrable_iff_bounded)
-  then show ?thesis
-    by (simp add: * not_integrable_integral_eq)
-next
-  assume "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>"
-  then have "integrable M f"
-    by (cases "\<integral>\<^sup>+ x. ennreal (f x) \<partial>M" rule: ennreal_cases)
-       (auto intro!: integrableI_nn_integral_finite assms)
-  from nn_integral_eq_integral[OF this] nonneg show ?thesis
-    by (simp add: integral_nonneg_AE)
-qed
+  by (metis assms enn2real_ennreal enn2real_top infinity_ennreal_def integrableI_nonneg integral_nonneg_AE
+      less_top nn_integral_eq_integral not_integrable_integral_eq)
 
 lemma enn2real_nn_integral_eq_integral:
-  assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \<le> g x"
-    and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < top"
-    and [measurable]: "g \<in> M \<rightarrow>\<^sub>M borel"
+  assumes "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \<le> g x"
+    and "g \<in> M \<rightarrow>\<^sub>M borel"
   shows "enn2real (\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>x. g x \<partial>M)"
-proof -
-  have "ennreal (enn2real (\<integral>\<^sup>+x. f x \<partial>M)) = (\<integral>\<^sup>+x. f x \<partial>M)"
-    using fin by (intro ennreal_enn2real) auto
-  also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M)"
-    using eq by (rule nn_integral_cong_AE)
-  also have "\<dots> = (\<integral>x. g x \<partial>M)"
-  proof (rule nn_integral_eq_integral)
-    show "integrable M g"
-    proof (rule integrableI_bounded)
-      have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)"
-        using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2)
-      also note fin
-      finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>"
-        by simp
-    qed simp
-  qed fact
-  finally show ?thesis
-    using nn by (simp add: integral_nonneg_AE)
-qed
+  by (metis assms integral_eq_nn_integral nn nn_integral_cong_AE)
 
 lemma has_bochner_integral_nn_integral:
   assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> x"
   assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
   shows "has_bochner_integral M f x"
-  unfolding has_bochner_integral_iff
-  using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite)
+  by (metis assms has_bochner_integral_iff nn_integral_eq_integrable)
 
 lemma integrableI_simple_bochner_integrable:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   shows "simple_bochner_integrable M f \<Longrightarrow> integrable M f"
-  by (intro integrableI_sequence[where s="\<lambda>_. f"] borel_measurable_simple_function)
-     (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps)
+  using has_bochner_integral_simple_bochner_integrable integrable.intros by blast
 
 proposition integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -1896,7 +1829,7 @@
   { fix i
     have "\<And>z. {y. s' i z = y \<and> y \<in> s' i ` space M \<and> y \<noteq> 0 \<and> z \<in> space M} =
         (if z \<in> space M \<and> s' i z \<noteq> 0 then {s' i z} else {})"
-      by (auto simp add: s'_def split: split_indicator)
+      by (auto simp: s'_def split: split_indicator)
     then have "\<And>z. s' i = (\<lambda>z. \<Sum>y\<in>s' i`space M - {0}. indicator {x\<in>space M. s' i x = y} z *\<^sub>R y)"
       using sf by (auto simp: fun_eq_iff simple_function_def s'_def) }
   note s'_eq = this
@@ -1947,7 +1880,7 @@
     by (simp add: nn_integral_0_iff_AE)
   with nonneg show "AE x in M. f x = 0"
     by auto
-qed (auto simp add: integral_eq_zero_AE)
+qed (auto simp: integral_eq_zero_AE)
 
 lemma integral_mono_AE:
   fixes f :: "'a \<Rightarrow> real"
@@ -1971,13 +1904,7 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{banach,second_countable_topology}"
   assumes "integrable M f" "integrable M g" "\<And>x. x \<in> space M \<Longrightarrow> norm(f x) \<le> g x"
   shows "norm (\<integral>x. f x \<partial>M) \<le> (\<integral>x. g x \<partial>M)"
-proof -
-  have "norm (\<integral>x. f x \<partial>M) \<le> (\<integral>x. norm (f x) \<partial>M)"
-    by (rule integral_norm_bound)
-  also have "... \<le> (\<integral>x. g x \<partial>M)"
-    using assms integrable_norm integral_mono by blast
-  finally show ?thesis .
-qed
+  by (smt (verit, best) assms integrable_norm integral_mono integral_norm_bound)
 
 lemma integral_abs_bound_integral:
   fixes f :: "'a::euclidean_space \<Rightarrow> real"
@@ -1993,15 +1920,7 @@
   fixes f::"_ \<Rightarrow> real"
   assumes "integrable M f" "AE x in M. g x \<le> f x" "AE x in M. 0 \<le> f x"
   shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)"
-proof (cases "integrable M g")
-  case True
-  show ?thesis by (rule integral_mono_AE, auto simp add: assms True)
-next
-  case False
-  then have "(\<integral>x. g x \<partial>M) = 0" by (simp add: not_integrable_integral_eq)
-  also have "... \<le> (\<integral>x. f x \<partial>M)" by (simp add: integral_nonneg_AE[OF assms(3)])
-  finally show ?thesis by simp
-qed
+  by (metis assms integral_mono_AE integral_nonneg_AE not_integrable_integral_eq)
 
 lemma integral_mono':
   fixes f::"_ \<Rightarrow> real"
@@ -2022,21 +1941,11 @@
     by (auto intro!: integrableI_bounded)
 qed
 
-lemma integrableI_real_bounded:
-  assumes f: "f \<in> borel_measurable M" and ae: "AE x in M. 0 \<le> f x" and fin: "integral\<^sup>N M f < \<infinity>"
-  shows "integrable M f"
-proof (rule integrableI_bounded)
-  have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = \<integral>\<^sup>+ x. ennreal (f x) \<partial>M"
-    using ae by (auto intro: nn_integral_cong_AE)
-  also note fin
-  finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
-qed fact
-
 lemma nn_integral_nonneg_infinite:
   fixes f::"'a \<Rightarrow> real"
   assumes "f \<in> borel_measurable M" "\<not> integrable M f" "AE x in M. f x \<ge> 0"
   shows "(\<integral>\<^sup>+x. f x \<partial>M) = \<infinity>"
-using assms integrableI_real_bounded less_top by auto
+  using assms integrableI_nonneg less_top by auto
 
 lemma integral_real_bounded:
   assumes "0 \<le> r" "integral\<^sup>N M f \<le> ennreal r"
@@ -2066,22 +1975,22 @@
   fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real"
   shows "\<lbrakk> finite I;  I \<noteq> {}; \<And>i. i \<in> I \<Longrightarrow> integrable M (f i) \<rbrakk>
    \<Longrightarrow> integrable M (\<lambda>x. MIN i\<in>I. f i x)"
-by (induct rule: finite_ne_induct) simp+
+  by (induct rule: finite_ne_induct) simp+
 
 lemma integrable_MAX:
   fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real"
   shows "\<lbrakk> finite I;  I \<noteq> {};  \<And>i. i \<in> I \<Longrightarrow> integrable M (f i) \<rbrakk>
   \<Longrightarrow> integrable M (\<lambda>x. MAX i\<in>I. f i x)"
-by (induct rule: finite_ne_induct) simp+
+  by (induct rule: finite_ne_induct) simp+
 
 theorem integral_Markov_inequality:
   assumes [measurable]: "integrable M u" and "AE x in M. 0 \<le> u x" "0 < (c::real)"
   shows "(emeasure M) {x\<in>space M. u x \<ge> c} \<le> (1/c) * (\<integral>x. u x \<partial>M)"
 proof -
   have "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>\<^sup>+ x. u x \<partial>M)"
-    by (rule nn_integral_mono_AE, auto simp add: \<open>c>0\<close> less_eq_real_def)
-  also have "... = (\<integral>x. u x \<partial>M)"
-    by (rule nn_integral_eq_integral, auto simp add: assms)
+    by (rule nn_integral_mono_AE, auto simp: \<open>c>0\<close> less_eq_real_def)
+  also have "\<dots> = (\<integral>x. u x \<partial>M)"
+    by (rule nn_integral_eq_integral, auto simp: assms)
   finally have *: "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>x. u x \<partial>M)"
     by simp
 
@@ -2089,9 +1998,9 @@
     using \<open>c>0\<close> by (auto simp: ennreal_mult'[symmetric])
   then have "emeasure M {x \<in> space M. u x \<ge> c} = emeasure M ({x \<in> space M. ennreal(1/c) * u x \<ge> 1})"
     by simp
-  also have "... \<le> ennreal(1/c) * (\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M)"
-    by (rule nn_integral_Markov_inequality) (auto simp add: assms)
-  also have "... \<le> ennreal(1/c) * (\<integral>x. u x \<partial>M)"
+  also have "\<dots> \<le> ennreal(1/c) * (\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M)"
+    by (rule nn_integral_Markov_inequality) (auto simp: assms)
+  also have "\<dots> \<le> ennreal(1/c) * (\<integral>x. u x \<partial>M)"
     by (rule mult_left_mono) (use * \<open>c > 0\<close> in auto)
   finally show ?thesis
     using \<open>0<c\<close> by (simp add: ennreal_mult'[symmetric])
@@ -2109,8 +2018,7 @@
     by (intro emeasure_eq_ennreal_measure [symmetric]) auto
   also note le
   finally show ?thesis
-    by (subst (asm) ennreal_le_iff)
-       (auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms)
+    by (simp add: assms divide_nonneg_pos integral_nonneg_AE)
 qed
 
 theorem%important (in finite_measure) second_moment_method:
@@ -2149,9 +2057,10 @@
   assumes "AE x in M. f x \<ge> 0" "(\<integral>x. f x \<partial>M) > 0"
       and [measurable]: "f \<in> borel_measurable M"
   shows "\<exists>A e. A \<in> sets M \<and> e>0 \<and> emeasure M A > 0 \<and> (\<forall>x \<in> A. f x \<ge> e)"
-proof (rule not_AE_zero_E, auto simp add: assms)
+proof (rule not_AE_zero_E, auto simp: assms)
   assume *: "AE x in M. f x = 0"
-  have "(\<integral>x. f x \<partial>M) = (\<integral>x. 0 \<partial>M)" by (rule integral_cong_AE, auto simp add: *)
+  have "(\<integral>x. f x \<partial>M) = (\<integral>x. 0 \<partial>M)" 
+    by (rule integral_cong_AE, auto simp: *)
   then have "(\<integral>x. f x \<partial>M) = 0" by simp
   then show False using assms(2) by simp
 qed
@@ -2163,28 +2072,31 @@
   shows "((\<lambda>n. (\<integral>x. u n x \<partial>M)) \<longlongrightarrow> (\<integral>x. f x \<partial>M)) F"
 proof -
   have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> (0::ennreal)) F"
-  proof (rule tendsto_sandwich[of "\<lambda>_. 0", where ?h = "\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"], auto simp add: assms)
+  proof (rule tendsto_sandwich[of "\<lambda>_. 0", where ?h = "\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"], auto simp: assms)
     {
       fix n
       have "(\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M) = (\<integral>x. u n x - f x \<partial>M)"
-        apply (rule Bochner_Integration.integral_diff[symmetric]) using assms by auto
+        by (simp add: assms)
       then have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) = norm (\<integral>x. u n x - f x \<partial>M)"
         by auto
-      also have "... \<le> (\<integral>x. norm(u n x - f x) \<partial>M)"
+      also have "\<dots> \<le> (\<integral>x. norm(u n x - f x) \<partial>M)"
         by (rule integral_norm_bound)
       finally have "ennreal(norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<le> (\<integral>x. norm(u n x - f x) \<partial>M)"
         by simp
-      also have "... = (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"
-        apply (rule nn_integral_eq_integral[symmetric]) using assms by auto
-      finally have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)" by simp
+      also have "\<dots> = (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"
+        by (simp add: assms nn_integral_eq_integral)
+      finally have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)" 
+        by simp
     }
     then show "eventually (\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) F"
       by auto
   qed
   then have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F"
     by (simp flip: ennreal_0)
-  then have "((\<lambda>n. ((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" using tendsto_norm_zero_iff by blast
-  then show ?thesis using Lim_null by auto
+  then have "((\<lambda>n. ((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" 
+    using tendsto_norm_zero_iff by blast
+  then show ?thesis 
+    using Lim_null by auto
 qed
 
 text \<open>The next lemma asserts that, if a sequence of functions converges in \<open>L\<^sup>1\<close>, then
@@ -2219,7 +2131,7 @@
     then have "(\<integral>x. norm(u (r n) x) \<partial>M) < (1/2)^n" 
       unfolding r_def by auto
     moreover have "(\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M) = (\<integral>x. norm(u (r n) x) \<partial>M)"
-      by (rule nn_integral_eq_integral, auto simp add: integrable_norm[OF assms(1)[of "r n"]])
+      by (rule nn_integral_eq_integral, auto simp: integrable_norm[OF assms(1)[of "r n"]])
     ultimately show ?thesis by (auto intro: ennreal_lessI)
   qed
 
@@ -2231,31 +2143,32 @@
     have A_bound: "emeasure M (A n) < (1/e) * ennreal((1/2)^n)" for n
     proof -
       have *: "indicator (A n) x \<le> (1/e) * ennreal(norm(u (r n) x))" for x
-        apply (cases "x \<in> A n") unfolding A_def using \<open>0 < e\<close> by (auto simp: ennreal_mult[symmetric])
+        using \<open>0 < e\<close> by (cases "x \<in> A n") (auto simp: A_def ennreal_mult[symmetric])
       have "emeasure M (A n) = (\<integral>\<^sup>+x. indicator (A n) x \<partial>M)" by auto
-      also have "... \<le> (\<integral>\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \<partial>M)"
-        apply (rule nn_integral_mono) using * by auto
-      also have "... = (1/e) * (\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M)"
-        apply (rule nn_integral_cmult) using \<open>e > 0\<close> by auto
-      also have "... < (1/e) * ennreal((1/2)^n)"
+      also have "\<dots> \<le> (\<integral>\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \<partial>M)"
+        by (meson "*" nn_integral_mono)
+      also have "\<dots> = (1/e) * (\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M)"
+        using \<open>e > 0\<close> by (force simp add: intro: nn_integral_cmult)
+      also have "\<dots> < (1/e) * ennreal((1/2)^n)"
         using I[of n] \<open>e > 0\<close> by (intro ennreal_mult_strict_left_mono) auto
       finally show ?thesis by simp
     qed
     have A_fin: "emeasure M (A n) < \<infinity>" for n
       using \<open>e > 0\<close> A_bound[of n]
-      by (auto simp add: ennreal_mult less_top[symmetric])
+      by (auto simp: ennreal_mult less_top[symmetric])
 
     have A_sum: "summable (\<lambda>n. measure M (A n))"
-    proof (rule summable_comparison_test'[of "\<lambda>n. (1/e) * (1/2)^n" 0])
-      have "summable (\<lambda>n. (1/(2::real))^n)" by (simp add: summable_geometric)
+    proof (rule summable_comparison_test')
+      have "summable (\<lambda>n. (1/(2::real))^n)" 
+        by (simp add: summable_geometric)
       then show "summable (\<lambda>n. (1/e) * (1/2)^n)" using summable_mult by blast
       fix n::nat assume "n \<ge> 0"
       have "norm(measure M (A n)) = measure M (A n)" by simp
-      also have "... = enn2real(emeasure M (A n))" unfolding measure_def by simp
-      also have "... < enn2real((1/e) * (1/2)^n)"
+      also have "\<dots> = enn2real(emeasure M (A n))" unfolding measure_def by simp
+      also have "\<dots> < enn2real((1/e) * (1/2)^n)"
         using A_bound[of n] \<open>emeasure M (A n) < \<infinity>\<close> \<open>0 < e\<close>
         by (auto simp: emeasure_eq_ennreal_measure ennreal_mult[symmetric] ennreal_less_iff)
-      also have "... = (1/e) * (1/2)^n"
+      also have "\<dots> = (1/e) * (1/2)^n"
         using \<open>0<e\<close> by auto
       finally show "norm(measure M (A n)) \<le> (1/e) * (1/2)^n" by simp
     qed
@@ -2272,13 +2185,14 @@
       then have "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> e"
         by (simp add: Limsup_bounded)
     }
-    ultimately show "AE x in M. limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0 + ennreal e" by auto
+    ultimately show "AE x in M. limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0 + ennreal e" 
+      by auto
   qed
   moreover
   {
-    fix x assume "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
-    moreover then have "liminf (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
-      by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup)
+    fix x assume x: "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
+    moreover have "liminf (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
+      by (metis x Liminf_le_Limsup le_zero_eq sequentially_bot)
     ultimately have "(\<lambda>n. ennreal (norm(u (r n) x))) \<longlonglongrightarrow> 0"
       using tendsto_Limsup[of sequentially "\<lambda>n. ennreal (norm(u (r n) x))"] by auto
     then have "(\<lambda>n. norm(u (r n) x)) \<longlonglongrightarrow> 0"
@@ -2326,7 +2240,7 @@
         unfolding lim
         using lim
         by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (indicator \<Omega> x *\<^sub>R f x)"])
-           (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR
+           (auto simp: space_restrict_space integrable_restrict_space simp del: norm_scaleR
                  split: split_indicator)
     qed
   qed
@@ -2335,11 +2249,7 @@
 lemma integral_empty:
   assumes "space M = {}"
   shows "integral\<^sup>L M f = 0"
-proof -
-  have "(\<integral> x. f x \<partial>M) = (\<integral> x. 0 \<partial>M)"
-    by(rule integral_cong)(simp_all add: assms)
-  thus ?thesis by simp
-qed
+  by (metis AE_I2 assms empty_iff integral_eq_zero_AE)
 
 subsection \<open>Measure spaces with an associated density\<close>
 
@@ -2379,7 +2289,7 @@
     also have "\<dots> = integral\<^sup>L (density M g) (indicator A)"
       using base by simp
     finally show ?case
-      using base g
+      using base g 
       apply (simp add: int integral_nonneg_AE)
       apply (subst (asm) ennreal_inj)
       apply (auto intro!: integral_nonneg_AE)
@@ -2409,7 +2319,7 @@
       show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L (density M g) f"
         unfolding lim(2)[symmetric]
         by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
-           (insert lim(3-5), auto)
+           (use lim in auto)
     qed
   qed
 qed (simp add: f g integrable_density)
@@ -2438,8 +2348,7 @@
 lemma integrable_distr:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   shows "T \<in> measurable M M' \<Longrightarrow> integrable (distr M M' T) f \<Longrightarrow> integrable M (\<lambda>x. f (T x))"
-  by (subst integrable_distr_eq[symmetric, where g=T])
-     (auto dest: borel_measurable_integrable)
+  by (metis integrable_distr_eq integrable_iff_bounded measurable_distr_eq1)
 
 lemma integral_distr:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -2488,7 +2397,7 @@
       show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L (distr M N g) f"
         unfolding lim(2)[symmetric]
         by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
-           (insert lim(3-5), auto)
+           (use lim in auto)
     qed
   qed
 qed (simp add: f g integrable_distr_eq)
@@ -2532,8 +2441,8 @@
 lemma integrable_count_space_nat_iff:
   fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
   shows "integrable (count_space UNIV) f \<longleftrightarrow> summable (\<lambda>x. norm (f x))"
-  by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top
-           intro:  summable_suminf_not_top)
+  by (auto simp: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top
+           intro: summable_suminf_not_top)
 
 lemma sums_integral_count_space_nat:
   fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
@@ -2563,7 +2472,7 @@
 lemma integral_count_space_nat:
   fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
   shows "integrable (count_space UNIV) f \<Longrightarrow> integral\<^sup>L (count_space UNIV) f = (\<Sum>x. f x)"
-  using sums_integral_count_space_nat by (rule sums_unique)
+  using sums_integral_count_space_nat sums_unique by auto
 
 lemma integrable_bij_count_space:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -2575,11 +2484,8 @@
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes g: "bij_betw g A B"
   shows "integral\<^sup>L (count_space A) (\<lambda>x. f (g x)) = integral\<^sup>L (count_space B) f"
-  using g[THEN bij_betw_imp_funcset]
-  apply (subst distr_bij_count_space[OF g, symmetric])
-  apply (intro integral_distr[symmetric])
-  apply auto
-  done
+  using g[THEN bij_betw_imp_funcset] distr_bij_count_space [OF g]
+  by (metis borel_measurable_count_space integral_distr measurable_count_space_eq1 space_count_space)
 
 lemma has_bochner_integral_count_space_nat:
   fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
@@ -2596,29 +2502,31 @@
 
 proposition integrable_point_measure_finite:
   fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and f :: "'a \<Rightarrow> real"
-  shows "finite A \<Longrightarrow> integrable (point_measure A f) g"
-  unfolding point_measure_def
-  apply (subst density_cong[where f'="\<lambda>x. ennreal (max 0 (f x))"])
-  apply (auto split: split_max simp: ennreal_neg)
-  apply (subst integrable_density)
-  apply (auto simp: AE_count_space integrable_count_space)
-  done
+  assumes "finite A"
+  shows "integrable (point_measure A f) g"
+proof -
+  have "integrable (density (count_space A) (\<lambda>x. ennreal (max 0 (f x)))) g"
+    using assms
+    by (simp add: integrable_count_space integrable_density )
+  then show ?thesis
+    by (smt (verit) AE_I2 borel_measurable_count_space density_cong ennreal_neg point_measure_def)
+qed
 
 subsection \<open>Lebesgue integration on \<^const>\<open>null_measure\<close>\<close>
 
 lemma has_bochner_integral_null_measure_iff[iff]:
   "has_bochner_integral (null_measure M) f 0 \<longleftrightarrow> f \<in> borel_measurable M"
-  by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def]
+  by (auto simp: has_bochner_integral.simps simple_bochner_integral_def[abs_def]
            intro!: exI[of _ "\<lambda>n x. 0"] simple_bochner_integrable.intros)
 
 lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \<longleftrightarrow> f \<in> borel_measurable M"
-  by (auto simp add: integrable.simps)
+  by (auto simp: integrable.simps)
 
 lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0"
-  by (cases "integrable (null_measure M) f")
-     (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq)
+  using integral_norm_bound_ennreal not_integrable_integral_eq by fastforce
 
 subsection \<open>Legacy lemmas for the real-valued Lebesgue integral\<close>
+
 theorem real_lebesgue_integral_def:
   assumes f[measurable]: "integrable M f"
   shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+x. f x \<partial>M) - enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)"
@@ -2701,18 +2609,18 @@
       using i by auto
   next
     show "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ennreal (f i x)) \<partial>M"
-      apply (rule nn_integral_cong_AE)
-      using lim mono nn u_nn
-      apply eventually_elim
-      apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def)
-      done
+    proof (rule nn_integral_cong_AE)
+      show "AE x in M. ennreal (u x) = (SUP i. ennreal (f i x))"
+        using lim mono nn u_nn
+        by eventually_elim (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def)
+    qed
   qed
   also have "\<dots> = ennreal x"
     using mono i nn unfolding nn_integral_eq_integral[OF i pos]
     by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim)
   finally have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = ennreal x" .
   moreover have "(\<integral>\<^sup>+ x. ennreal (- u x) \<partial>M) = 0"
-    using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg)
+    using u u_nn by (subst nn_integral_0_iff_AE) (auto simp: ennreal_neg)
   ultimately show "integrable M u" "integral\<^sup>L M u = x"
     by (auto simp: real_integrable_def real_lebesgue_integral_def u)
 qed
@@ -2777,7 +2685,7 @@
 proof -
   { assume "emeasure M (space M) = \<infinity>" "a \<noteq> 0"
     then have ?thesis
-      by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) }
+      by (auto simp: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) }
   moreover
   { assume "a = 0" then have ?thesis by simp }
   moreover
@@ -2795,17 +2703,14 @@
 lemma (in finite_measure) integrable_const_bound:
   fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
   shows "AE x in M. norm (f x) \<le> B \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> integrable M f"
-  apply (rule integrable_bound[OF integrable_const[of B], of f])
-  apply assumption
-  apply (cases "0 \<le> B")
-  apply auto
-  done
+  using integrable_bound[OF integrable_const[of B], of f]
+  by (smt (verit, ccfv_SIG) eventually_elim2 real_norm_def)
 
 lemma (in finite_measure) integral_bounded_eq_bound_then_AE:
   assumes "AE x in M. f x \<le> (c::real)"
     "integrable M f" "(\<integral>x. f x \<partial>M) = c * measure M (space M)"
   shows "AE x in M. f x = c"
-  apply (rule integral_ineq_eq_0_then_AE) using assms by auto
+  using assms  by (intro integral_ineq_eq_0_then_AE) auto
 
 lemma integral_indicator_finite_real:
   fixes f :: "'a \<Rightarrow> real"
@@ -3030,13 +2935,7 @@
   fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
   assumes "integrable (M1 \<Otimes>\<^sub>M M2) f"
   shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x))"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 ..
-  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
-  show ?thesis unfolding *
-    by (rule integrable_distr[OF measurable_pair_swap'])
-       (simp add: distr_pair_swap[symmetric] assms)
-qed
+  by (smt (verit) assms distr_pair_swap integrable_cong integrable_distr measurable_pair_swap' prod.case_distrib split_cong)
 
 lemma (in pair_sigma_finite) integrable_product_swap_iff:
   fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
@@ -3051,11 +2950,7 @@
   fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
   assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
   shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
-proof -
-  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
-  show ?thesis unfolding *
-    by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric])
-qed
+  by (smt (verit) distr_pair_swap f integral_cong integral_distr measurable_pair_swap' prod.case_distrib split_cong)
 
 theorem (in pair_sigma_finite) Fubini_integrable:
   fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
@@ -3180,14 +3075,13 @@
   case (add f g)
   then have [measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "g \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
     by auto
-  have "(\<integral> x. \<integral> y. f (x, y) + g (x, y) \<partial>M2 \<partial>M1) =
+  have "AE x in M1. LINT y|M2. f (x, y) + g (x, y) =
+                (LINT y|M2. f (x, y)) + (LINT y|M2. g (x, y))"
+    using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)]
+    by eventually_elim simp
+  then have "(\<integral> x. \<integral> y. f (x, y) + g (x, y) \<partial>M2 \<partial>M1) =
     (\<integral> x. (\<integral> y. f (x, y) \<partial>M2) + (\<integral> y. g (x, y) \<partial>M2) \<partial>M1)"
-    apply (rule integral_cong_AE)
-    apply simp_all
-    using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)]
-    apply eventually_elim
-    apply simp
-    done
+    by (intro integral_cong_AE) auto
   also have "\<dots> = (\<integral> x. f x \<partial>(M1 \<Otimes>\<^sub>M M2)) + (\<integral> x. g x \<partial>(M1 \<Otimes>\<^sub>M M2))"
     using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp
   finally show ?case
@@ -3278,10 +3172,7 @@
 lemma (in product_sigma_finite) product_integral_singleton:
   fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
   shows "f \<in> borel_measurable (M i) \<Longrightarrow> (\<integral>x. f (x i) \<partial>Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f"
-  apply (subst distr_singleton[symmetric])
-  apply (subst integral_distr)
-  apply simp_all
-  done
+  by (metis (no_types) distr_singleton insert_iff integral_distr measurable_component_singleton)
 
 lemma (in product_sigma_finite) product_integral_fold:
   fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
@@ -3302,12 +3193,12 @@
     using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def)
   have f_int: "integrable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) ?f"
     by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f)
-  show ?thesis
+  have "LINT x|(Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M). f (merge I J x) =
+        LINT x|Pi\<^sub>M I M. LINT y|Pi\<^sub>M J M. f (merge I J (x, y))"
+    by (simp add: P.integral_fst'[symmetric, OF f_int])
+  then show ?thesis
     apply (subst distr_merge[symmetric, OF IJ fin])
-    apply (subst integral_distr[OF measurable_merge f_borel])
-    apply (subst P.integral_fst'[symmetric, OF f_int])
-    apply simp
-    done
+    by (simp add: integral_distr[OF measurable_merge f_borel])
 qed
 
 lemma (in product_sigma_finite) product_integral_insert:
@@ -3413,15 +3304,11 @@
     proof (intro LIMSEQ_unique)
       show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L N f"
         apply (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
-        using lim
-        apply auto
-        done
+        using lim by auto
       show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"
         unfolding lim
         apply (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
-        using lim M N(2)
-        apply auto
-        done
+        using lim M N by auto
     qed
   qed
 qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms])
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -391,22 +391,22 @@
 qed fact
 
 locale kuhn_simplex =
-  fixes p n and base upd and s :: "(nat \<Rightarrow> nat) set"
+  fixes p n and base upd and S :: "(nat \<Rightarrow> nat) set"
   assumes base: "base \<in> {..< n} \<rightarrow> {..< p}"
   assumes base_out: "\<And>i. n \<le> i \<Longrightarrow> base i = p"
   assumes upd: "bij_betw upd {..< n} {..< n}"
-  assumes s_pre: "s = (\<lambda>i j. if j \<in> upd`{..< i} then Suc (base j) else base j) ` {.. n}"
+  assumes s_pre: "S = (\<lambda>i j. if j \<in> upd`{..< i} then Suc (base j) else base j) ` {.. n}"
 begin
 
 definition "enum i j = (if j \<in> upd`{..< i} then Suc (base j) else base j)"
 
-lemma s_eq: "s = enum ` {.. n}"
+lemma s_eq: "S = enum ` {.. n}"
   unfolding s_pre enum_def[abs_def] ..
 
 lemma upd_space: "i < n \<Longrightarrow> upd i < n"
   using upd by (auto dest!: bij_betwE)
 
-lemma s_space: "s \<subseteq> {..< n} \<rightarrow> {.. p}"
+lemma s_space: "S \<subseteq> {..< n} \<rightarrow> {.. p}"
 proof -
   { fix i assume "i \<le> n" then have "enum i \<in> {..< n} \<rightarrow> {.. p}"
     proof (induct i)
@@ -437,19 +437,19 @@
 lemma enum_0: "enum 0 = base"
   by (simp add: enum_def[abs_def])
 
-lemma base_in_s: "base \<in> s"
+lemma base_in_s: "base \<in> S"
   unfolding s_eq by (subst enum_0[symmetric]) auto
 
-lemma enum_in: "i \<le> n \<Longrightarrow> enum i \<in> s"
+lemma enum_in: "i \<le> n \<Longrightarrow> enum i \<in> S"
   unfolding s_eq by auto
 
 lemma one_step:
-  assumes a: "a \<in> s" "j < n"
-  assumes *: "\<And>a'. a' \<in> s \<Longrightarrow> a' \<noteq> a \<Longrightarrow> a' j = p'"
+  assumes a: "a \<in> S" "j < n"
+  assumes *: "\<And>a'. a' \<in> S \<Longrightarrow> a' \<noteq> a \<Longrightarrow> a' j = p'"
   shows "a j \<noteq> p'"
 proof
   assume "a j = p'"
-  with * a have "\<And>a'. a' \<in> s \<Longrightarrow> a' j = p'"
+  with * a have "\<And>a'. a' \<in> S \<Longrightarrow> a' j = p'"
     by auto
   then have "\<And>i. i \<le> n \<Longrightarrow> enum i j = p'"
     unfolding s_eq by auto
@@ -481,16 +481,16 @@
 lemma enum_strict_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i < enum j \<longleftrightarrow> i < j"
   using enum_mono[of i j] enum_inj[of i j] by (auto simp: le_less)
 
-lemma chain: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a \<le> b \<or> b \<le> a"
+lemma chain: "a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> a \<le> b \<or> b \<le> a"
   by (auto simp: s_eq enum_mono)
 
-lemma less: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a i < b i \<Longrightarrow> a < b"
+lemma less: "a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> a i < b i \<Longrightarrow> a < b"
   using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric])
 
-lemma enum_0_bot: "a \<in> s \<Longrightarrow> a = enum 0 \<longleftrightarrow> (\<forall>a'\<in>s. a \<le> a')"
+lemma enum_0_bot: "a \<in> S \<Longrightarrow> a = enum 0 \<longleftrightarrow> (\<forall>a'\<in>S. a \<le> a')"
   unfolding s_eq by (auto simp: enum_mono Ball_def)
 
-lemma enum_n_top: "a \<in> s \<Longrightarrow> a = enum n \<longleftrightarrow> (\<forall>a'\<in>s. a' \<le> a)"
+lemma enum_n_top: "a \<in> S \<Longrightarrow> a = enum n \<longleftrightarrow> (\<forall>a'\<in>S. a' \<le> a)"
   unfolding s_eq by (auto simp: enum_mono Ball_def)
 
 lemma enum_Suc: "i < n \<Longrightarrow> enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))"
@@ -499,51 +499,51 @@
 lemma enum_eq_p: "i \<le> n \<Longrightarrow> n \<le> j \<Longrightarrow> enum i j = p"
   by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric])
 
-lemma out_eq_p: "a \<in> s \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p"
+lemma out_eq_p: "a \<in> S \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p"
   unfolding s_eq by (auto simp: enum_eq_p)
 
-lemma s_le_p: "a \<in> s \<Longrightarrow> a j \<le> p"
+lemma s_le_p: "a \<in> S \<Longrightarrow> a j \<le> p"
   using out_eq_p[of a j] s_space by (cases "j < n") auto
 
-lemma le_Suc_base: "a \<in> s \<Longrightarrow> a j \<le> Suc (base j)"
+lemma le_Suc_base: "a \<in> S \<Longrightarrow> a j \<le> Suc (base j)"
   unfolding s_eq by (auto simp: enum_def)
 
-lemma base_le: "a \<in> s \<Longrightarrow> base j \<le> a j"
+lemma base_le: "a \<in> S \<Longrightarrow> base j \<le> a j"
   unfolding s_eq by (auto simp: enum_def)
 
 lemma enum_le_p: "i \<le> n \<Longrightarrow> j < n \<Longrightarrow> enum i j \<le> p"
   using enum_in[of i] s_space by auto
 
-lemma enum_less: "a \<in> s \<Longrightarrow> i < n \<Longrightarrow> enum i < a \<longleftrightarrow> enum (Suc i) \<le> a"
+lemma enum_less: "a \<in> S \<Longrightarrow> i < n \<Longrightarrow> enum i < a \<longleftrightarrow> enum (Suc i) \<le> a"
   unfolding s_eq by (auto simp: enum_strict_mono enum_mono)
 
 lemma ksimplex_0:
-  "n = 0 \<Longrightarrow> s = {(\<lambda>x. p)}"
+  "n = 0 \<Longrightarrow> S = {(\<lambda>x. p)}"
   using s_eq enum_def base_out by auto
 
 lemma replace_0:
-  assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = 0" and "x \<in> s"
+  assumes "j < n" "a \<in> S" and p: "\<forall>x\<in>S - {a}. x j = 0" and "x \<in> S"
   shows "x \<le> a"
 proof cases
   assume "x \<noteq> a"
   have "a j \<noteq> 0"
     using assms by (intro one_step[where a=a]) auto
-  with less[OF \<open>x\<in>s\<close> \<open>a\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close>
+  with less[OF \<open>x\<in>S\<close> \<open>a\<in>S\<close>, of j] p[rule_format, of x] \<open>x \<in> S\<close> \<open>x \<noteq> a\<close>
   show ?thesis
     by auto
 qed simp
 
 lemma replace_1:
-  assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = p" and "x \<in> s"
+  assumes "j < n" "a \<in> S" and p: "\<forall>x\<in>S - {a}. x j = p" and "x \<in> S"
   shows "a \<le> x"
 proof cases
   assume "x \<noteq> a"
   have "a j \<noteq> p"
     using assms by (intro one_step[where a=a]) auto
-  with enum_le_p[of _ j] \<open>j < n\<close> \<open>a\<in>s\<close>
+  with enum_le_p[of _ j] \<open>j < n\<close> \<open>a\<in>S\<close>
   have "a j < p"
     by (auto simp: less_le s_eq)
-  with less[OF \<open>a\<in>s\<close> \<open>x\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close>
+  with less[OF \<open>a\<in>S\<close> \<open>x\<in>S\<close>, of j] p[rule_format, of x] \<open>x \<in> S\<close> \<open>x \<noteq> a\<close>
   show ?thesis
     by auto
 qed simp
--- a/src/HOL/Analysis/Change_Of_Vars.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Change_Of_Vars.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -48,10 +48,8 @@
       by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
     moreover have "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})) \<subseteq> {x. ?mn \<bullet> x = b$m}"
       using \<open>m \<noteq> n\<close> ab_ne
-      apply (auto simp: algebra_simps mem_box_cart inner_axis')
-      apply (drule_tac x=m in spec)+
-      apply simp
-      done
+      apply (clarsimp simp: algebra_simps mem_box_cart inner_axis')
+      by (smt (verit, ccfv_SIG) interval_ne_empty_cart(1))
     ultimately show "negligible (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}))"
       by (rule negligible_subset)
     show "?f ` cbox a b \<union> cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<union> cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} = cbox a ?c" (is "?lhs = _")
@@ -59,9 +57,8 @@
       show "?lhs \<subseteq> cbox a ?c"
         by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans)
       show "cbox a ?c \<subseteq> ?lhs"
-        apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \<open>m \<noteq> n\<close>])
-        apply (auto simp: mem_box_cart split: if_split_asm)
-        done
+        apply (clarsimp simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \<open>m \<noteq> n\<close>])
+        by (smt (verit, del_insts) mem_box_cart(2) vec_lambda_beta)
     qed
   qed (fact fab)
   let ?d = "\<chi> i. if i = m then a $ m - b $ m else 0"
@@ -106,17 +103,15 @@
       by (rule negligible_subset)
   qed
   have ac_ne: "cbox a ?c \<noteq> {}"
-    using ab_ne an
-    by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans)
+    by (smt (verit, del_insts) ab_ne an interval_ne_empty_cart(1) vec_lambda_beta)
   have ax_ne: "cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i) \<noteq> {}"
     using ab_ne an
-    by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans)
+    by (smt (verit, ccfv_threshold) interval_ne_empty_cart(1) vec_lambda_beta)
   have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\<chi> i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)"
     by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove
              if_distrib [of "\<lambda>u. u - z" for z] prod.remove)
   show ?Q
-    using eq1 eq2 eq3
-    by (simp add: algebra_simps)
+    using eq1 eq2 eq3 by (simp add: algebra_simps)
 qed
 
 
@@ -214,10 +209,8 @@
       by auto
     have nfS: "negligible (?f ` S)"
       by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \<open>m k = 0\<close> in auto)
-    then have "(?f ` S) \<in> lmeasurable"
-      by (simp add: negligible_iff_measure)
-    with nfS show ?thesis
-      by (simp add: prm negligible_iff_measure0)
+    then show ?thesis
+      by (simp add: negligible_iff_measure prm)
   qed
   then show "(?f ` S) \<in> lmeasurable" ?MEQ
     by metis+
@@ -297,10 +290,7 @@
     then have 1: "\<bar>det (matrix ?h)\<bar> = 1"
       by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult)
     show "?h ` S \<in> lmeasurable \<and> ?Q ?h S"
-    proof
-      show "?h ` S \<in> lmeasurable" "?Q ?h S"
-        using measure_linear_sufficient [OF lin \<open>S \<in> lmeasurable\<close>] meq 1 by force+
-    qed
+      using measure_linear_sufficient [OF lin \<open>S \<in> lmeasurable\<close>] meq 1 by force
   next
     fix m n :: "'n" and S :: "(real, 'n) vec set"
     assume "m \<noteq> n" and "S \<in> lmeasurable"
@@ -355,7 +345,7 @@
         by (rule ssubst) (rule measure_translation)
       also have "\<dots> = measure lebesgue ((\<lambda>v. \<chi> i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))"
         by (metis (no_types, lifting) cbox_translation)
-      also have "\<dots> = measure lebesgue ((+) (\<chi> i. if i = n then - a $ n else 0) ` cbox a b)"
+      also have "\<dots> = measure lebesgue ((+) ?v ` cbox a b)"
         apply (subst measure_shear_interval)
         using \<open>m \<noteq> n\<close> ne apply auto
         apply (simp add: cbox_translation)
@@ -617,9 +607,7 @@
           have inj: "inj_on (\<lambda>(x, y). ball x y) K"
             by (force simp: inj_on_def ball_eq_ball_iff dest: gt0)
           have disjnt: "disjoint ((\<lambda>(x, y). ball x y) ` K)"
-            using pwC that
-            apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff)
-            by (metis subsetD fst_conv snd_conv)
+            using pwC that pairwise_image pairwise_mono by fastforce
           have "?l \<le> (\<Sum>i\<in>K. ?\<mu> (case i of (x, s) \<Rightarrow> f ` (S \<inter> ball x s)))"
           proof (rule measure_UNION_le [OF \<open>finite K\<close>], clarify)
             fix x r
@@ -630,8 +618,7 @@
               by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue)
           qed
           also have "\<dots> \<le> (\<Sum>(x,s) \<in> K. (B + e) * ?\<mu> (ball x s))"
-            apply (rule sum_mono)
-            using Csub r \<open>K \<subseteq> C\<close> by auto
+            using Csub r \<open>K \<subseteq> C\<close>  by (intro sum_mono) auto
           also have "\<dots> = (B + e) * (\<Sum>(x,s) \<in> K. ?\<mu> (ball x s))"
             by (simp add: prod.case_distrib sum_distrib_left)
           also have "\<dots> = (B + e) * sum ?\<mu> ((\<lambda>(x, y). ball x y) ` K)"
@@ -641,8 +628,8 @@
             by (subst measure_Union') (auto simp: disjnt measure_Union')
           also have "\<dots> \<le> (B + e) * ?\<mu> T"
             using \<open>B \<ge> 0\<close> \<open>e > 0\<close> that apply simp
-            apply (rule measure_mono_fmeasurable [OF _ _ \<open>T \<in> lmeasurable\<close>])
-            using Csub rT by force+
+            using measure_mono_fmeasurable [OF _ _ \<open>T \<in> lmeasurable\<close>] Csub rT
+            by (smt (verit) SUP_least measure_nonneg measure_notin_sets mem_Collect_eq old.prod.case subset_iff)
           also have "\<dots> \<le> (B + e) * (?\<mu> S + d)"
             using \<open>B \<ge> 0\<close> \<open>e > 0\<close> Tless by simp
           finally show ?thesis .
@@ -732,10 +719,8 @@
       using set_integrable_subset [OF aint_S] meas_t T_def by blast
     have Seq: "S = (\<Union>n. T n)"
       apply (auto simp: T_def)
-      apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI)
-      using that apply auto
-      using of_int_floor_le pos_le_divide_eq apply blast
-      by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt)
+      apply (rule_tac x = "nat \<lfloor>\<bar>det (matrix (f' x))\<bar> / e\<rfloor>" in exI)
+      by (smt (verit, del_insts) divide_nonneg_nonneg floor_eq_iff of_nat_nat pos_divide_less_eq that zero_le_floor)
     have meas_ft: "f ` T n \<in> lmeasurable" for n
     proof (rule measurable_bounded_differentiable_image)
       show "T n \<in> lmeasurable"
@@ -779,8 +764,8 @@
         case (less m n)
         have False if "T n \<subseteq> T m" "x \<in> T n" for x
           using \<open>e > 0\<close> \<open>m < n\<close> that
-          apply (auto simp: T_def  mult.commute intro: less_le_trans dest!: subsetD)
-          by (metis add.commute less_le_trans nat_less_real_le not_le mult_le_cancel_iff2)
+          apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD)
+          by (smt (verit, best) mult_less_cancel_left_disj nat_less_real_le)
         then show ?case
           using less.prems by blast
       qed auto
@@ -839,10 +824,10 @@
           unfolding m_def
         proof (rule integral_subset_le)
           have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on (\<Union>k\<le>n. T k)"
-            apply (rule set_integrable_subset [OF aint_S])
-             apply (intro measurable meas_t fmeasurableD)
-            apply (force simp: Seq)
-            done
+          proof (rule set_integrable_subset [OF aint_S])
+            show "\<Union> (T ` {..n}) \<in> sets lebesgue"
+              by (intro measurable meas_t fmeasurableD)
+          qed (force simp: Seq)
           then show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on (\<Union>k\<le>n. T k)"
             using absolutely_integrable_on_def by blast
         qed (use Seq int in auto)
@@ -905,9 +890,9 @@
   let ?I = "\<lambda>n::nat. cbox (vec (-n)) (vec n) \<inter> S"
   let ?\<mu> = "measure lebesgue"
   have "x \<in> cbox (vec (- real (nat \<lceil>norm x\<rceil>))) (vec (real (nat \<lceil>norm x\<rceil>)))" for x :: "real^'n::_"
-    apply (auto simp: mem_box_cart)
-    apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans)
-    by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge)
+    apply (simp add: mem_box_cart)
+    by (smt (verit, best) Finite_Cartesian_Product.norm_nth_le nat_ceiling_le_eq 
+        real_nat_ceiling_ge real_norm_def)
   then have Seq: "S = (\<Union>n. ?I n)"
     by auto
   have fIn: "f ` ?I n \<in> lmeasurable"
@@ -918,14 +903,7 @@
     moreover have "\<And>x. x \<in> ?I n \<Longrightarrow> (f has_derivative f' x) (at x within ?I n)"
       by (meson Int_iff deriv has_derivative_subset subsetI)
     moreover have int_In: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on ?I n"
-    proof -
-      have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
-        using int absolutely_integrable_integrable_bound by force
-      then have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on ?I n"
-        by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset)
-      then show ?thesis
-        using absolutely_integrable_on_def by blast
-    qed
+      by (metis (mono_tags) Int_commute int integrable_altD(1) integrable_restrict_Int)
     ultimately have "f ` ?I n \<in> lmeasurable" "?\<mu> (f ` ?I n) \<le> integral (?I n) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
       using m_diff_image_weak by metis+
     moreover have "integral (?I n) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
@@ -972,14 +950,16 @@
     proof -
       have g: "g n x \<le> g (N + n) x" for N
         by (rule transitive_stepwise_le) (use inc_g in auto)
-      have "\<exists>na\<ge>N. g n x - f x \<le> dist (g na x) (f x)" for N
-        apply (rule_tac x="N+n" in exI)
-        using g [of N] by (auto simp: dist_norm)
+      have "\<exists>m\<ge>N. g n x - f x \<le> dist (g m x) (f x)" for N
+      proof
+        show "N \<le> N + n \<and> g n x - f x \<le> dist (g (N + n) x) (f x)"
+          using g [of N] by (auto simp: dist_norm)
+      qed
       with that show ?thesis
         using diff_gt_0_iff_gt by blast
     qed
     with lim show ?thesis
-      apply (auto simp: lim_sequentially)
+      unfolding lim_sequentially
       by (meson less_le_not_le not_le_imp_less)
   qed
   moreover
@@ -993,10 +973,8 @@
       fix k::real
       assume "k \<in> \<int>" and k: "\<bar>k\<bar> \<le> 2 ^ (2*n)"
       show "0 \<le> k/2^n * ?\<Omega> n k x"
-        using f \<open>k \<in> \<int>\<close> apply (auto simp: indicator_def field_split_simps Ints_def)
-        apply (drule spec [where x=x])
-        using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"]
-        by linarith
+        using f \<open>k \<in> \<int>\<close> apply (clarsimp simp: indicator_def field_split_simps Ints_def)
+        by (smt (verit) int_less_real_le mult_nonneg_nonneg of_int_0 zero_le_power)
     qed
     show "?g n x \<le> ?g (Suc n) x" for n x
     proof -
@@ -1049,19 +1027,12 @@
             show "finite {k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)}"
               by (rule finite_abs_int_segment)
             show "(*) 2 ` {k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2^(2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2^(2*n)} \<subseteq> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)}"
-              apply auto
+              apply (clarsimp simp: image_subset_iff)
               using one_le_power [of "2::real" "2*n"]  by linarith
             have *: "\<lbrakk>x \<in> (S \<union> T) - U; \<And>x. x \<in> S \<Longrightarrow> x \<in> U; \<And>x. x \<in> T \<Longrightarrow> x \<in> U\<rbrakk> \<Longrightarrow> P x" for S T U P
               by blast
             have "0 \<le> b" if "b \<in> \<int>" "f x * (2 * 2^n) < b + 1" for b
-            proof -
-              have "0 \<le> f x * (2 * 2^n)"
-                by (simp add: f)
-              also have "\<dots> < b+1"
-                by (simp add: that)
-              finally show "0 \<le> b"
-                using \<open>b \<in> \<int>\<close> by (auto simp: elim!: Ints_cases)
-            qed
+              by (smt (verit, ccfv_SIG) Ints_cases f int_le_real_less mult_nonneg_nonneg of_int_add one_le_power that)
             then show "0 \<le> b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \<le> f y \<and> f y < (b + 1)/2 ^ Suc n} x"
                   if "b \<in> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)} -
                           ((*) 2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})" for b
@@ -1125,10 +1096,8 @@
           using N2 by (simp add: divide_simps mult.commute) linarith
         also have "\<dots> \<le> \<bar>2^n\<bar> * e"
           using that \<open>e > 0\<close> by auto
-        finally have "dist (?m/2^n) (f x) < e"
-          by (simp add: dist_norm)
-        then show ?thesis
-          using eq by linarith
+        finally show ?thesis
+          using eq by (simp add: dist_real_def)
       qed
       then show "\<exists>no. \<forall>n\<ge>no. dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k * ?\<Omega> n k x/2^n) (f x) < e"
         by force
@@ -1213,9 +1182,7 @@
   ultimately have dim: "dim {x. f x = 0} = DIM('a)"
     by force
   then show ?thesis
-    using dim_eq_full
-    by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis
-        mem_Collect_eq module_hom_zero span_base span_raw_def)
+    by (metis (mono_tags, lifting) dim_eq_full UNIV_I eq_0_on_span mem_Collect_eq span_raw_def)
 qed
 
 lemma lemma_partial_derivatives:
@@ -1368,8 +1335,7 @@
             next
               case False
               have "\<bar>v \<bullet> (y - x)\<bar> < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)"
-                apply (rule dless)
-                using False \<open>y \<in> S\<close> d by (auto simp: norm_minus_commute)
+                by (metis Diff_iff False \<open>y \<in> S\<close> d dless empty_iff insert_iff norm_minus_commute)
               also have "\<dots> \<le> norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))"
                 using d r \<open>e > 0\<close> by (simp add: field_simps norm_minus_commute mult_left_mono)
               finally show ?thesis .
@@ -1500,8 +1466,7 @@
             also have "\<dots> = UNIV"
             proof -
               have "dim ?CA \<le> CARD('m)"
-                using dim_subset_UNIV[of ?CA]
-                by auto
+                using dim_subset_UNIV[of ?CA] by auto
               moreover have "False" if less: "dim ?CA < CARD('m)"
               proof -
                 obtain d where "d \<noteq> 0" and d: "\<And>y. y \<in> span ?CA \<Longrightarrow> orthogonal d y"
@@ -1570,7 +1535,7 @@
                 proof (rule eventuallyI)
                   fix n
                   show "0 \<le> \<bar>d \<bullet> ((\<gamma> n - x) /\<^sub>R norm (\<gamma> n - x))\<bar> - \<xi>"
-                  using \<gamma>le [of n] \<gamma>Sx by (auto simp: abs_mult divide_simps)
+                    using \<gamma>le [of n] \<gamma>Sx by (auto simp: abs_mult divide_simps)
                 qed
                 ultimately have "\<xi> \<le> \<bar>d \<bullet> z\<bar>"
                   using tendsto_lowerbound [where a=0] by fastforce
@@ -1599,8 +1564,7 @@
                       also have "\<dots> \<le> 1/(Suc j) * norm (\<gamma> p - x) + 1/(Suc i) * norm (\<gamma> p - x)"
                         by (metis A Diff_iff \<gamma>Sx dist_norm p add_mono)
                       also have "\<dots> \<le> 1/N * norm (\<gamma> p - x) + 1/N * norm (\<gamma> p - x)"
-                        apply (intro add_mono mult_right_mono)
-                        using ij \<open>N > 0\<close> by (auto simp: field_simps)
+                        using ij \<open>N > 0\<close> by (intro add_mono mult_right_mono) (auto simp: field_simps)
                       also have "\<dots> = 2 / N * norm (\<gamma> p - x)"
                         by simp
                       finally have no_le: "norm ((A i - A j) *v (\<gamma> p - x)) \<le> 2 / N * norm (\<gamma> p - x)" .
@@ -1807,8 +1771,12 @@
                       by metis
                   qed
                   also have "\<dots> \<le> e * norm (y - x) / 4"
-                    using \<open>e > 0\<close> apply (simp add: norm_mult abs_mult)
-                    by (metis component_le_norm_cart vector_minus_component)
+                  proof -
+                    have "\<bar>y $ n - x $ n\<bar> \<le> norm (y - x)"
+                      by (metis component_le_norm_cart vector_minus_component)
+                    with \<open>e > 0\<close> show ?thesis
+                      by (simp add: norm_mult abs_mult)
+                  qed
                   finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \<le> e * norm (y - x) / 4" .
                   show "0 < e * norm (y - x)"
                     by (simp add: False \<open>e > 0\<close>)
@@ -1869,12 +1837,7 @@
 lemma sets_lebesgue_almost_borel:
   assumes "S \<in> sets lebesgue"
   obtains B N where "B \<in> sets borel" "negligible N" "B \<union> N = S"
-proof -
-  obtain T N N' where "S = T \<union> N" "N \<subseteq> N'" "N' \<in> null_sets lborel" "T \<in> sets borel"
-    using sets_completionE [OF assms] by auto
-  then show thesis
-    by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that)
-qed
+  by (metis assms negligible_iff_null_sets negligible_subset null_sets_completionI sets_completionE sets_lborel)
 
 lemma double_lebesgue_sets:
  assumes S: "S \<in> sets lebesgue" and T: "T \<in> sets lebesgue" and fim: "f ` S \<subseteq> T"
@@ -1906,8 +1869,7 @@
      and 2 [rule_format]: "\<forall>U. negligible U \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
      and "U \<in> sets lebesgue" "U \<subseteq> T"
   then obtain C N where C: "C \<in> sets borel \<and> negligible N \<and> C \<union> N = U"
-    using sets_lebesgue_almost_borel
-    by metis
+    using sets_lebesgue_almost_borel by metis
   then have "{x \<in> S. f x \<in> C} \<in> sets lebesgue"
     by (blast intro: 1)
   moreover have "{x \<in> S. f x \<in> N} \<in> sets lebesgue"
@@ -1974,12 +1936,7 @@
     and mS: "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
   proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e])
     have "norm a *\<^sub>R axis k 1 \<bullet> x = 0" if "T x \<in> P" for x
-    proof -
-      have "a \<bullet> T x = 0"
-        using P that by blast
-      then show ?thesis
-        by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def)
-    qed
+      by (smt (verit, del_insts) P T a mem_Collect_eq orthogonal_transformation_def subset_eq that)
     then show "T -` P \<subseteq> {x. norm a *\<^sub>R axis k 1 \<bullet> x = 0}"
       by auto
   qed (use assms T in auto)
@@ -1990,22 +1947,11 @@
     have "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> T ` {z. norm z \<le> m \<and> (\<exists>t\<in>T -` P. norm (z - t) \<le> e)}"
     proof clarsimp
       fix x t
-      assume "norm x \<le> m" "t \<in> P" "norm (x - t) \<le> e"
+      assume \<section>: "norm x \<le> m" "t \<in> P" "norm (x - t) \<le> e"
       then have "norm (inv T x) \<le> m"
         using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm)
       moreover have "\<exists>t\<in>T -` P. norm (inv T x - t) \<le> e"
-      proof
-        have "T (inv T x - inv T t) = x - t"
-          using T linear_diff orthogonal_transformation_def
-          by (metis (no_types, opaque_lifting) Tinv)
-        then have "norm (inv T x - inv T t) = norm (x - t)"
-          by (metis T orthogonal_transformation_norm)
-        then show "norm (inv T x - inv T t) \<le> e"
-          using \<open>norm (x - t) \<le> e\<close> by linarith
-       next
-         show "inv T t \<in> T -` P"
-           using \<open>t \<in> P\<close> by force
-      qed
+        by (smt (verit, del_insts) T Tinv \<section> linear_diff orthogonal_transformation_def orthogonal_transformation_norm vimage_eq)
       ultimately show "x \<in> T ` {z. norm z \<le> m \<and> (\<exists>t\<in>T -` P. norm (z - t) \<le> e)}"
         by force
     qed
@@ -2124,8 +2070,8 @@
             and subT: "{z. norm(z - f x) \<le> (2 * B) * norm(v - u) \<and> (\<exists>t \<in> range (f' x). norm(z - f x - t) \<le> d * norm(v - u))} \<subseteq> T"
             and measT: "?\<mu> T \<le> (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)"
                         (is "_ \<le> ?DVU")
-        apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"])
-        using \<open>B > 0\<close> \<open>d > 0\<close> by simp_all
+        using Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)"]
+        using \<open>B > 0\<close> \<open>d > 0\<close> by auto
       show ?thesis
       proof (intro exI conjI)
         have "f ` (K \<inter> S) \<subseteq> {z. norm(z - f x) \<le> (2 * B) * norm(v - u) \<and> (\<exists>t \<in> range (f' x). norm(z - f x - t) \<le> d * norm(v - u))}"
@@ -2150,18 +2096,13 @@
             have "norm (f' x (y - x)) \<le> onorm (f' x) * norm (y - x)"
               using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1))
             also have "\<dots> \<le> B * norm (v - u)"
-            proof (rule mult_mono)
-              show "onorm (f' x) \<le> B"
-                using B x by blast
-            qed (use \<open>B > 0\<close> yx_le in auto)
+              by (meson B IntE lin_f' linear_linear mult_mono' norm_ge_zero onorm_pos_le x(1) yx_le)
             finally show "norm (f' x (y - x)) \<le> B * norm (v - u)" .
             show "d * norm (y - x) \<le> B * norm (v - u)"
               using \<open>B > 0\<close> by (auto intro: mult_mono [OF \<open>d \<le> B\<close> yx_le])
           qed
           show "\<exists>t. norm (f y - f x - f' x t) \<le> d * norm (v - u)"
-            apply (rule_tac x="y-x" in exI)
-            using \<open>d > 0\<close> yx_le le_dyx mult_left_mono [where c=d]
-            by (meson order_trans mult_le_cancel_iff2)
+            by (smt (verit, best) \<open>0 < d\<close> le_dyx mult_le_cancel_left_pos yx_le)
         qed
         with subT show "f ` (K \<inter> S) \<subseteq> T" by blast
         show "?\<mu> T \<le> e / (2*c) ^ ?m * ?\<mu> K"
@@ -2355,8 +2296,12 @@
       have "(\<lambda>x. if x \<in> {t. h n (g t) = y} then ?D x else 0) integrable_on S"
       proof (rule measurable_bounded_by_integrable_imp_integrable)
         have "(\<lambda>x. ?D x) \<in> borel_measurable (lebesgue_on ({t. h n (g t) = y} \<inter> S))"
-          apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g])
-          by (meson der_g IntD2 has_derivative_subset inf_le2)
+        proof -
+          have "(\<lambda>v. det (matrix (g' v))) \<in> borel_measurable (lebesgue_on (S \<inter> {v. h n (g v) = y}))"
+            by (metis Int_lower1 S assms(4) borel_measurable_det_Jacobian measurable_restrict_mono)
+          then show ?thesis
+            by (simp add: Int_commute)
+        qed
         then have "(\<lambda>x. if x \<in> {t. h n (g t) = y} \<inter> S then ?D x else 0) \<in> borel_measurable lebesgue"
           by (rule borel_measurable_if_I [OF _ h_lmeas])
         then show "(\<lambda>x. if x \<in> {t. h n (g t) = y} then ?D x else 0) \<in> borel_measurable (lebesgue_on S)"
@@ -2371,28 +2316,23 @@
       then have int_det: "(\<lambda>t. \<bar>det (matrix (g' t))\<bar>) integrable_on ({t. h n (g t) = y} \<inter> S)"
         using integrable_restrict_Int by force
       have "(g ` ({t. h n (g t) = y} \<inter> S)) \<in> lmeasurable"
-        apply (rule measurable_differentiable_image [OF h_lmeas])
-         apply (blast intro: has_derivative_subset [OF der_g])
-        apply (rule int_det)
-        done
+        by (blast intro: has_derivative_subset [OF der_g]  measurable_differentiable_image [OF h_lmeas] int_det)
       moreover have "g ` ({t. h n (g t) = y} \<inter> S) = {x. h n x = y} \<inter> g ` S"
         by blast
       moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \<inter> S))
                      \<le> integral ({t. h n (g t) = y} \<inter> S) (\<lambda>t. \<bar>det (matrix (g' t))\<bar>)"
-        apply (rule measure_differentiable_image [OF h_lmeas _ int_det])
-        apply (blast intro: has_derivative_subset [OF der_g])
-        done
+        by (blast intro: has_derivative_subset [OF der_g] measure_differentiable_image [OF h_lmeas _ int_det])
       ultimately show ?thesis
         using \<open>y > 0\<close> integral_restrict_Int [of S "{t. h n (g t) = y}" "\<lambda>t. \<bar>det (matrix (g' t))\<bar> * y"]
         apply (simp add: integrable_on_indicator integral_indicator)
         apply (simp add: indicator_def of_bool_def if_distrib cong: if_cong)
         done
     qed
-    have hn_int: "h n integrable_on g ` S"
-      apply (subst hn_eq)
-      using yind by (force intro: integrable_sum [OF fin_R])
-    then show ?thesis
+    show ?thesis
     proof
+      show "h n integrable_on g ` S"
+        apply (subst hn_eq)
+        using yind by (force intro: integrable_sum [OF fin_R])
       have "?lhs = integral (g ` S) (\<lambda>x. \<Sum>y\<in>range (h n). y * indicat_real {x. h n x = y} x)"
         by (metis hn_eq)
       also have "\<dots> = (\<Sum>y\<in>range (h n). integral (g ` S) (\<lambda>x. y * indicat_real {x. h n x = y} x))"
@@ -2414,8 +2354,8 @@
         next
           fix x
           assume "x \<in> S"
-          have "y * indicat_real {x. h n x = y} (g x) \<le> f (g x)"
-            by (metis \<open>x \<in> S\<close> h_le_f indicator_simps(1) indicator_simps(2) mem_Collect_eq mult.commute mult.left_neutral mult_zero_left nonneg_fg)
+          then have "y * indicat_real {x. h n x = y} (g x) \<le> f (g x)"
+            by (metis (full_types) h_le_f indicator_simps mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg)
           with \<open>y \<ge> 0\<close> show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \<le> ?D x * f(g x)"
             by (simp add: abs_mult mult.assoc mult_left_mono)
         qed (use S det_int_fg in auto)
@@ -2522,8 +2462,10 @@
         by (rule Df_borel)
       finally have *: "(\<lambda>x. \<bar>?D x\<bar> * f (g x)) \<in> borel_measurable (lebesgue_on S')"
         by (simp add: borel_measurable_if_D)
-      have "?h \<in> borel_measurable (lebesgue_on S')"
-        by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS')
+      have "(\<lambda>v. det (matrix (g' v))) \<in> borel_measurable (lebesgue_on S')"
+        using S' borel_measurable_det_Jacobian der_gS' by blast
+      then have "?h \<in> borel_measurable (lebesgue_on S')"
+        using "*" borel_measurable_abs borel_measurable_inverse borel_measurable_scaleR by blast
       moreover have "?h x = f(g x)" if "x \<in> S'" for x
         using that by (auto simp: S'_def)
       ultimately have "(\<lambda>x. f(g x)) \<in> borel_measurable (lebesgue_on S')"
@@ -2586,8 +2528,8 @@
     then have "f integrable_on g ` {x \<in> S. ?D x \<noteq> 0}"
       by (auto simp: image_iff elim!: integrable_eq)
     then show "f integrable_on g ` S"
-      apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset])
-      using negg null by auto
+      using negg null
+      by (auto intro: integrable_spike_set [OF _ empty_imp_negligible negligible_subset])
     have "integral (g ` S) f = integral (g ` {x \<in> S. ?D x \<noteq> 0}) f"
       using negg by (auto intro: negligible_subset integral_spike_set)
     also have "\<dots> = integral (g ` {x \<in> S. ?D x \<noteq> 0}) (\<lambda>x. if x \<in> g ` ?F then f x else 0)"
@@ -2653,15 +2595,15 @@
 
   have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \<in> ?P" for x
       using der_g has_derivative_subset that by force
-  have "f integrable_on g ` ?P \<and> integral (g ` ?P) f \<le> integral ?P ?D"
-  proof (rule integral_on_image_ubound_nonneg [OF _ der_gP])
-    have "?D integrable_on {x \<in> S. 0 < ?D x}"
-      using Dgt
-      by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS)
-    then show "?D integrable_on ?P"
-      apply (rule integrable_spike_set)
-      by (auto simp: zero_less_mult_iff empty_imp_negligible)
-  qed auto
+    have "f integrable_on g ` ?P \<and> integral (g ` ?P) f \<le> integral ?P ?D"
+    proof (rule integral_on_image_ubound_nonneg [OF _ der_gP])
+      show "?D integrable_on ?P"
+      proof (rule integrable_spike_set)
+        show "?D integrable_on {x \<in> S. 0 < ?D x}"
+          using Dgt
+          by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS)
+      qed (auto simp: zero_less_mult_iff empty_imp_negligible)
+    qed auto
   then have "f integrable_on g ` ?P"
     by metis
   moreover have "g ` ?P = {y \<in> g ` S. f y > 0}"
@@ -2730,8 +2672,7 @@
         if "x \<in> T" for x
       proof -
         have "matrix (h' x) ** matrix (g' (h x)) = mat 1"
-          using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear
-          by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+
+          by (metis der_g der_h gh has_derivative_linear local.id matrix_compose matrix_id_mat_1 that)
         then have "\<bar>det (matrix (h' x))\<bar> * \<bar>det (matrix (g' (h x)))\<bar> = 1"
           by (metis abs_1 abs_mult det_I det_mul)
         then show ?thesis
@@ -2740,10 +2681,7 @@
       have "?D integrable_on (h ` T)"
       proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real)
         show "(\<lambda>x. ?fgh x) absolutely_integrable_on T"
-        proof (subst absolutely_integrable_on_iff_nonneg)
-          show "(\<lambda>x. ?fgh x) integrable_on T"
-            using ddf fT integrable_eq by force
-        qed (simp add: zero_le_mult_iff f0 gh)
+          by (smt (verit, del_insts) abs_absolutely_integrableI_1 ddf f0 fT integrable_eq)
       qed (use der_h in auto)
       with Seq show "(\<lambda>x. ?D x) integrable_on S"
         by simp
@@ -2755,9 +2693,8 @@
       qed (use f0 gh der_h in auto)
       also have "\<dots> = integral T f"
         by (force simp: ddf intro: integral_cong)
-      also have "\<dots> \<le> b"
-        by (rule intf)
-      finally show "integral S (\<lambda>x. ?D x) \<le> b" .
+      finally show "integral S (\<lambda>x. ?D x) \<le> b"
+        using intf by linarith 
     qed
   next
     assume R: ?rhs
@@ -2850,15 +2787,16 @@
       using "-" [of "integral {x \<in> S. f(g x) < 0} ?DN"] aN
       by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff)
     have faN: "f absolutely_integrable_on {y \<in> T. f y < 0}"
-      apply (rule absolutely_integrable_integrable_bound [where g = "\<lambda>x. - f x"])
-      using fN by (auto simp: integrable_neg_iff)
+    proof (rule absolutely_integrable_integrable_bound)
+      show "(\<lambda>x. - f x) integrable_on {y \<in> T. f y < 0}"
+        using fN by (auto simp: integrable_neg_iff)
+    qed (use fN in auto)
     have fP: "f integrable_on {y \<in> T. f y > 0}"
              "integral {y \<in> T. f y > 0} f = integral {x \<in> S. f (g x) > 0} ?DP"
       using "+" [of "integral {x \<in> S. f(g x) > 0} ?DP"] aP
       by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff)
     have faP: "f absolutely_integrable_on {y \<in> T. f y > 0}"
-      apply (rule absolutely_integrable_integrable_bound [where g = f])
-      using fP by auto
+      using fP(1) nonnegative_absolutely_integrable_1 by fastforce
     have fa: "f absolutely_integrable_on ({y \<in> T. f y < 0} \<union> {y \<in> T. f y > 0})"
       by (rule absolutely_integrable_Un [OF faN faP])
     show ?rhs
@@ -2904,8 +2842,7 @@
     have aint: "f absolutely_integrable_on {y. y \<in> T \<and> 0 < (f y)}"
                "f absolutely_integrable_on {y. y \<in> T \<and> (f y) < 0}"
          and intT: "integral T f = b"
-      using set_integrable_subset [of _ T] TP TN RHS
-      by blast+
+      using set_integrable_subset [of _ T] TP TN RHS by blast+
     show ?lhs
     proof
       have fN: "f integrable_on {v \<in> T. f v < 0}"
@@ -3162,10 +3099,7 @@
             by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI)
         qed auto
       qed
-    next
-      show "(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D"
-        by (rule DU)
-    qed
+    qed (use DU in metis)
   next
     assume fs: "f absolutely_integrable_on (\<Union>x. g ` F x)"
       and b: "b = integral ((\<Union>x. g ` F x)) f"
@@ -3194,8 +3128,7 @@
           unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN)
         { fix n::nat
           have "(norm \<circ> f) absolutely_integrable_on (\<Union>m\<le>n. g ` F m)"
-            apply (rule absolutely_integrable_norm)
-            using fgU by blast
+            using absolutely_integrable_norm fgU by blast
           then have "integral (?U n) (norm \<circ> ?D) = integral (g ` ?U n) (norm \<circ> f)"
             using iff [of n "?lift \<circ> norm \<circ> f" "integral (g ` ?U n) (?lift \<circ> norm \<circ> f)"]
             unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def)
@@ -3233,11 +3166,9 @@
       unfolding absolutely_integrable_restrict_UNIV by simp
     show "integral (\<Union>n. F n) ?D = integral ((\<Union>x. g ` F x)) f"
     proof (rule LIMSEQ_unique)
-      show "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>x. g ` F x) f"
-        by (rule fgU)
       show "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D"
         unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb])
-    qed
+    qed (use fgU in metis)
   qed
 qed
 
--- a/src/HOL/Analysis/Complex_Transcendental.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Complex_Transcendental.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -75,7 +75,7 @@
 lemma continuous_within_exp:
   fixes z::"'a::{real_normed_field,banach}"
   shows "continuous (at z within s) exp"
-by (simp add: continuous_at_imp_continuous_within)
+  by (simp add: continuous_at_imp_continuous_within)
 
 lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s"
   by (simp add: field_differentiable_within_exp holomorphic_on_def)
@@ -153,7 +153,7 @@
 
 theorem Euler: "exp(z) = of_real(exp(Re z)) *
               (of_real(cos(Im z)) + \<i> * of_real(sin(Im z)))"
-by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real Complex_eq)
+  by (simp add: Complex_eq cis.code exp_eq_polar)
 
 lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
   by (simp add: sin_exp_eq field_simps Re_divide Im_exp)
@@ -568,8 +568,8 @@
 
 lemma cos_eq:
    "cos x = cos y \<longleftrightarrow> (\<exists>n \<in> \<int>. x = y + 2*n*pi \<or> x = -y + 2*n*pi)"
-  using complex_cos_eq [of x y]
-  by (simp only: cos_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff)
+  using complex_cos_eq [of x y] unfolding cos_of_real 
+  by (metis Re_complex_of_real of_real_add of_real_minus)
 
 lemma sinh_complex:
   fixes z :: complex
@@ -635,7 +635,7 @@
 lemma norm_cos_plus1_le:
   fixes z::complex
   shows "norm(1 + cos z) \<le> 2 * exp(norm z)"
-  by (smt (verit, best) exp_ge_add_one_self norm_cos_le norm_ge_zero norm_one norm_triangle_ineq)
+  by (metis mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono one_le_exp_iff)
 
 lemma sinh_conv_sin: "sinh z = -\<i> * sin (\<i>*z)"
   by (simp add: sinh_field_def sin_i_times exp_minus)
@@ -3854,8 +3854,7 @@
              \<i> * (of_nat k * (of_real pi * 2)) + \<i> * (of_int z * (of_nat n * (of_real pi * 2)))) \<longleftrightarrow> j mod n = k mod n" .
    note * = this
   show ?thesis
-    using assms
-    by (simp add: exp_eq field_split_simps *)
+    using assms by (simp add: exp_eq field_split_simps *)
 qed
 
 corollary bij_betw_roots_unity:
--- a/src/HOL/Analysis/Connected.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Connected.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -20,23 +20,12 @@
       e1 \<inter> e2 = {} \<and>
       e1 \<noteq> {} \<and>
       e2 \<noteq> {})"
-  unfolding connected_def openin_open
-  by safe blast+
+  using connected_openin by blast
 
 lemma exists_diff:
   fixes P :: "'a set \<Rightarrow> bool"
   shows "(\<exists>S. P (- S)) \<longleftrightarrow> (\<exists>S. P S)"
-    (is "?lhs \<longleftrightarrow> ?rhs")
-proof -
-  have ?rhs if ?lhs
-    using that by blast
-  moreover have "P (- (- S))" if "P S" for S
-  proof -
-    have "S = - (- S)" by simp
-    with that show ?thesis by metis
-  qed
-  ultimately show ?thesis by metis
-qed
+  by (metis boolean_algebra_class.boolean_algebra.double_compl)
 
 lemma connected_clopen: "connected S \<longleftrightarrow>
   (\<forall>T. openin (top_of_set S) T \<and>
@@ -80,10 +69,10 @@
   by (auto simp: connected_component_def)
 
 lemma connected_component_refl: "x \<in> S \<Longrightarrow> connected_component S x x"
-  by (auto simp: connected_component_def) (use connected_sing in blast)
-
+  using connected_component_def connected_sing by blast
+ 
 lemma connected_component_refl_eq [simp]: "connected_component S x x \<longleftrightarrow> x \<in> S"
-  by (auto simp: connected_component_refl) (auto simp: connected_component_def)
+  using connected_component_in connected_component_refl by blast
 
 lemma connected_component_sym: "connected_component S x y \<Longrightarrow> connected_component S y x"
   by (auto simp: connected_component_def)
@@ -105,23 +94,7 @@
 
 lemma connected_iff_eq_connected_component_set:
   "connected S \<longleftrightarrow> (\<forall>x \<in> S. connected_component_set S x = S)"
-proof (cases "S = {}")
-  case True
-  then show ?thesis by simp
-next
-  case False
-  then obtain x where "x \<in> S" by auto
-  show ?thesis
-  proof
-    assume "connected S"
-    then show "\<forall>x \<in> S. connected_component_set S x = S"
-      by (force simp: connected_component_def)
-  next
-    assume "\<forall>x \<in> S. connected_component_set S x = S"
-    then show "connected S"
-      by (metis \<open>x \<in> S\<close> connected_connected_component)
-  qed
-qed
+  by (metis connected_component_def connected_component_in connected_connected_component mem_Collect_eq subsetI subset_antisym)
 
 lemma connected_component_subset: "connected_component_set S x \<subseteq> S"
   using connected_component_in by blast
@@ -165,34 +138,27 @@
     unfolding closure_eq [symmetric]
   proof
     show "closure (connected_component_set S x) \<subseteq> connected_component_set S x"
-      apply (rule connected_component_maximal)
-        apply (simp add: closure_def True)
-       apply (simp add: connected_imp_connected_closure)
-      apply (simp add: S closure_minimal connected_component_subset)
-      done
-  next
-    show "connected_component_set S x \<subseteq> closure (connected_component_set S x)"
-      by (simp add: closure_subset)
-  qed
+    proof (rule connected_component_maximal)
+      show "x \<in> closure (connected_component_set S x)"
+        by (simp add: closure_def True)
+      show "connected (closure (connected_component_set S x))"
+        by (simp add: connected_imp_connected_closure)
+      show "closure (connected_component_set S x) \<subseteq> S"
+        by (simp add: S closure_minimal connected_component_subset)
+    qed  
+  qed (simp add: closure_subset)
 qed
 
 lemma connected_component_disjoint:
   "connected_component_set S a \<inter> connected_component_set S b = {} \<longleftrightarrow>
     a \<notin> connected_component_set S b"
-  apply (auto simp: connected_component_eq)
-  using connected_component_eq connected_component_sym
-  apply blast
-  done
+  by (smt (verit) connected_component_eq connected_component_eq_empty connected_component_refl_eq
+      disjoint_iff_not_equal mem_Collect_eq)
 
 lemma connected_component_nonoverlap:
   "connected_component_set S a \<inter> connected_component_set S b = {} \<longleftrightarrow>
     a \<notin> S \<or> b \<notin> S \<or> connected_component_set S a \<noteq> connected_component_set S b"
-  apply (auto simp: connected_component_in)
-  using connected_component_refl_eq
-    apply blast
-   apply (metis connected_component_eq mem_Collect_eq)
-  apply (metis connected_component_eq mem_Collect_eq)
-  done
+  by (metis connected_component_disjoint connected_component_eq connected_component_eq_empty inf.idem)
 
 lemma connected_component_overlap:
   "connected_component_set S a \<inter> connected_component_set S b \<noteq> {} \<longleftrightarrow>
@@ -205,13 +171,8 @@
 lemma connected_component_eq_eq:
   "connected_component_set S x = connected_component_set S y \<longleftrightarrow>
     x \<notin> S \<and> y \<notin> S \<or> x \<in> S \<and> y \<in> S \<and> connected_component S x y"
-  apply (cases "y \<in> S", simp)
-   apply (metis connected_component_eq connected_component_eq_empty connected_component_refl_eq mem_Collect_eq)
-  apply (cases "x \<in> S", simp)
-   apply (metis connected_component_eq_empty)
-  using connected_component_eq_empty
-  apply blast
-  done
+  by (metis connected_component_eq connected_component_eq_empty connected_component_refl mem_Collect_eq)
+
 
 lemma connected_iff_connected_component_eq:
   "connected S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. connected_component_set S x = connected_component_set S y)"
@@ -219,19 +180,13 @@
 
 lemma connected_component_idemp:
   "connected_component_set (connected_component_set S x) x = connected_component_set S x"
-  apply (rule subset_antisym)
-   apply (simp add: connected_component_subset)
-  apply (metis connected_component_eq_empty connected_component_maximal
-      connected_component_refl_eq connected_connected_component mem_Collect_eq set_eq_subset)
-  done
+  by (metis Int_absorb connected_component_disjoint connected_component_eq_empty connected_component_eq_self connected_connected_component)
 
 lemma connected_component_unique:
   "\<lbrakk>x \<in> c; c \<subseteq> S; connected c;
     \<And>c'. \<lbrakk>x \<in> c'; c' \<subseteq> S; connected c'\<rbrakk> \<Longrightarrow> c' \<subseteq> c\<rbrakk>
         \<Longrightarrow> connected_component_set S x = c"
-  apply (rule subset_antisym)
-   apply (meson connected_component_maximal connected_component_subset connected_connected_component contra_subsetD)
-  by (simp add: connected_component_maximal)
+  by (meson connected_component_maximal connected_component_subset connected_connected_component subsetD subset_antisym)
 
 lemma joinable_connected_component_eq:
   "\<lbrakk>connected T; T \<subseteq> S;
@@ -241,25 +196,27 @@
   by (metis (full_types) subsetD connected_component_eq connected_component_maximal disjoint_iff_not_equal)
 
 lemma Union_connected_component: "\<Union>(connected_component_set S ` S) = S"
-  apply (rule subset_antisym)
-  apply (simp add: SUP_least connected_component_subset)
-  using connected_component_refl_eq
-  by force
-
+proof
+  show "\<Union>(connected_component_set S ` S) \<subseteq> S"
+    by (simp add: SUP_least connected_component_subset)
+qed (use connected_component_refl_eq in force)
 
 lemma complement_connected_component_unions:
     "S - connected_component_set S x =
      \<Union>(connected_component_set S ` S - {connected_component_set S x})"
-  apply (subst Union_connected_component [symmetric], auto)
-  apply (metis connected_component_eq_eq connected_component_in)
-  by (metis connected_component_eq mem_Collect_eq)
+    (is "?lhs = ?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    by (metis Diff_subset Diff_subset_conv Sup_insert Union_connected_component insert_Diff_single)
+  show "?rhs \<subseteq> ?lhs"
+    by clarsimp (metis connected_component_eq_eq connected_component_in)
+qed
 
 lemma connected_component_intermediate_subset:
         "\<lbrakk>connected_component_set U a \<subseteq> T; T \<subseteq> U\<rbrakk>
         \<Longrightarrow> connected_component_set T a = connected_component_set U a"
   by (metis connected_component_idemp connected_component_mono subset_antisym)
 
-
 lemma connected_component_homeomorphismI:
   assumes "homeomorphism A B f g" "connected_component A x y"
   shows   "connected_component B (f x) (f y)"
@@ -304,148 +261,121 @@
   obtains x where "x \<in> U" "S = connected_component_set U x"
   using assms by (auto simp: components_def)
 
-lemma Union_components [simp]: "\<Union>(components u) = u"
-  apply (rule subset_antisym)
-  using Union_connected_component components_def apply fastforce
-  apply (metis Union_connected_component components_def set_eq_subset)
-  done
+lemma Union_components [simp]: "\<Union>(components U) = U"
+  by (simp add: Union_connected_component components_def)
 
-lemma pairwise_disjoint_components: "pairwise (\<lambda>X Y. X \<inter> Y = {}) (components u)"
-  apply (simp add: pairwise_def)
-  apply (auto simp: components_iff)
-  apply (metis connected_component_eq_eq connected_component_in)+
-  done
+lemma pairwise_disjoint_components: "pairwise (\<lambda>X Y. X \<inter> Y = {}) (components U)"
+  unfolding pairwise_def
+  by (metis (full_types) components_iff connected_component_nonoverlap)
 
-lemma in_components_nonempty: "c \<in> components s \<Longrightarrow> c \<noteq> {}"
+lemma in_components_nonempty: "C \<in> components S \<Longrightarrow> C \<noteq> {}"
     by (metis components_iff connected_component_eq_empty)
 
-lemma in_components_subset: "c \<in> components s \<Longrightarrow> c \<subseteq> s"
+lemma in_components_subset: "C \<in> components S \<Longrightarrow> C \<subseteq> S"
   using Union_components by blast
 
-lemma in_components_connected: "c \<in> components s \<Longrightarrow> connected c"
+lemma in_components_connected: "C \<in> components S \<Longrightarrow> connected C"
   by (metis components_iff connected_connected_component)
 
 lemma in_components_maximal:
-  "c \<in> components s \<longleftrightarrow>
-    c \<noteq> {} \<and> c \<subseteq> s \<and> connected c \<and> (\<forall>d. d \<noteq> {} \<and> c \<subseteq> d \<and> d \<subseteq> s \<and> connected d \<longrightarrow> d = c)"
-  apply (rule iffI)
-   apply (simp add: in_components_nonempty in_components_connected)
-   apply (metis (full_types) components_iff connected_component_eq_self connected_component_intermediate_subset connected_component_refl in_components_subset mem_Collect_eq rev_subsetD)
-  apply (metis bot.extremum_uniqueI components_iff connected_component_eq_empty connected_component_maximal connected_component_subset connected_connected_component subset_emptyI)
-  done
+  "C \<in> components S \<longleftrightarrow>
+    C \<noteq> {} \<and> C \<subseteq> S \<and> connected C \<and> (\<forall>d. d \<noteq> {} \<and> C \<subseteq> d \<and> d \<subseteq> S \<and> connected d \<longrightarrow> d = C)"
+(is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume L: ?lhs
+  then have "C \<subseteq> S" "connected C"
+    by (simp_all add: in_components_subset in_components_connected)
+  then show ?rhs
+    by (metis (full_types) L components_iff connected_component_maximal connected_component_refl empty_iff mem_Collect_eq subsetD subset_antisym)
+next
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (metis bot.extremum_uniqueI components_iff connected_component_eq_empty connected_component_maximal connected_component_subset connected_connected_component subset_emptyI)
+qed
+
 
 lemma joinable_components_eq:
-  "connected t \<and> t \<subseteq> s \<and> c1 \<in> components s \<and> c2 \<in> components s \<and> c1 \<inter> t \<noteq> {} \<and> c2 \<inter> t \<noteq> {} \<Longrightarrow> c1 = c2"
+  "connected T \<and> T \<subseteq> S \<and> c1 \<in> components S \<and> c2 \<in> components S \<and> c1 \<inter> T \<noteq> {} \<and> c2 \<inter> T \<noteq> {} \<Longrightarrow> c1 = c2"
   by (metis (full_types) components_iff joinable_connected_component_eq)
 
-lemma closed_components: "\<lbrakk>closed s; c \<in> components s\<rbrakk> \<Longrightarrow> closed c"
+lemma closed_components: "\<lbrakk>closed S; C \<in> components S\<rbrakk> \<Longrightarrow> closed C"
   by (metis closed_connected_component components_iff)
 
 lemma components_nonoverlap:
-    "\<lbrakk>c \<in> components s; c' \<in> components s\<rbrakk> \<Longrightarrow> (c \<inter> c' = {}) \<longleftrightarrow> (c \<noteq> c')"
-  apply (auto simp: in_components_nonempty components_iff)
-    using connected_component_refl apply blast
-   apply (metis connected_component_eq_eq connected_component_in)
-  by (metis connected_component_eq mem_Collect_eq)
+    "\<lbrakk>C \<in> components S; C' \<in> components S\<rbrakk> \<Longrightarrow> (C \<inter> C' = {}) \<longleftrightarrow> (C \<noteq> C')"
+  by (metis (full_types) components_iff connected_component_nonoverlap)
 
-lemma components_eq: "\<lbrakk>c \<in> components s; c' \<in> components s\<rbrakk> \<Longrightarrow> (c = c' \<longleftrightarrow> c \<inter> c' \<noteq> {})"
+lemma components_eq: "\<lbrakk>C \<in> components S; C' \<in> components S\<rbrakk> \<Longrightarrow> (C = C' \<longleftrightarrow> C \<inter> C' \<noteq> {})"
   by (metis components_nonoverlap)
 
-lemma components_eq_empty [simp]: "components s = {} \<longleftrightarrow> s = {}"
+lemma components_eq_empty [simp]: "components S = {} \<longleftrightarrow> S = {}"
   by (simp add: components_def)
 
 lemma components_empty [simp]: "components {} = {}"
   by simp
 
-lemma connected_eq_connected_components_eq: "connected s \<longleftrightarrow> (\<forall>c \<in> components s. \<forall>c' \<in> components s. c = c')"
+lemma connected_eq_connected_components_eq: "connected S \<longleftrightarrow> (\<forall>C \<in> components S. \<forall>C' \<in> components S. C = C')"
   by (metis (no_types, opaque_lifting) components_iff connected_component_eq_eq connected_iff_connected_component)
 
-lemma components_eq_sing_iff: "components s = {s} \<longleftrightarrow> connected s \<and> s \<noteq> {}"
-  apply (rule iffI)
-  using in_components_connected apply fastforce
-  apply safe
-  using Union_components apply fastforce
-   apply (metis components_iff connected_component_eq_self)
-  using in_components_maximal
-  apply auto
-  done
+lemma components_eq_sing_iff: "components S = {S} \<longleftrightarrow> connected S \<and> S \<noteq> {}" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (metis components_iff connected_component_eq_self equals0I insert_iff mk_disjoint_insert)
+qed (use in_components_connected in fastforce)
 
-lemma components_eq_sing_exists: "(\<exists>a. components s = {a}) \<longleftrightarrow> connected s \<and> s \<noteq> {}"
-  apply (rule iffI)
-  using connected_eq_connected_components_eq apply fastforce
-  apply (metis components_eq_sing_iff)
-  done
+lemma components_eq_sing_exists: "(\<exists>a. components S = {a}) \<longleftrightarrow> connected S \<and> S \<noteq> {}"
+  by (metis Union_components ccpo_Sup_singleton components_eq_sing_iff)
 
-lemma connected_eq_components_subset_sing: "connected s \<longleftrightarrow> components s \<subseteq> {s}"
-  by (metis Union_components components_empty components_eq_sing_iff connected_empty insert_subset order_refl subset_singletonD)
+lemma connected_eq_components_subset_sing: "connected S \<longleftrightarrow> components S \<subseteq> {S}"
+  by (metis components_eq_empty components_eq_sing_iff connected_empty subset_singleton_iff)
 
-lemma connected_eq_components_subset_sing_exists: "connected s \<longleftrightarrow> (\<exists>a. components s \<subseteq> {a})"
-  by (metis components_eq_sing_exists connected_eq_components_subset_sing empty_iff subset_iff subset_singletonD)
+lemma connected_eq_components_subset_sing_exists: "connected S \<longleftrightarrow> (\<exists>a. components S \<subseteq> {a})"
+  by (metis components_eq_sing_exists connected_eq_components_subset_sing subset_singleton_iff)
 
-lemma in_components_self: "s \<in> components s \<longleftrightarrow> connected s \<and> s \<noteq> {}"
+lemma in_components_self: "S \<in> components S \<longleftrightarrow> connected S \<and> S \<noteq> {}"
   by (metis components_empty components_eq_sing_iff empty_iff in_components_connected insertI1)
 
-lemma components_maximal: "\<lbrakk>c \<in> components s; connected t; t \<subseteq> s; c \<inter> t \<noteq> {}\<rbrakk> \<Longrightarrow> t \<subseteq> c"
-  apply (simp add: components_def ex_in_conv [symmetric], clarify)
-  by (meson connected_component_def connected_component_trans)
+lemma components_maximal: "\<lbrakk>C \<in> components S; connected T; T \<subseteq> S; C \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow> T \<subseteq> C"
+  by (smt (verit, best) Int_Un_eq(4) Un_upper1 bot_eq_sup_iff connected_Un in_components_maximal inf.orderE sup.mono sup.orderI)
 
-lemma exists_component_superset: "\<lbrakk>t \<subseteq> s; s \<noteq> {}; connected t\<rbrakk> \<Longrightarrow> \<exists>c. c \<in> components s \<and> t \<subseteq> c"
-  apply (cases "t = {}", force)
-  apply (metis components_def ex_in_conv connected_component_maximal contra_subsetD image_eqI)
-  done
+lemma exists_component_superset: "\<lbrakk>T \<subseteq> S; S \<noteq> {}; connected T\<rbrakk> \<Longrightarrow> \<exists>C. C \<in> components S \<and> T \<subseteq> C"
+  by (meson componentsI connected_component_maximal equals0I subset_eq)
 
-lemma components_intermediate_subset: "\<lbrakk>s \<in> components u; s \<subseteq> t; t \<subseteq> u\<rbrakk> \<Longrightarrow> s \<in> components t"
-  apply (auto simp: components_iff)
-  apply (metis connected_component_eq_empty connected_component_intermediate_subset)
-  done
+lemma components_intermediate_subset: "\<lbrakk>S \<in> components U; S \<subseteq> T; T \<subseteq> U\<rbrakk> \<Longrightarrow> S \<in> components T"
+  by (smt (verit, best) dual_order.trans in_components_maximal)
 
-lemma in_components_unions_complement: "c \<in> components s \<Longrightarrow> s - c = \<Union>(components s - {c})"
+lemma in_components_unions_complement: "C \<in> components S \<Longrightarrow> S - C = \<Union>(components S - {C})"
   by (metis complement_connected_component_unions components_def components_iff)
 
 lemma connected_intermediate_closure:
-  assumes cs: "connected s" and st: "s \<subseteq> t" and ts: "t \<subseteq> closure s"
-  shows "connected t"
-proof (rule connectedI)
-  fix A B
-  assume A: "open A" and B: "open B" and Alap: "A \<inter> t \<noteq> {}" and Blap: "B \<inter> t \<noteq> {}"
-    and disj: "A \<inter> B \<inter> t = {}" and cover: "t \<subseteq> A \<union> B"
-  have disjs: "A \<inter> B \<inter> s = {}"
-    using disj st by auto
-  have "A \<inter> closure s \<noteq> {}"
-    using Alap Int_absorb1 ts by blast
-  then have Alaps: "A \<inter> s \<noteq> {}"
-    by (simp add: A open_Int_closure_eq_empty)
-  have "B \<inter> closure s \<noteq> {}"
-    using Blap Int_absorb1 ts by blast
-  then have Blaps: "B \<inter> s \<noteq> {}"
-    by (simp add: B open_Int_closure_eq_empty)
-  then show False
-    using cs [unfolded connected_def] A B disjs Alaps Blaps cover st
-    by blast
-qed
+  assumes cs: "connected S" and st: "S \<subseteq> T" and ts: "T \<subseteq> closure S"
+  shows "connected T"
+  using assms unfolding connected_def
+  by (smt (verit) Int_assoc inf.absorb_iff2 inf_bot_left open_Int_closure_eq_empty)
 
-lemma closedin_connected_component: "closedin (top_of_set s) (connected_component_set s x)"
-proof (cases "connected_component_set s x = {}")
+lemma closedin_connected_component: "closedin (top_of_set S) (connected_component_set S x)"
+proof (cases "connected_component_set S x = {}")
   case True
   then show ?thesis
     by (metis closedin_empty)
 next
   case False
-  then obtain y where y: "connected_component s x y"
-    by blast
-  have *: "connected_component_set s x \<subseteq> s \<inter> closure (connected_component_set s x)"
+  then obtain y where y: "connected_component S x y" and "x \<in> S"
+    using connected_component_eq_empty by blast
+  have *: "connected_component_set S x \<subseteq> S \<inter> closure (connected_component_set S x)"
     by (auto simp: closure_def connected_component_in)
-  have "connected_component s x y \<Longrightarrow> s \<inter> closure (connected_component_set s x) \<subseteq> connected_component_set s x"
-    apply (rule connected_component_maximal, simp)
-    using closure_subset connected_component_in apply fastforce
-    using * connected_intermediate_closure apply blast+
-    done
+  have **: "x \<in> closure (connected_component_set S x)"
+    by (simp add: \<open>x \<in> S\<close> closure_def)
+  have "S \<inter> closure (connected_component_set S x) \<subseteq> connected_component_set S x" if "connected_component S x y"
+  proof (rule connected_component_maximal)
+    show "connected (S \<inter> closure (connected_component_set S x))"
+      using "*" connected_intermediate_closure by blast
+  qed (use \<open>x \<in> S\<close> ** in auto)
   with y * show ?thesis
     by (auto simp: closedin_closed)
 qed
 
 lemma closedin_component:
-   "C \<in> components s \<Longrightarrow> closedin (top_of_set s) C"
+   "C \<in> components S \<Longrightarrow> closedin (top_of_set S) C"
   using closedin_connected_component componentsE by blast
 
 
@@ -454,37 +384,34 @@
 
 lemma continuous_levelset_openin_cases:
   fixes f :: "_ \<Rightarrow> 'b::t1_space"
-  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
-        openin (top_of_set s) {x \<in> s. f x = a}
-        \<Longrightarrow> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"
+  shows "connected S \<Longrightarrow> continuous_on S f \<Longrightarrow>
+        openin (top_of_set S) {x \<in> S. f x = a}
+        \<Longrightarrow> (\<forall>x \<in> S. f x \<noteq> a) \<or> (\<forall>x \<in> S. f x = a)"
   unfolding connected_clopen
   using continuous_closedin_preimage_constant by auto
 
 lemma continuous_levelset_openin:
   fixes f :: "_ \<Rightarrow> 'b::t1_space"
-  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
-        openin (top_of_set s) {x \<in> s. f x = a} \<Longrightarrow>
-        (\<exists>x \<in> s. f x = a)  \<Longrightarrow> (\<forall>x \<in> s. f x = a)"
-  using continuous_levelset_openin_cases[of s f ]
+  shows "connected S \<Longrightarrow> continuous_on S f \<Longrightarrow>
+        openin (top_of_set S) {x \<in> S. f x = a} \<Longrightarrow>
+        (\<exists>x \<in> S. f x = a)  \<Longrightarrow> (\<forall>x \<in> S. f x = a)"
+  using continuous_levelset_openin_cases[of S f ]
   by meson
 
 lemma continuous_levelset_open:
   fixes f :: "_ \<Rightarrow> 'b::t1_space"
-  assumes "connected s"
-    and "continuous_on s f"
-    and "open {x \<in> s. f x = a}"
-    and "\<exists>x \<in> s.  f x = a"
-  shows "\<forall>x \<in> s. f x = a"
-  using continuous_levelset_openin[OF assms(1,2), of a, unfolded openin_open]
-  using assms (3,4)
+  assumes S: "connected S" "continuous_on S f"
+    and a: "open {x \<in> S. f x = a}" "\<exists>x \<in> S.  f x = a"
+  shows "\<forall>x \<in> S. f x = a"
+  using a continuous_levelset_openin[OF S, of a, unfolded openin_open]
   by fast
 
 
 subsection\<^marker>\<open>tag unimportant\<close> \<open>Preservation of Connectedness\<close>
 
 lemma homeomorphic_connectedness:
-  assumes "s homeomorphic t"
-  shows "connected s \<longleftrightarrow> connected t"
+  assumes "S homeomorphic T"
+  shows "connected S \<longleftrightarrow> connected T"
 using assms unfolding homeomorphic_def homeomorphism_def by (metis connected_continuous_image)
 
 lemma connected_monotone_quotient_preimage:
@@ -539,14 +466,7 @@
   show ?thesis
   proof (rule connected_monotone_quotient_preimage [OF \<open>connected C\<close> contf' eqC])
     show "connected (S \<inter> f -` C \<inter> f -` {y})" if "y \<in> C" for y
-    proof -
-      have "S \<inter> f -` C \<inter> f -` {y} = S \<inter> f -` {y}"
-        using that by blast
-      moreover have "connected (S \<inter> f -` {y})"
-        using \<open>C \<subseteq> T\<close> connT that by blast
-      ultimately show ?thesis
-        by metis
-    qed
+      by (metis Int_assoc Int_empty_right Int_insert_right_if1 assms(6) connT in_mono that vimage_Int)
     have "\<And>U. openin (top_of_set (S \<inter> f -` C)) U
                \<Longrightarrow> openin (top_of_set C) (f ` U)"
       using open_map_restrict [OF _ ST \<open>C \<subseteq> T\<close>] by metis
@@ -572,14 +492,7 @@
   show ?thesis
   proof (rule connected_monotone_quotient_preimage [OF \<open>connected C\<close> contf' eqC])
     show "connected (S \<inter> f -` C \<inter> f -` {y})" if "y \<in> C" for y
-    proof -
-      have "S \<inter> f -` C \<inter> f -` {y} = S \<inter> f -` {y}"
-        using that by blast
-      moreover have "connected (S \<inter> f -` {y})"
-        using \<open>C \<subseteq> T\<close> connT that by blast
-      ultimately show ?thesis
-        by metis
-    qed
+      by (metis Int_assoc Int_empty_right Int_insert_right_if1 \<open>C \<subseteq> T\<close> connT subsetD that vimage_Int)
     have "\<And>U. closedin (top_of_set (S \<inter> f -` C)) U
                \<Longrightarrow> closedin (top_of_set C) (f ` U)"
       using closed_map_restrict [OF _ ST \<open>C \<subseteq> T\<close>] by metis
@@ -615,15 +528,12 @@
     then have clo: "closedin (top_of_set S) (S \<inter> H1)"
                    "closedin (top_of_set S) (S \<inter> H2)"
       by (metis Un_upper1 closedin_closed_subset inf_commute)+
-    have Seq: "S \<inter> (H1 \<union> H2) = S"
-      by (simp add: H)
-    have "S \<inter> ((S \<union> T) \<inter> H1) \<union> S \<inter> ((S \<union> T) \<inter> H2) = S"
-      using Seq by auto
+    moreover have "S \<inter> ((S \<union> T) \<inter> H1) \<union> S \<inter> ((S \<union> T) \<inter> H2) = S"
+      using H by blast
     moreover have "H1 \<inter> (S \<inter> ((S \<union> T) \<inter> H2)) = {}"
       using H by blast
     ultimately have "S \<inter> H1 = {} \<or> S \<inter> H2 = {}"
-      by (metis (no_types) H Int_assoc \<open>S \<inter> (H1 \<union> H2) = S\<close> \<open>connected S\<close>
-          clo Seq connected_closedin inf_bot_right inf_le1)
+      by (smt (verit) Int_assoc \<open>connected S\<close> connected_closedin_eq inf_commute inf_sup_absorb)
     then show "S \<subseteq> H1 \<or> S \<subseteq> H2"
       using H \<open>connected S\<close> unfolding connected_closedin by blast
   next
@@ -641,8 +551,7 @@
     have "openin (top_of_set ((U - S) \<union> (S \<union> T))) H2"
     proof (rule openin_subtopology_Un)
       show "openin (top_of_set (S \<union> T)) H2"
-        using \<open>H2 \<subseteq> T\<close> apply (auto simp: openin_closedin_eq)
-        by (metis Diff_Diff_Int Diff_disjoint Diff_partition Diff_subset H Int_absorb1 Un_Diff)
+        by (metis Diff_cancel H Un_Diff Un_Diff_Int closedin_subset openin_closedin_eq topspace_euclidean_subtopology)
       then show "openin (top_of_set (U - S)) H2"
         by (meson H2T Un_upper2 opeT openin_subset_trans openin_trans)
     qed
@@ -652,8 +561,7 @@
         using H H2T cloT closedin_subset_trans 
         by (blast intro: closedin_subtopology_Un closedin_trans)
     qed (simp add: H)
-    ultimately
-    have H2: "H2 = {} \<or> H2 = U"
+    ultimately have H2: "H2 = {} \<or> H2 = U"
       using Ueq \<open>connected U\<close> unfolding connected_clopen by metis   
     then have "H2 \<subseteq> S"
       by (metis Diff_partition H Un_Diff_cancel Un_subset_iff \<open>H2 \<subseteq> T\<close> assms(3) inf.orderE opeT openin_imp_subset)
@@ -674,7 +582,7 @@
   fix H3 H4 
   assume clo3: "closedin (top_of_set (U - C)) H3" 
     and clo4: "closedin (top_of_set (U - C)) H4" 
-    and "H3 \<union> H4 = U - C" and "H3 \<inter> H4 = {}" and "H3 \<noteq> {}" and "H4 \<noteq> {}"
+    and H34: "H3 \<union> H4 = U - C" "H3 \<inter> H4 = {}" and "H3 \<noteq> {}" and "H4 \<noteq> {}"
     and * [rule_format]:
     "\<forall>H1 H2. \<not> closedin (top_of_set S) H1 \<or>
                       \<not> closedin (top_of_set S) H2 \<or>
@@ -687,16 +595,12 @@
   have cCH3: "connected (C \<union> H3)"
   proof (rule connected_Un_clopen_in_complement [OF \<open>connected C\<close> \<open>connected U\<close> _ _ clo3])
     show "openin (top_of_set (U - C)) H3"
-      apply (simp add: openin_closedin_eq \<open>H3 \<subseteq> U - C\<close>)
-      apply (simp add: closedin_subtopology)
-      by (metis Diff_cancel Diff_triv Un_Diff clo4 \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> closedin_closed inf_commute sup_bot.left_neutral)
+      by (metis Diff_cancel Un_Diff Un_Diff_Int \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> ope4)
   qed (use clo3 \<open>C \<subseteq> U - S\<close> in auto)
   have cCH4: "connected (C \<union> H4)"
   proof (rule connected_Un_clopen_in_complement [OF \<open>connected C\<close> \<open>connected U\<close> _ _ clo4])
     show "openin (top_of_set (U - C)) H4"
-      apply (simp add: openin_closedin_eq \<open>H4 \<subseteq> U - C\<close>)
-      apply (simp add: closedin_subtopology)
-      by (metis Diff_cancel Int_commute Un_Diff Un_Diff_Int \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> clo3 closedin_closed)
+      by (metis Diff_cancel Diff_triv Int_Un_eq(2) Un_Diff H34 inf_commute ope3)
   qed (use clo4 \<open>C \<subseteq> U - S\<close> in auto)
   have "closedin (top_of_set S) (S \<inter> H3)" "closedin (top_of_set S) (S \<inter> H4)"
     using clo3 clo4 \<open>S \<subseteq> U\<close> \<open>C \<subseteq> U - S\<close> by (auto simp: closedin_closed)
@@ -717,18 +621,17 @@
 lemma continuous_disconnected_range_constant:
   assumes S: "connected S"
       and conf: "continuous_on S f"
-      and fim: "f ` S \<subseteq> t"
-      and cct: "\<And>y. y \<in> t \<Longrightarrow> connected_component_set t y = {y}"
+      and fim: "f \<in> S \<rightarrow> T"
+      and cct: "\<And>y. y \<in> T \<Longrightarrow> connected_component_set T y = {y}"
     shows "f constant_on S"
 proof (cases "S = {}")
   case True then show ?thesis
     by (simp add: constant_on_def)
 next
   case False
-  { fix x assume "x \<in> S"
-    then have "f ` S \<subseteq> {f x}"
-    by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI S cct)
-  }
+  then have "f ` S \<subseteq> {f x}" if "x \<in> S" for x
+    by (metis PiE S cct connected_component_maximal connected_continuous_image [OF conf] fim image_eqI 
+        image_subset_iff that)
   with False show ?thesis
     unfolding constant_on_def by blast
 qed
@@ -740,19 +643,17 @@
               \<lbrakk>continuous_on S f; finite(f ` S)\<rbrakk> \<Longrightarrow> f constant_on S"
     shows "connected S"
 proof -
-  { fix t u
-    assume clt: "closedin (top_of_set S) t"
-       and clu: "closedin (top_of_set S) u"
-       and tue: "t \<inter> u = {}" and tus: "t \<union> u = S"
-    have conif: "continuous_on S (\<lambda>x. if x \<in> t then 0 else 1)"
-      apply (subst tus [symmetric])
-      apply (rule continuous_on_cases_local)
-      using clt clu tue
-      apply (auto simp: tus)
-      done
-    have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` S)"
+  { fix T U
+    assume clt: "closedin (top_of_set S) T"
+       and clu: "closedin (top_of_set S) U"
+       and tue: "T \<inter> U = {}" and tus: "T \<union> U = S"
+    have "continuous_on (T \<union> U) (\<lambda>x. if x \<in> T then 0 else 1)"
+      using clt clu tue by (intro continuous_on_cases_local) (auto simp: tus)
+    then have conif: "continuous_on S (\<lambda>x. if x \<in> T then 0 else 1)"
+      using tus by blast
+    have fi: "finite ((\<lambda>x. if x \<in> T then 0 else 1) ` S)"
       by (rule finite_subset [of _ "{0,1}"]) auto
-    have "t = {} \<or> u = {}"
+    have "T = {} \<or> U = {}"
       using assms [OF conif fi] tus [symmetric]
       by (auto simp: Ball_def constant_on_def) (metis IntI empty_iff one_neq_zero tue)
   }
--- a/src/HOL/Analysis/Convex.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Convex.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -22,7 +22,7 @@
 lemma convexI:
   assumes "\<And>x y u v. x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> 0 \<le> u \<Longrightarrow> 0 \<le> v \<Longrightarrow> u + v = 1 \<Longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s"
   shows "convex s"
-  using assms unfolding convex_def by fast
+  by (simp add: assms convex_def)
 
 lemma convexD:
   assumes "convex s" and "x \<in> s" and "y \<in> s" and "0 \<le> u" and "0 \<le> v" and "u + v = 1"
@@ -31,25 +31,7 @@
 
 lemma convex_alt: "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> ((1 - u) *\<^sub>R x + u *\<^sub>R y) \<in> s)"
   (is "_ \<longleftrightarrow> ?alt")
-proof
-  show "convex s" if alt: ?alt
-  proof -
-    {
-      fix x y and u v :: real
-      assume mem: "x \<in> s" "y \<in> s"
-      assume "0 \<le> u" "0 \<le> v"
-      moreover
-      assume "u + v = 1"
-      then have "u = 1 - v" by auto
-      ultimately have "u *\<^sub>R x + v *\<^sub>R y \<in> s"
-        using alt [rule_format, OF mem] by auto
-    }
-    then show ?thesis
-      unfolding convex_def by auto
-  qed
-  show ?alt if "convex s"
-    using that by (auto simp: convex_def)
-qed
+  by (smt (verit) convexD convexI)
 
 lemma convexD_alt:
   assumes "convex s" "a \<in> s" "b \<in> s" "0 \<le> u" "u \<le> 1"
@@ -193,44 +175,31 @@
   fixes C :: "'a::real_vector set"
   assumes "finite S"
     and "convex C"
-    and "(\<Sum> i \<in> S. a i) = 1"
-  assumes "\<And>i. i \<in> S \<Longrightarrow> a i \<ge> 0"
-    and "\<And>i. i \<in> S \<Longrightarrow> y i \<in> C"
+    and a: "(\<Sum> i \<in> S. a i) = 1" "\<And>i. i \<in> S \<Longrightarrow> a i \<ge> 0"
+    and C: "\<And>i. i \<in> S \<Longrightarrow> y i \<in> C"
   shows "(\<Sum> j \<in> S. a j *\<^sub>R y j) \<in> C"
-  using assms(1,3,4,5)
-proof (induct arbitrary: a set: finite)
+  using \<open>finite S\<close> a C
+proof (induction arbitrary: a set: finite)
   case empty
   then show ?case by simp
 next
-  case (insert i S) note IH = this(3)
-  have "a i + sum a S = 1"
-    and "0 \<le> a i"
-    and "\<forall>j\<in>S. 0 \<le> a j"
-    and "y i \<in> C"
-    and "\<forall>j\<in>S. y j \<in> C"
-    using insert.hyps(1,2) insert.prems by simp_all
+  case (insert i S) 
   then have "0 \<le> sum a S"
     by (simp add: sum_nonneg)
   have "a i *\<^sub>R y i + (\<Sum>j\<in>S. a j *\<^sub>R y j) \<in> C"
   proof (cases "sum a S = 0")
-    case True
-    with \<open>a i + sum a S = 1\<close> have "a i = 1"
-      by simp
-    from sum_nonneg_0 [OF \<open>finite S\<close> _ True] \<open>\<forall>j\<in>S. 0 \<le> a j\<close> have "\<forall>j\<in>S. a j = 0"
-      by simp
-    show ?thesis using \<open>a i = 1\<close> and \<open>\<forall>j\<in>S. a j = 0\<close> and \<open>y i \<in> C\<close>
-      by simp
+    case True with insert show ?thesis
+      by (simp add: sum_nonneg_eq_0_iff)
   next
     case False
     with \<open>0 \<le> sum a S\<close> have "0 < sum a S"
       by simp
     then have "(\<Sum>j\<in>S. (a j / sum a S) *\<^sub>R y j) \<in> C"
-      using \<open>\<forall>j\<in>S. 0 \<le> a j\<close> and \<open>\<forall>j\<in>S. y j \<in> C\<close>
-      by (simp add: IH sum_divide_distrib [symmetric])
-    from \<open>convex C\<close> and \<open>y i \<in> C\<close> and this and \<open>0 \<le> a i\<close>
-      and \<open>0 \<le> sum a S\<close> and \<open>a i + sum a S = 1\<close>
+      using insert
+      by (simp add: insert.IH flip: sum_divide_distrib)
+    with \<open>convex C\<close> insert \<open>0 \<le> sum a S\<close> 
     have "a i *\<^sub>R y i + sum a S *\<^sub>R (\<Sum>j\<in>S. (a j / sum a S) *\<^sub>R y j) \<in> C"
-      by (rule convexD)
+      by (simp add: convex_def)
     then show ?thesis
       by (simp add: scaleR_sum_right False)
   qed
@@ -239,18 +208,13 @@
 qed
 
 lemma convex:
-  "convex S \<longleftrightarrow> (\<forall>(k::nat) u x. (\<forall>i. 1\<le>i \<and> i\<le>k \<longrightarrow> 0 \<le> u i \<and> x i \<in>S) \<and> (sum u {1..k} = 1)
-      \<longrightarrow> sum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> S)"
-proof safe
-  fix k :: nat
-  fix u :: "nat \<Rightarrow> real"
-  fix x
-  assume "convex S"
-    "\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> S"
-    "sum u {1..k} = 1"
-  with convex_sum[of "{1 .. k}" S] show "(\<Sum>j\<in>{1 .. k}. u j *\<^sub>R x j) \<in> S"
-    by auto
-next
+  "convex S \<longleftrightarrow> 
+    (\<forall>(k::nat) u x. (\<forall>i. 1\<le>i \<and> i\<le>k \<longrightarrow> 0 \<le> u i \<and> x i \<in>S) \<and> (sum u {1..k} = 1)
+      \<longrightarrow> sum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> S)"  
+  (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (metis (full_types) atLeastAtMost_iff convex_sum finite_atLeastAtMost)
   assume *: "\<forall>k u x. (\<forall> i :: nat. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> S) \<and> sum u {1..k} = 1
     \<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R (x i :: 'a)) \<in> S"
   {
@@ -262,18 +226,14 @@
     let ?x = "\<lambda>i. if (i :: nat) = 1 then x else y"
     have "{1 :: nat .. 2} \<inter> - {x. x = 1} = {2}"
       by auto
-    then have card: "card ({1 :: nat .. 2} \<inter> - {x. x = 1}) = 1"
-      by simp
-    then have "sum ?u {1 .. 2} = 1"
-      using sum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda> x. x = 1" "\<lambda> x. \<mu>" "\<lambda> x. 1 - \<mu>"]
-      by auto
-    with *[rule_format, of "2" ?u ?x] have S: "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) \<in> S"
-      using mu xy by auto
+    then have S: "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) \<in> S"
+      using sum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda>x. x = 1" "\<lambda>x. \<mu>" "\<lambda>x. 1 - \<mu>"]
+      using mu xy "*" by auto
     have grarr: "(\<Sum>j \<in> {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \<mu>) *\<^sub>R y"
       using sum.atLeast_Suc_atMost[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto
-    from sum.atLeast_Suc_atMost[of "Suc 0" 2 "\<lambda> j. ?u j *\<^sub>R ?x j", simplified this]
+    with sum.atLeast_Suc_atMost
     have "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) = \<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y"
-      by auto
+      by (smt (verit, best) Suc_1 Suc_eq_plus1 add_0 le_add1)
     then have "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x \<in> S"
       using S by (auto simp: add.commute)
   }
@@ -293,7 +253,7 @@
     and "finite t"
     and "t \<subseteq> S" "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1"
   then show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> S"
-    using convex_sum[of t S u "\<lambda> x. x"] by auto
+    by (simp add: convex_sum subsetD)
 next
   assume *: "\<forall>t. \<forall> u. finite t \<and> t \<subseteq> S \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and>
     sum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> S"
@@ -330,7 +290,8 @@
     assume *: "\<forall>x\<in>T. 0 \<le> u x" "sum u T = 1"
     assume "T \<subseteq> S"
     then have "S \<inter> T = T" by auto
-    with sum[THEN spec[where x="\<lambda>x. if x\<in>T then u x else 0"]] * have "(\<Sum>x\<in>T. u x *\<^sub>R x) \<in> S"
+    with sum[THEN spec[where x="\<lambda>x. if x\<in>T then u x else 0"]] *
+    have "(\<Sum>x\<in>T. u x *\<^sub>R x) \<in> S"
       by (auto simp: assms sum.If_cases if_distrib if_distrib_arg) }
   moreover assume ?rhs
   ultimately show ?lhs
@@ -357,30 +318,14 @@
     f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y"
   shows "convex_on A f"
   unfolding convex_on_def
-proof clarify
-  fix x y
-  fix u v :: real
-  assume A: "x \<in> A" "y \<in> A" "u \<ge> 0" "v \<ge> 0" "u + v = 1"
-  from A(5) have [simp]: "v = 1 - u"
-    by (simp add: algebra_simps)
-  from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y"
-    using assms[of u y x]
-    by (cases "u = 0 \<or> u = 1") (auto simp: algebra_simps)
-qed
+  by (smt (verit, del_insts) assms mult_cancel_right1 mult_eq_0_iff scaleR_collapse scaleR_eq_0_iff)
 
 lemma convex_on_linorderI [intro?]:
   fixes A :: "('a::{linorder,real_vector}) set"
   assumes "\<And>t x y. t > 0 \<Longrightarrow> t < 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x < y \<Longrightarrow>
     f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y"
   shows "convex_on A f"
-proof
-  fix x y
-  fix t :: real
-  assume A: "x \<in> A" "y \<in> A" "t > 0" "t < 1"
-  with assms [of t x y] assms [of "1 - t" y x]
-  show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y"
-    by (cases x y rule: linorder_cases) (auto simp: algebra_simps)
-qed
+  by (smt (verit, best) add.commute assms convex_onI distrib_left linorder_cases mult.commute mult_cancel_left2 scaleR_collapse)
 
 lemma convex_onD:
   assumes "convex_on A f"
@@ -452,7 +397,7 @@
 lemma convex_on_dist [intro]:
   fixes S :: "'a::real_normed_vector set"
   shows "convex_on S (\<lambda>x. dist a x)"
-proof (auto simp: convex_on_def dist_norm)
+proof (clarsimp simp: convex_on_def dist_norm)
   fix x y
   assume "x \<in> S" "y \<in> S"
   fix u v :: real
@@ -460,20 +405,18 @@
   assume "0 \<le> v"
   assume "u + v = 1"
   have "a = u *\<^sub>R a + v *\<^sub>R a"
-    unfolding scaleR_left_distrib[symmetric] and \<open>u + v = 1\<close> by simp
-  then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))"
+    by (metis \<open>u + v = 1\<close> scaleR_left.add scaleR_one)
+  then have "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))"
     by (auto simp: algebra_simps)
-  show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \<le> u * norm (a - x) + v * norm (a - y)"
-    unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"]
-    using \<open>0 \<le> u\<close> \<open>0 \<le> v\<close> by auto
+  then show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \<le> u * norm (a - x) + v * norm (a - y)"
+    by (smt (verit, best) \<open>0 \<le> u\<close> \<open>0 \<le> v\<close> norm_scaleR norm_triangle_ineq)
 qed
 
 
 subsection\<^marker>\<open>tag unimportant\<close> \<open>Arithmetic operations on sets preserve convexity\<close>
 
 lemma convex_linear_image:
-  assumes "linear f"
-    and "convex S"
+  assumes "linear f" and "convex S"
   shows "convex (f ` S)"
 proof -
   interpret f: linear f by fact
@@ -482,8 +425,7 @@
 qed
 
 lemma convex_linear_vimage:
-  assumes "linear f"
-    and "convex S"
+  assumes "linear f" and "convex S"
   shows "convex (f -` S)"
 proof -
   interpret f: linear f by fact
@@ -494,32 +436,17 @@
 lemma convex_scaling:
   assumes "convex S"
   shows "convex ((\<lambda>x. c *\<^sub>R x) ` S)"
-proof -
-  have "linear (\<lambda>x. c *\<^sub>R x)"
-    by (simp add: linearI scaleR_add_right)
-  then show ?thesis
-    using \<open>convex S\<close> by (rule convex_linear_image)
-qed
+  by (simp add: assms convex_linear_image)
 
 lemma convex_scaled:
   assumes "convex S"
   shows "convex ((\<lambda>x. x *\<^sub>R c) ` S)"
-proof -
-  have "linear (\<lambda>x. x *\<^sub>R c)"
-    by (simp add: linearI scaleR_add_left)
-  then show ?thesis
-    using \<open>convex S\<close> by (rule convex_linear_image)
-qed
+  by (simp add: assms convex_linear_image)
 
 lemma convex_negations:
   assumes "convex S"
   shows "convex ((\<lambda>x. - x) ` S)"
-proof -
-  have "linear (\<lambda>x. - x)"
-    by (simp add: linearI)
-  then show ?thesis
-    using \<open>convex S\<close> by (rule convex_linear_image)
-qed
+  by (simp add: assms convex_linear_image module_hom_uminus)
 
 lemma convex_sums:
   assumes "convex S"
@@ -572,124 +499,79 @@
   fixes a :: "'a \<Rightarrow> real"
     and y :: "'a \<Rightarrow> 'b::real_vector"
     and f :: "'b \<Rightarrow> real"
-  assumes "finite s" "s \<noteq> {}"
+  assumes "finite S" "S \<noteq> {}"
     and "convex_on C f"
     and "convex C"
-    and "(\<Sum> i \<in> s. a i) = 1"
-    and "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0"
-    and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C"
-  shows "f (\<Sum> i \<in> s. a i *\<^sub>R y i) \<le> (\<Sum> i \<in> s. a i * f (y i))"
+    and "(\<Sum> i \<in> S. a i) = 1"
+    and "\<And>i. i \<in> S \<Longrightarrow> a i \<ge> 0"
+    and "\<And>i. i \<in> S \<Longrightarrow> y i \<in> C"
+  shows "f (\<Sum> i \<in> S. a i *\<^sub>R y i) \<le> (\<Sum> i \<in> S. a i * f (y i))"
   using assms
-proof (induct s arbitrary: a rule: finite_ne_induct)
+proof (induct S arbitrary: a rule: finite_ne_induct)
   case (singleton i)
-  then have ai: "a i = 1"
-    by auto
   then show ?case
     by auto
 next
-  case (insert i s)
-  then have "convex_on C f"
-    by simp
-  from this[unfolded convex_on_def, rule_format]
-  have conv: "\<And>x y \<mu>. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> 0 \<le> \<mu> \<Longrightarrow> \<mu> \<le> 1 \<Longrightarrow>
+  case (insert i S)
+  then have yai: "y i \<in> C" "a i \<ge> 0"
+    by auto
+  with insert have conv: "\<And>x y \<mu>. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> 0 \<le> \<mu> \<Longrightarrow> \<mu> \<le> 1 \<Longrightarrow>
       f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y"
-    by simp
+    by (simp add: convex_on_def)
   show ?case
   proof (cases "a i = 1")
     case True
-    then have "(\<Sum> j \<in> s. a j) = 0"
-      using insert by auto
-    then have "\<And>j. j \<in> s \<Longrightarrow> a j = 0"
-      using insert by (fastforce simp: sum_nonneg_eq_0_iff)
-    then show ?thesis
-      using insert by auto
+    with insert have "(\<Sum> j \<in> S. a j) = 0"
+      by auto
+    with insert show ?thesis
+      by (simp add: sum_nonneg_eq_0_iff)
   next
     case False
-    from insert have yai: "y i \<in> C" "a i \<ge> 0"
-      by auto
-    have fis: "finite (insert i s)"
-      using insert by auto
-    then have ai1: "a i \<le> 1"
-      using sum_nonneg_leq_bound[of "insert i s" a] insert by simp
-    then have "a i < 1"
-      using False by auto
+    then have ai1: "a i < 1"
+      using sum_nonneg_leq_bound[of "insert i S" a] insert by force
     then have i0: "1 - a i > 0"
       by auto
     let ?a = "\<lambda>j. a j / (1 - a i)"
-    have a_nonneg: "?a j \<ge> 0" if "j \<in> s" for j
+    have a_nonneg: "?a j \<ge> 0" if "j \<in> S" for j
       using i0 insert that by fastforce
-    have "(\<Sum> j \<in> insert i s. a j) = 1"
+    have "(\<Sum> j \<in> insert i S. a j) = 1"
       using insert by auto
-    then have "(\<Sum> j \<in> s. a j) = 1 - a i"
+    then have "(\<Sum> j \<in> S. a j) = 1 - a i"
       using sum.insert insert by fastforce
-    then have "(\<Sum> j \<in> s. a j) / (1 - a i) = 1"
-      using i0 by auto
-    then have a1: "(\<Sum> j \<in> s. ?a j) = 1"
-      unfolding sum_divide_distrib by simp
-    have "convex C" using insert by auto
-    then have asum: "(\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<in> C"
-      using insert convex_sum [OF \<open>finite s\<close> \<open>convex C\<close> a1 a_nonneg] by auto
-    have asum_le: "f (\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<le> (\<Sum> j \<in> s. ?a j * f (y j))"
-      using a_nonneg a1 insert by blast
-    have "f (\<Sum> j \<in> insert i s. a j *\<^sub>R y j) = f ((\<Sum> j \<in> s. a j *\<^sub>R y j) + a i *\<^sub>R y i)"
-      using sum.insert[of s i "\<lambda> j. a j *\<^sub>R y j", OF \<open>finite s\<close> \<open>i \<notin> s\<close>] insert
-      by (auto simp only: add.commute)
-    also have "\<dots> = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\<Sum> j \<in> s. a j *\<^sub>R y j) + a i *\<^sub>R y i)"
+    then have "(\<Sum> j \<in> S. a j) / (1 - a i) = 1"
       using i0 by auto
-    also have "\<dots> = f ((1 - a i) *\<^sub>R (\<Sum> j \<in> s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)"
-      using scaleR_right.sum[of "inverse (1 - a i)" "\<lambda> j. a j *\<^sub>R y j" s, symmetric]
+    then have a1: "(\<Sum> j \<in> S. ?a j) = 1"
+      unfolding sum_divide_distrib by simp
+    have asum: "(\<Sum> j \<in> S. ?a j *\<^sub>R y j) \<in> C"
+      using insert convex_sum [OF \<open>finite S\<close> \<open>convex C\<close> a1 a_nonneg] by auto
+    have asum_le: "f (\<Sum> j \<in> S. ?a j *\<^sub>R y j) \<le> (\<Sum> j \<in> S. ?a j * f (y j))"
+      using a_nonneg a1 insert by blast
+    have "f (\<Sum> j \<in> insert i S. a j *\<^sub>R y j) = f ((\<Sum> j \<in> S. a j *\<^sub>R y j) + a i *\<^sub>R y i)"
+      by (simp add: add.commute insert.hyps)
+    also have "\<dots> = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\<Sum> j \<in> S. a j *\<^sub>R y j) + a i *\<^sub>R y i)"
+      using i0 by auto
+    also have "\<dots> = f ((1 - a i) *\<^sub>R (\<Sum> j \<in> S. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)"
+      using scaleR_right.sum[of "inverse (1 - a i)" "\<lambda> j. a j *\<^sub>R y j" S, symmetric]
       by (auto simp: algebra_simps)
-    also have "\<dots> = f ((1 - a i) *\<^sub>R (\<Sum> j \<in> s. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)"
+    also have "\<dots> = f ((1 - a i) *\<^sub>R (\<Sum> j \<in> S. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)"
       by (auto simp: divide_inverse)
-    also have "\<dots> \<le> (1 - a i) *\<^sub>R f ((\<Sum> j \<in> s. ?a j *\<^sub>R y j)) + a i * f (y i)"
-      using conv[of "y i" "(\<Sum> j \<in> s. ?a j *\<^sub>R y j)" "a i", OF yai(1) asum yai(2) ai1]
-      by (auto simp: add.commute)
-    also have "\<dots> \<le> (1 - a i) * (\<Sum> j \<in> s. ?a j * f (y j)) + a i * f (y i)"
-      using add_right_mono [OF mult_left_mono [of _ _ "1 - a i",
-            OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"]
-      by simp
-    also have "\<dots> = (\<Sum> j \<in> s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)"
-      unfolding sum_distrib_left[of "1 - a i" "\<lambda> j. ?a j * f (y j)"]
-      using i0 by auto
-    also have "\<dots> = (\<Sum> j \<in> s. a j * f (y j)) + a i * f (y i)"
-      using i0 by auto
-    also have "\<dots> = (\<Sum> j \<in> insert i s. a j * f (y j))"
+    also have "\<dots> \<le> (1 - a i) *\<^sub>R f ((\<Sum> j \<in> S. ?a j *\<^sub>R y j)) + a i * f (y i)"
+      using ai1 by (smt (verit) asum conv real_scaleR_def yai)
+    also have "\<dots> \<le> (1 - a i) * (\<Sum> j \<in> S. ?a j * f (y j)) + a i * f (y i)"
+      using asum_le i0 by fastforce
+    also have "\<dots> = (\<Sum> j \<in> S. a j * f (y j)) + a i * f (y i)"
+      using i0 by (auto simp: sum_distrib_left)
+    finally show ?thesis
       using insert by auto
-    finally show ?thesis
-      by simp
   qed
 qed
 
 lemma convex_on_alt:
   fixes C :: "'a::real_vector set"
   shows "convex_on C f \<longleftrightarrow>
-    (\<forall>x \<in> C. \<forall> y \<in> C. \<forall> \<mu> :: real. \<mu> \<ge> 0 \<and> \<mu> \<le> 1 \<longrightarrow>
-      f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y)"
-proof safe
-  fix x y
-  fix \<mu> :: real
-  assume *: "convex_on C f" "x \<in> C" "y \<in> C" "0 \<le> \<mu>" "\<mu> \<le> 1"
-  from this[unfolded convex_on_def, rule_format]
-  have "0 \<le> u \<Longrightarrow> 0 \<le> v \<Longrightarrow> u + v = 1 \<Longrightarrow> f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y" for u v
-    by auto
-  from this [of "\<mu>" "1 - \<mu>", simplified] *
-  show "f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y"
-    by auto
-next
-  assume *: "\<forall>x\<in>C. \<forall>y\<in>C. \<forall>\<mu>. 0 \<le> \<mu> \<and> \<mu> \<le> 1 \<longrightarrow>
-    f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y"
-  {
-    fix x y
-    fix u v :: real
-    assume **: "x \<in> C" "y \<in> C" "u \<ge> 0" "v \<ge> 0" "u + v = 1"
-    then have[simp]: "1 - u = v" by auto
-    from *[rule_format, of x y u]
-    have "f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y"
-      using ** by auto
-  }
-  then show "convex_on C f"
-    unfolding convex_on_def by auto
-qed
+         (\<forall>x \<in> C. \<forall>y \<in> C. \<forall> \<mu> :: real. \<mu> \<ge> 0 \<and> \<mu> \<le> 1 \<longrightarrow>
+          f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y)"
+  by (smt (verit) convex_on_def)
 
 lemma convex_on_diff:
   fixes f :: "real \<Rightarrow> real"
@@ -784,12 +666,10 @@
   shows "f' x * (y - x) \<le> f y - f x"
   using assms
 proof -
-  have less_imp: "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y"
+  have "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y"
     if *: "x \<in> C" "y \<in> C" "y > x" for x y :: real
   proof -
-    from * have ge: "y - x > 0" "y - x \<ge> 0"
-      by auto
-    from * have le: "x - y < 0" "x - y \<le> 0"
+    from * have ge: "y - x > 0" "y - x \<ge> 0" and le: "x - y < 0" "x - y \<le> 0"
       by auto
     then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1"
       using subsetD[OF atMostAtLeast_subset_convex[OF \<open>convex C\<close> \<open>x \<in> C\<close> \<open>y \<in> C\<close> \<open>x < y\<close>],
@@ -798,8 +678,6 @@
     then have "z1 \<in> C"
       using atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>y \<in> C\<close> \<open>x < y\<close>
       by fastforce
-    from z1 have z1': "f x - f y = (x - y) * f' z1"
-      by (simp add: field_simps)
     obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2"
       using subsetD[OF atMostAtLeast_subset_convex[OF \<open>convex C\<close> \<open>x \<in> C\<close> \<open>z1 \<in> C\<close> \<open>x < z1\<close>],
           THEN f'', THEN MVT2[OF \<open>x < z1\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1
@@ -808,75 +686,41 @@
       using subsetD[OF atMostAtLeast_subset_convex[OF \<open>convex C\<close> \<open>z1 \<in> C\<close> \<open>y \<in> C\<close> \<open>z1 < y\<close>],
           THEN f'', THEN MVT2[OF \<open>z1 < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1
       by auto
-    have "f' y - (f x - f y) / (x - y) = f' y - f' z1"
-      using * z1' by auto
-    also have "\<dots> = (y - z1) * f'' z3"
-      using z3 by auto
-    finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3"
-      by simp
-    have A': "y - z1 \<ge> 0"
-      using z1 by auto
+    from z1 have "f x - f y = (x - y) * f' z1"
+      by (simp add: field_simps)
+    then have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3"
+      using le(1) z3(3) by auto
     have "z3 \<in> C"
       using z3 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>z1 \<in> C\<close> \<open>x < z1\<close>
       by fastforce
     then have B': "f'' z3 \<ge> 0"
       using assms by auto
-    from A' B' have "(y - z1) * f'' z3 \<ge> 0"
-      by auto
-    from cool' this have "f' y - (f x - f y) / (x - y) \<ge> 0"
-      by auto
-    from mult_right_mono_neg[OF this le(2)]
-    have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \<le> 0 * (x - y)"
-      by (simp add: algebra_simps)
-    then have "f' y * (x - y) - (f x - f y) \<le> 0"
-      using le by auto
+    with cool' have "f' y - (f x - f y) / (x - y) \<ge> 0"
+      using z1 by auto
     then have res: "f' y * (x - y) \<le> f x - f y"
-      by auto
-    have "(f y - f x) / (y - x) - f' x = f' z1 - f' x"
-      using * z1 by auto
-    also have "\<dots> = (z1 - x) * f'' z2"
-      using z2 by auto
-    finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2"
-      by simp
-    have A: "z1 - x \<ge> 0"
-      using z1 by auto
+      by (meson diff_ge_0_iff_ge le(1) neg_divide_le_eq)
+    have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2"
+      using le(1) z1(3) z2(3) by auto
     have "z2 \<in> C"
       using z2 z1 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>z1 \<in> C\<close> \<open>y \<in> C\<close> \<open>z1 < y\<close>
       by fastforce
-    then have B: "f'' z2 \<ge> 0"
-      using assms by auto
-    from A B have "(z1 - x) * f'' z2 \<ge> 0"
-      by auto
-    with cool have "(f y - f x) / (y - x) - f' x \<ge> 0"
+    with z1 assms have "(z1 - x) * f'' z2 \<ge> 0"
       by auto
-    from mult_right_mono[OF this ge(2)]
-    have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \<ge> 0 * (y - x)"
-      by (simp add: algebra_simps)
-    then have "f y - f x - f' x * (y - x) \<ge> 0"
-      using ge by auto
     then show "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y"
-      using res by auto
+      using that(3) z1(3) res cool by auto
   qed
-  show ?thesis
-  proof (cases "x = y")
-    case True
-    with x y show ?thesis by auto
-  next
-    case False
-    with less_imp x y show ?thesis
-      by (auto simp: neq_iff)
-  qed
+  then show ?thesis
+    using x y by fastforce
 qed
 
 lemma f''_ge0_imp_convex:
   fixes f :: "real \<Rightarrow> real"
-  assumes conv: "convex C"
-    and f': "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)"
-    and f'': "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)"
-    and 0: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0"
+  assumes "convex C"
+    and "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)"
+    and "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)"
+    and "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0"
   shows "convex_on C f"
-  using f''_imp_f'[OF conv f' f'' 0] assms pos_convex_function
-  by fastforce
+  by (metis assms f''_imp_f' pos_convex_function)
 
 lemma f''_le0_imp_concave:
   fixes f :: "real \<Rightarrow> real"
@@ -965,14 +809,23 @@
 qed
 
 lemma convex_on_inverse:
+  fixes A :: "real set"
   assumes "A \<subseteq> {0<..}"
-  shows "convex_on A (inverse :: real \<Rightarrow> real)"
-proof (rule convex_on_subset[OF _ assms], intro convex_on_realI[of _ _ "\<lambda>x. -inverse (x^2)"])
-  fix u v :: real
-  assume "u \<in> {0<..}" "v \<in> {0<..}" "u \<le> v"
-  with assms show "-inverse (u^2) \<le> -inverse (v^2)"
-    by (intro le_imp_neg_le le_imp_inverse_le power_mono) (simp_all)
-qed (insert assms, auto intro!: derivative_eq_intros simp: field_split_simps power2_eq_square)
+  shows "convex_on A inverse"
+proof -
+  have "convex_on {0::real<..} inverse"
+  proof (intro convex_on_realI)
+    fix u v :: real
+    assume "u \<in> {0<..}" "v \<in> {0<..}" "u \<le> v"
+    with assms show "-inverse (u^2) \<le> -inverse (v^2)"
+      by simp
+  next
+    show "\<And>x. x \<in> {0<..} \<Longrightarrow> (inverse has_real_derivative - inverse (x\<^sup>2)) (at x)"
+      by (rule derivative_eq_intros | simp add: power2_eq_square)+
+  qed auto
+  then show ?thesis
+    using assms convex_on_subset by blast
+qed
 
 lemma convex_onD_Icc':
   assumes "convex_on {x..y} f" "c \<in> {x..y}"
@@ -995,7 +848,7 @@
   also from d have "\<dots> = (f y - f x) / d * (c - x) + f x"
     by (simp add: field_simps)
   finally show ?thesis .
-qed (insert assms(2), simp_all)
+qed (use assms in auto)
 
 lemma convex_onD_Icc'':
   assumes "convex_on {x..y} f" "c \<in> {x..y}"
@@ -1018,7 +871,7 @@
   also from d have "\<dots> = (f x - f y) / d * (y - c) + f y"
     by (simp add: field_simps)
   finally show ?thesis .
-qed (insert assms(2), simp_all)
+qed (use assms in auto)
 
 subsection \<open>Some inequalities\<close>
 
@@ -1161,59 +1014,34 @@
 
 lemma cone_iff:
   assumes "S \<noteq> {}"
-  shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> ((*\<^sub>R) c) ` S = S)"
-proof -
+  shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> ((*\<^sub>R) c) ` S = S)"  (is "_ = ?rhs")
+proof 
+  assume "cone S"
   {
-    assume "cone S"
-    {
-      fix c :: real
-      assume "c > 0"
-      {
-        fix x
-        assume "x \<in> S"
-        then have "x \<in> ((*\<^sub>R) c) ` S"
-          unfolding image_def
-          using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"]
-            exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"]
-          by auto
-      }
-      moreover
-      {
-        fix x
-        assume "x \<in> ((*\<^sub>R) c) ` S"
-        then have "x \<in> S"
-          using \<open>0 < c\<close> \<open>cone S\<close> mem_cone by fastforce
-      }
-      ultimately have "((*\<^sub>R) c) ` S = S" by blast
-    }
-    then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> ((*\<^sub>R) c) ` S = S)"
-      using \<open>cone S\<close> cone_contains_0[of S] assms by auto
+    fix c :: real
+    assume "c > 0"
+    have "x \<in> ((*\<^sub>R) c) ` S" if "x \<in> S" for x
+        using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"] that
+          exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"]
+        by auto
+    then have "((*\<^sub>R) c) ` S = S" 
+        using \<open>0 < c\<close> \<open>cone S\<close> mem_cone by fastforce
   }
-  moreover
-  {
-    assume a: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> ((*\<^sub>R) c) ` S = S)"
-    {
-      fix x
-      assume "x \<in> S"
-      fix c1 :: real
-      assume "c1 \<ge> 0"
-      then have "c1 = 0 \<or> c1 > 0" by auto
-      then have "c1 *\<^sub>R x \<in> S" using a \<open>x \<in> S\<close> by auto
-    }
-    then have "cone S" unfolding cone_def by auto
-  }
-  ultimately show ?thesis by blast
+  then show "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> ((*\<^sub>R) c) ` S = S)"
+    using \<open>cone S\<close> cone_contains_0[of S] assms by auto
+next
+  show "?rhs \<Longrightarrow> cone S"
+    by (metis Convex.cone_def imageI order_neq_le_trans scaleR_zero_left)
 qed
 
 lemma cone_hull_empty: "cone hull {} = {}"
   by (metis cone_empty cone_hull_eq)
 
 lemma cone_hull_empty_iff: "S = {} \<longleftrightarrow> cone hull S = {}"
-  by (metis bot_least cone_hull_empty hull_subset xtrans(5))
+  by (metis cone_hull_empty hull_subset subset_empty)
 
 lemma cone_hull_contains_0: "S \<noteq> {} \<longleftrightarrow> 0 \<in> cone hull S"
-  using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S]
-  by auto
+  by (metis cone_cone_hull cone_contains_0 cone_hull_empty_iff)
 
 lemma mem_cone_hull:
   assumes "x \<in> S" "c \<ge> 0"
@@ -1222,65 +1050,27 @@
 
 proposition cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}"
   (is "?lhs = ?rhs")
-proof -
-  {
-    fix x
-    assume "x \<in> ?rhs"
-    then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
-      by auto
-    fix c :: real
-    assume c: "c \<ge> 0"
-    then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx"
-      using x by (simp add: algebra_simps)
-    moreover
-    have "c * cx \<ge> 0" using c x by auto
-    ultimately
-    have "c *\<^sub>R x \<in> ?rhs" using x by auto
-  }
-  then have "cone ?rhs"
-    unfolding cone_def by auto
-  then have "?rhs \<in> Collect cone"
-    unfolding mem_Collect_eq by auto
-  {
-    fix x
-    assume "x \<in> S"
-    then have "1 *\<^sub>R x \<in> ?rhs"
-      using zero_le_one by blast
-    then have "x \<in> ?rhs" by auto
-  }
-  then have "S \<subseteq> ?rhs" by auto
-  then have "?lhs \<subseteq> ?rhs"
-    using \<open>?rhs \<in> Collect cone\<close> hull_minimal[of S "?rhs" "cone"] by auto
-  moreover
-  {
-    fix x
-    assume "x \<in> ?rhs"
-    then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
-      by auto
-    then have "xx \<in> cone hull S"
-      using hull_subset[of S] by auto
-    then have "x \<in> ?lhs"
-      using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto
-  }
-  ultimately show ?thesis by auto
-qed
+proof 
+  have "?rhs \<in> Collect cone"
+    using Convex.cone_def by fastforce
+  moreover have "S \<subseteq> ?rhs"
+    by (smt (verit) mem_Collect_eq scaleR_one subsetI)
+  ultimately show "?lhs \<subseteq> ?rhs"
+    using hull_minimal by blast
+qed (use mem_cone_hull in auto)
 
 lemma convex_cone:
-  "convex s \<and> cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. (x + y) \<in> s) \<and> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)"
+  "convex S \<and> cone S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. (x + y) \<in> S) \<and> (\<forall>x\<in>S. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> S)"
   (is "?lhs = ?rhs")
 proof -
   {
     fix x y
-    assume "x\<in>s" "y\<in>s" and ?lhs
-    then have "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s"
+    assume "x\<in>S" "y\<in>S" and ?lhs
+    then have "2 *\<^sub>R x \<in>S" "2 *\<^sub>R y \<in> S" "convex S"
       unfolding cone_def by auto
-    then have "x + y \<in> s"
-      using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1]
-      apply (erule_tac x="2*\<^sub>R x" in ballE)
-      apply (erule_tac x="2*\<^sub>R y" in ballE)
-      apply (erule_tac x="1/2" in allE, simp)
-      apply (erule_tac x="1/2" in allE, auto)
-      done
+    then have "x + y \<in> S"
+      using convexD [OF \<open>convex S\<close>, of "2*\<^sub>R x" "2*\<^sub>R y"]
+      by (smt (verit, ccfv_threshold) field_sum_of_halves scaleR_2 scaleR_half_double)
   }
   then show ?thesis
     unfolding convex_def cone_def by blast
@@ -1315,13 +1105,12 @@
 qed
 
 corollary%unimportant connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)"
-by (simp add: convex_connected)
+  by (simp add: convex_connected)
 
 lemma convex_prod:
   assumes "\<And>i. i \<in> Basis \<Longrightarrow> convex {x. P i x}"
   shows "convex {x. \<forall>i\<in>Basis. P i (x\<bullet>i)}"
-  using assms unfolding convex_def
-  by (auto simp: inner_add_left)
+  using assms by (auto simp: inner_add_left convex_def)
 
 lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}"
 by (rule convex_prod) (simp flip: atLeast_def)
@@ -1329,9 +1118,7 @@
 subsection \<open>Convex hull\<close>
 
 lemma convex_convex_hull [iff]: "convex (convex hull s)"
-  unfolding hull_def
-  using convex_Inter[of "{t. convex t \<and> s \<subseteq> t}"]
-  by auto
+  by (metis (mono_tags) convex_Inter hull_def mem_Collect_eq)
 
 lemma convex_hull_subset:
     "s \<subseteq> convex hull t \<Longrightarrow> convex hull s \<subseteq> convex hull t"
@@ -1344,56 +1131,50 @@
 
 lemma convex_hull_linear_image:
   assumes f: "linear f"
-  shows "f ` (convex hull s) = convex hull (f ` s)"
+  shows "f ` (convex hull S) = convex hull (f ` S)"
 proof
-  show "convex hull (f ` s) \<subseteq> f ` (convex hull s)"
+  show "convex hull (f ` S) \<subseteq> f ` (convex hull S)"
     by (intro hull_minimal image_mono hull_subset convex_linear_image assms convex_convex_hull)
-  show "f ` (convex hull s) \<subseteq> convex hull (f ` s)"
-  proof (unfold image_subset_iff_subset_vimage, rule hull_minimal)
-    show "s \<subseteq> f -` (convex hull (f ` s))"
-      by (fast intro: hull_inc)
-    show "convex (f -` (convex hull (f ` s)))"
-      by (intro convex_linear_vimage [OF f] convex_convex_hull)
-  qed
+  show "f ` (convex hull S) \<subseteq> convex hull (f ` S)"
+    by (meson convex_convex_hull convex_linear_vimage f hull_minimal hull_subset image_subset_iff_subset_vimage)
 qed
 
 lemma in_convex_hull_linear_image:
-  assumes "linear f"
-    and "x \<in> convex hull s"
-  shows "f x \<in> convex hull (f ` s)"
-  using convex_hull_linear_image[OF assms(1)] assms(2) by auto
+  assumes "linear f" "x \<in> convex hull S"
+  shows "f x \<in> convex hull (f ` S)"
+  using assms convex_hull_linear_image image_eqI by blast
 
 lemma convex_hull_Times:
-  "convex hull (s \<times> t) = (convex hull s) \<times> (convex hull t)"
+  "convex hull (S \<times> T) = (convex hull S) \<times> (convex hull T)"
 proof
-  show "convex hull (s \<times> t) \<subseteq> (convex hull s) \<times> (convex hull t)"
+  show "convex hull (S \<times> T) \<subseteq> (convex hull S) \<times> (convex hull T)"
     by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull)
-  have "(x, y) \<in> convex hull (s \<times> t)" if x: "x \<in> convex hull s" and y: "y \<in> convex hull t" for x y
+  have "(x, y) \<in> convex hull (S \<times> T)" if x: "x \<in> convex hull S" and y: "y \<in> convex hull T" for x y
   proof (rule hull_induct [OF x], rule hull_induct [OF y])
-    fix x y assume "x \<in> s" and "y \<in> t"
-    then show "(x, y) \<in> convex hull (s \<times> t)"
+    fix x y assume "x \<in> S" and "y \<in> T"
+    then show "(x, y) \<in> convex hull (S \<times> T)"
       by (simp add: hull_inc)
   next
-    fix x let ?S = "((\<lambda>y. (0, y)) -` (\<lambda>p. (- x, 0) + p) ` (convex hull s \<times> t))"
+    fix x let ?S = "((\<lambda>y. (0, y)) -` (\<lambda>p. (- x, 0) + p) ` (convex hull S \<times> T))"
     have "convex ?S"
       by (intro convex_linear_vimage convex_translation convex_convex_hull,
         simp add: linear_iff)
-    also have "?S = {y. (x, y) \<in> convex hull (s \<times> t)}"
+    also have "?S = {y. (x, y) \<in> convex hull (S \<times> T)}"
       by (auto simp: image_def Bex_def)
-    finally show "convex {y. (x, y) \<in> convex hull (s \<times> t)}" .
+    finally show "convex {y. (x, y) \<in> convex hull (S \<times> T)}" .
   next
-    show "convex {x. (x, y) \<in> convex hull s \<times> t}"
+    show "convex {x. (x, y) \<in> convex hull S \<times> T}"
     proof -
-      fix y let ?S = "((\<lambda>x. (x, 0)) -` (\<lambda>p. (0, - y) + p) ` (convex hull s \<times> t))"
+      fix y let ?S = "((\<lambda>x. (x, 0)) -` (\<lambda>p. (0, - y) + p) ` (convex hull S \<times> T))"
       have "convex ?S"
       by (intro convex_linear_vimage convex_translation convex_convex_hull,
         simp add: linear_iff)
-      also have "?S = {x. (x, y) \<in> convex hull (s \<times> t)}"
+      also have "?S = {x. (x, y) \<in> convex hull (S \<times> T)}"
         by (auto simp: image_def Bex_def)
-      finally show "convex {x. (x, y) \<in> convex hull (s \<times> t)}" .
+      finally show "convex {x. (x, y) \<in> convex hull (S \<times> T)}" .
     qed
   qed
-  then show "(convex hull s) \<times> (convex hull t) \<subseteq> convex hull (s \<times> t)"
+  then show "(convex hull S) \<times> (convex hull T) \<subseteq> convex hull (S \<times> T)"
     unfolding subset_eq split_paired_Ball_Sigma by blast
 qed
 
@@ -1401,10 +1182,10 @@
 subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Stepping theorems for convex hulls of finite sets\<close>
 
 lemma convex_hull_empty[simp]: "convex hull {} = {}"
-  by (rule hull_unique) auto
+  by (simp add: hull_same)
 
 lemma convex_hull_singleton[simp]: "convex hull {a} = {a}"
-  by (rule hull_unique) auto
+  by (simp add: hull_same)
 
 lemma convex_hull_insert:
   fixes S :: "'a::real_vector set"
@@ -1415,19 +1196,17 @@
 proof (intro equalityI hull_minimal subsetI)
   fix x
   assume "x \<in> insert a S"
-  then have "\<exists>u\<ge>0. \<exists>v\<ge>0. u + v = 1 \<and> (\<exists>b. b \<in> convex hull S \<and> x = u *\<^sub>R a + v *\<^sub>R b)"
+  then show "x \<in> ?hull"
   unfolding insert_iff
   proof
     assume "x = a"
     then show ?thesis
-      by (rule_tac x=1 in exI) (use assms hull_subset in fastforce)
+      by (smt (verit, del_insts) add.right_neutral assms ex_in_conv hull_inc mem_Collect_eq scaleR_one scaleR_zero_left)
   next
     assume "x \<in> S"
-    with hull_subset[of S convex] show ?thesis
+    with hull_subset show ?thesis
       by force
   qed
-  then show "x \<in> ?hull"
-    by simp
 next
   fix x
   assume "x \<in> ?hull"
@@ -1468,44 +1247,35 @@
     next
       case False
       have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)"
-        using as(3) obt1(3) obt2(3) by (auto simp: field_simps)
-      also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)"
-        using as(3) obt1(3) obt2(3) by (auto simp: field_simps)
+        by (simp add: as(3))
       also have "\<dots> = u * v1 + v * v2"
-        by simp
-      finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
+        by (smt (verit, ccfv_SIG) distrib_left mult_cancel_left1 obt1(3) obt2(3))
+      finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" .
       let ?b = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2"
       have zeroes: "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2"
-        using as(1,2) obt1(1,2) obt2(1,2) by auto
+        using as obt1 obt2 by auto
       show ?thesis
       proof
         show "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (?b - (u * u1) *\<^sub>R ?b - (v * u2) *\<^sub>R ?b)"
           unfolding xeq yeq * **
           using False by (auto simp: scaleR_left_distrib scaleR_right_distrib)
         show "?b \<in> convex hull S"
-          using False zeroes obt1(4) obt2(4)
-          by (auto simp: convexD [OF convex_convex_hull] scaleR_left_distrib scaleR_right_distrib  add_divide_distrib[symmetric]  zero_le_divide_iff)
+          using False mem_convex_alt obt1(4) obt2(4) zeroes(2) zeroes(4) by fastforce
       qed
     qed
     then obtain b where b: "b \<in> convex hull S" 
        "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" ..
-
-    have u1: "u1 \<le> 1"
-      unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto
-    have u2: "u2 \<le> 1"
-      unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto
+    obtain u1: "u1 \<le> 1" and u2: "u2 \<le> 1"
+      using obt1 obt2 by auto
     have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v"
-    proof (rule add_mono)
-      show "u1 * u \<le> max u1 u2 * u" "u2 * v \<le> max u1 u2 * v"
-        by (simp_all add: as mult_right_mono)
-    qed
+      by (smt (verit, ccfv_SIG) as mult_right_mono)
     also have "\<dots> \<le> 1"
       unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto
     finally have le1: "u1 * u + u2 * v \<le> 1" .    
     show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull"
     proof (intro CollectI exI conjI)
       show "0 \<le> u * u1 + v * u2"
-        by (simp add: as(1) as(2) obt1(1) obt2(1))
+        by (simp add: as obt1(1) obt2(1))
       show "0 \<le> 1 - u * u1 - v * u2"
         by (simp add: le1 diff_diff_add mult.commute)
     qed (use b in \<open>auto simp: algebra_simps\<close>)
@@ -1516,9 +1286,9 @@
    "convex hull (insert a S) =
      (if S = {} then {a}
       else {(1 - u) *\<^sub>R a + u *\<^sub>R x |x u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> convex hull S})"
-  apply (auto simp: convex_hull_insert)
-  using diff_eq_eq apply fastforce
-  using diff_add_cancel diff_ge_0_iff_ge by blast
+  apply (simp add: convex_hull_insert)
+  using diff_add_cancel diff_ge_0_iff_ge
+  by (smt (verit, del_insts) Collect_cong) 
 
 subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Explicit expression for convex hull\<close>
 
@@ -1608,87 +1378,56 @@
   shows "convex hull p =
     {y. \<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}"
   (is "?lhs = ?rhs")
-proof -
+proof (intro subset_antisym subsetI)
+  fix x
+  assume "x \<in> convex hull p"
+  then obtain k u y where
+    obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
+    unfolding convex_hull_indexed by auto
+  have fin: "finite {1..k}" by auto
   {
-    fix x
-    assume "x\<in>?lhs"
-    then obtain k u y where
-        obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
-      unfolding convex_hull_indexed by auto
-
-    have fin: "finite {1..k}" by auto
-    have fin': "\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
-    {
-      fix j
-      assume "j\<in>{1..k}"
-      then have "y j \<in> p \<and> 0 \<le> sum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
-        using obt(1)[THEN bspec[where x=j]] and obt(2)
-        by (metis (no_types, lifting) One_nat_def atLeastAtMost_iff mem_Collect_eq obt(1) sum_nonneg)
-    }
-    moreover
-    have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v}) = 1"
-      unfolding sum.image_gen[OF fin, symmetric] using obt(2) by auto
-    moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
-      using sum.image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
-      unfolding scaleR_left.sum using obt(3) by auto
-    ultimately
-    have "\<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = x"
-      apply (rule_tac x="y ` {1..k}" in exI)
-      apply (rule_tac x="\<lambda>v. sum u {i\<in>{1..k}. y i = v}" in exI, auto)
-      done
-    then have "x\<in>?rhs" by auto
+    fix j
+    assume "j\<in>{1..k}"
+    then have "y j \<in> p \<and> 0 \<le> sum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
+      by (metis (mono_tags, lifting) One_nat_def atLeastAtMost_iff mem_Collect_eq obt(1) sum_nonneg)
   }
-  moreover
+  moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v}) = 1"
+    unfolding sum.image_gen[OF fin, symmetric] using obt(2) by auto
+  moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
+    using sum.image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
+    unfolding scaleR_left.sum using obt(3) by auto
+  ultimately
+  have "\<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = x"
+    by (smt (verit, ccfv_SIG) imageE mem_Collect_eq obt(1) subsetI sum.cong sum.infinite sum_nonneg)
+  then show "x \<in> ?rhs" by auto
+next
+  fix y
+  assume "y \<in> ?rhs"
+  then obtain S u where
+    S: "finite S" "S \<subseteq> p" "\<forall>x\<in>S. 0 \<le> u x" "sum u S = 1" "(\<Sum>v\<in>S. u v *\<^sub>R v) = y"
+    by auto
+  obtain f where f: "inj_on f {1..card S}" "f ` {1..card S} = S"
+    using ex_bij_betw_nat_finite_1[OF S(1)] unfolding bij_betw_def by auto
+  then have "0 \<le> u (f i)" "f i \<in> p" if "i \<in> {1..card S}" for i
+    using S \<open>i \<in> {1..card S}\<close> by blast+
+  moreover 
   {
     fix y
-    assume "y\<in>?rhs"
-    then obtain S u where
-      obt: "finite S" "S \<subseteq> p" "\<forall>x\<in>S. 0 \<le> u x" "sum u S = 1" "(\<Sum>v\<in>S. u v *\<^sub>R v) = y"
-      by auto
-
-    obtain f where f: "inj_on f {1..card S}" "f ` {1..card S} = S"
-      using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto
-    {
-      fix i :: nat
-      assume "i\<in>{1..card S}"
-      then have "f i \<in> S"
-        using f(2) by blast
-      then have "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto
-    }
-    moreover have *: "finite {1..card S}" by auto
-    {
-      fix y
-      assume "y\<in>S"
-      then obtain i where "i\<in>{1..card S}" "f i = y"
-        using f using image_iff[of y f "{1..card S}"]
-        by auto
-      then have "{x. Suc 0 \<le> x \<and> x \<le> card S \<and> f x = y} = {i}"
-        using f(1) inj_onD by fastforce
-      then have "card {x. Suc 0 \<le> x \<and> x \<le> card S \<and> f x = y} = 1" by auto
-      then have "(\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x)) = u y"
-          "(\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
-        by (auto simp: sum_constant_scaleR)
-    }
-    then have "(\<Sum>x = 1..card S. u (f x)) = 1" "(\<Sum>i = 1..card S. u (f i) *\<^sub>R f i) = y"
-      unfolding sum.image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
-        and sum.image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
-      unfolding f
-      using sum.cong [of S S "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
-      using sum.cong [of S S "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x))" u]
-      unfolding obt(4,5)
-      by auto
-    ultimately
-    have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> sum u {1..k} = 1 \<and>
-        (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
-      apply (rule_tac x="card S" in exI)
-      apply (rule_tac x="u \<circ> f" in exI)
-      apply (rule_tac x=f in exI, fastforce)
-      done
-    then have "y \<in> ?lhs"
-      unfolding convex_hull_indexed by auto
+    assume "y\<in>S"
+    then obtain i where "i\<in>{1..card S}" "f i = y"
+      by (metis f(2) image_iff)
+    then have "{x. Suc 0 \<le> x \<and> x \<le> card S \<and> f x = y} = {i}"
+      using f(1) inj_onD by fastforce
+    then have "(\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x)) = u y"
+      "(\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
+      by (simp_all add: sum_constant_scaleR \<open>f i = y\<close>)
   }
-  ultimately show ?thesis
-    unfolding set_eq_iff by blast
+  then have "(\<Sum>x = 1..card S. u (f x)) = 1" "(\<Sum>i = 1..card S. u (f i) *\<^sub>R f i) = y"
+    by (metis (mono_tags, lifting) S(4,5) f sum.reindex_cong)+
+  ultimately
+  show "y \<in> convex hull p"
+    unfolding convex_hull_indexed
+    by (smt (verit, del_insts) mem_Collect_eq sum.cong)
 qed
 
 
@@ -1832,10 +1571,7 @@
 lemma aff_dim_convex_hull:
   fixes S :: "'n::euclidean_space set"
   shows "aff_dim (convex hull S) = aff_dim S"
-  using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S]
-    hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"]
-    aff_dim_subset[of "convex hull S" "affine hull S"]
-  by auto
+  by (smt (verit) aff_dim_affine_hull aff_dim_subset convex_hull_subset_affine_hull hull_subset)
 
 
 subsection \<open>Caratheodory's theorem\<close>
@@ -1856,40 +1592,31 @@
     by (rule_tac ex_least_nat_le, auto)
   then obtain n where "?P n" and smallest: "\<forall>k<n. \<not> ?P k"
     by blast
-  then obtain S u where obt: "finite S" "card S = n" "S\<subseteq>p" "\<forall>x\<in>S. 0 \<le> u x"
-    "sum u S = 1"  "(\<Sum>v\<in>S. u v *\<^sub>R v) = y" by auto
+  then obtain S u where S: "finite S" "card S = n" "S\<subseteq>p" 
+    and u: "\<forall>x\<in>S. 0 \<le> u x" "sum u S = 1"  "(\<Sum>v\<in>S. u v *\<^sub>R v) = y" by auto
 
   have "card S \<le> aff_dim p + 1"
   proof (rule ccontr, simp only: not_le)
     assume "aff_dim p + 1 < card S"
     then have "affine_dependent S"
-      using affine_dependent_biggerset[OF obt(1)] independent_card_le_aff_dim not_less obt(3)
-      by blast
+      by (smt (verit) independent_card_le_aff_dim S(3))
     then obtain w v where wv: "sum w S = 0" "v\<in>S" "w v \<noteq> 0" "(\<Sum>v\<in>S. w v *\<^sub>R v) = 0"
-      using affine_dependent_explicit_finite[OF obt(1)] by auto
+      using affine_dependent_explicit_finite[OF S(1)] by auto
     define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>S. w v < 0}"
     define t where "t = Min i"
     have "\<exists>x\<in>S. w x < 0"
-    proof (rule ccontr, simp add: not_less)
-      assume as:"\<forall>x\<in>S. 0 \<le> w x"
-      then have "sum w (S - {v}) \<ge> 0"
-        by (meson Diff_iff sum_nonneg)
-      then have "sum w S > 0"
-        using as obt(1) sum_nonneg_eq_0_iff wv by blast
-      then show False using wv(1) by auto
-    qed
+      by (smt (verit, best) S(1) sum_pos2 wv)
     then have "i \<noteq> {}" unfolding i_def by auto
     then have "t \<ge> 0"
-      using Min_ge_iff[of i 0] and obt(1)
+      using Min_ge_iff[of i 0] and S(1) u[unfolded le_less]
       unfolding t_def i_def
-      using obt(4)[unfolded le_less]
       by (auto simp: divide_le_0_iff)
     have t: "\<forall>v\<in>S. u v + t * w v \<ge> 0"
     proof
       fix v
       assume "v \<in> S"
       then have v: "0 \<le> u v"
-        using obt(4)[THEN bspec[where x=v]] by auto
+        using u(1) by blast
       show "0 \<le> u v + t * w v"
       proof (cases "w v < 0")
         case False
@@ -1897,26 +1624,26 @@
       next
         case True
         then have "t \<le> u v / (- w v)"
-          using \<open>v\<in>S\<close> obt unfolding t_def i_def by (auto intro: Min_le)
+          using \<open>v\<in>S\<close> S unfolding t_def i_def by (auto intro: Min_le)
         then show ?thesis
           unfolding real_0_le_add_iff
           using True neg_le_minus_divide_eq by auto
       qed
     qed
     obtain a where "a \<in> S" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
-      using Min_in[OF _ \<open>i\<noteq>{}\<close>] and obt(1) unfolding i_def t_def by auto
+      using Min_in[OF _ \<open>i\<noteq>{}\<close>] and S(1) unfolding i_def t_def by auto
     then have a: "a \<in> S" "u a + t * w a = 0" by auto
     have *: "\<And>f. sum f (S - {a}) = sum f S - ((f a)::'b::ab_group_add)"
-      unfolding sum.remove[OF obt(1) \<open>a\<in>S\<close>] by auto
+      unfolding sum.remove[OF S(1) \<open>a\<in>S\<close>] by auto
     have "(\<Sum>v\<in>S. u v + t * w v) = 1"
-      unfolding sum.distrib wv(1) sum_distrib_left[symmetric] obt(5) by auto
+      by (metis add.right_neutral mult_zero_right sum.distrib sum_distrib_left u(2) wv(1))
     moreover have "(\<Sum>v\<in>S. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y"
-      unfolding sum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] wv(4)
+      unfolding sum.distrib u(3) scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] wv(4)
       using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
     ultimately have "?P (n - 1)"
       apply (rule_tac x="(S - {a})" in exI)
       apply (rule_tac x="\<lambda>v. u v + t * w v" in exI)
-      using obt(1-3) and t and a
+      using S t a
       apply (auto simp: * scaleR_left_distrib)
       done
     then show False
@@ -1924,7 +1651,7 @@
   qed
   then show "\<exists>S u. finite S \<and> S \<subseteq> p \<and> card S \<le> aff_dim p + 1 \<and>
       (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = y"
-    using obt by auto
+    using S u by auto
 qed auto
 
 lemma caratheodory_aff_dim:
@@ -1934,7 +1661,7 @@
 proof
   have "\<And>x S u. \<lbrakk>finite S; S \<subseteq> p; int (card S) \<le> aff_dim p + 1; \<forall>x\<in>S. 0 \<le> u x; sum u S = 1\<rbrakk>
                 \<Longrightarrow> (\<Sum>v\<in>S. u v *\<^sub>R v) \<in> convex hull S"
-    by (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull])
+    by (metis (mono_tags, lifting) convex_convex_hull convex_explicit hull_subset)
   then show "?lhs \<subseteq> ?rhs"
     by (subst convex_hull_caratheodory_aff_dim, auto)
 qed (use hull_mono in auto)
@@ -1986,23 +1713,19 @@
 
 lemma convex_hull_set_plus:
   "convex hull (S + T) = convex hull S + convex hull T"
-  unfolding set_plus_image 
-  apply (subst convex_hull_linear_image [symmetric])
-  apply (simp add: linear_iff scaleR_right_distrib)
-  apply (simp add: convex_hull_Times)
-  done
+  by (simp add: set_plus_image linear_iff scaleR_right_distrib convex_hull_Times 
+        flip: convex_hull_linear_image)
 
 lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` T = {a} + T"
   unfolding set_plus_def by auto
 
 lemma convex_hull_translation:
   "convex hull ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (convex hull S)"
-  unfolding translation_eq_singleton_plus
-  by (simp only: convex_hull_set_plus convex_hull_singleton)
+  by (simp add: convex_hull_set_plus translation_eq_singleton_plus)
 
 lemma convex_hull_scaling:
   "convex hull ((\<lambda>x. c *\<^sub>R x) ` S) = (\<lambda>x. c *\<^sub>R x) ` (convex hull S)"
-  using linear_scaleR by (rule convex_hull_linear_image [symmetric])
+  by (simp add: convex_hull_linear_image)
 
 lemma convex_hull_affinity:
   "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` S) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull S)"
@@ -2022,31 +1745,25 @@
   fix u v :: real
   assume uv: "u \<ge> 0" "v \<ge> 0" "u + v = 1"
   then have *: "u *\<^sub>R x \<in> cone hull S" "v *\<^sub>R y \<in> cone hull S"
-    using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto
-  from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
-    using cone_hull_expl[of S] by auto
-  from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S"
+    by (simp_all add: cone_cone_hull mem_cone uv xy)
+  then obtain cx :: real and xx
+      and cy :: real and yy  where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" 
+      and y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S"
     using cone_hull_expl[of S] by auto
-  {
-    assume "cx + cy \<le> 0"
-    then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0"
-      using x y by auto
-    then have "u *\<^sub>R x + v *\<^sub>R y = 0"
-      by auto
-    then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
-      using cone_hull_contains_0[of S] \<open>S \<noteq> {}\<close> by auto
-  }
+
+  have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" if "cx + cy \<le> 0"
+    using "*"(1) nless_le that x(2) y by fastforce
   moreover
-  {
-    assume "cx + cy > 0"
-    then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S"
-      using assms mem_convex_alt[of S xx yy cx cy] x y by auto
+  have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" if "cx + cy > 0"
+  proof -
+    have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S"
+      using assms mem_convex_alt[of S xx yy cx cy] x y that by auto
     then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S"
       using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \<open>cx+cy>0\<close>
       by (auto simp: scaleR_right_distrib)
-    then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
+    then show ?thesis
       using x y by auto
-  }
+  qed
   moreover have "cx + cy \<le> 0 \<or> cx + cy > 0" by auto
   ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" by blast
 qed
@@ -2054,28 +1771,7 @@
 lemma cone_convex_hull:
   assumes "cone S"
   shows "cone (convex hull S)"
-proof (cases "S = {}")
-  case True
-  then show ?thesis by auto
-next
-  case False
-  then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (*\<^sub>R) c ` S = S)"
-    using cone_iff[of S] assms by auto
-  {
-    fix c :: real
-    assume "c > 0"
-    then have "(*\<^sub>R) c ` (convex hull S) = convex hull ((*\<^sub>R) c ` S)"
-      using convex_hull_scaling[of _ S] by auto
-    also have "\<dots> = convex hull S"
-      using * \<open>c > 0\<close> by auto
-    finally have "(*\<^sub>R) c ` (convex hull S) = convex hull S"
-      by auto
-  }
-  then have "0 \<in> convex hull S" "\<And>c. c > 0 \<Longrightarrow> ((*\<^sub>R) c ` (convex hull S)) = (convex hull S)"
-    using * hull_subset[of S convex] by auto
-  then show ?thesis
-    using \<open>S \<noteq> {}\<close> cone_iff[of "convex hull S"] by auto
-qed
+  by (metis (no_types, lifting) affine_hull_convex_hull affine_hull_eq_empty assms cone_iff convex_hull_scaling hull_inc)
 
 subsection \<open>Radon's theorem\<close>
 
@@ -2084,29 +1780,17 @@
 lemma Radon_ex_lemma:
   assumes "finite c" "affine_dependent c"
   shows "\<exists>u. sum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) c = 0"
-proof -
-  from assms(2)[unfolded affine_dependent_explicit]
-  obtain S u where
-      "finite S" "S \<subseteq> c" "sum u S = 0" "\<exists>v\<in>S. u v \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
-    by blast
-  then show ?thesis
-    apply (rule_tac x="\<lambda>v. if v\<in>S then u v else 0" in exI)
-    unfolding if_smult scaleR_zero_left 
-    by (auto simp: Int_absorb1 sum.inter_restrict[OF \<open>finite c\<close>, symmetric])
-qed
+  using affine_dependent_explicit_finite assms by blast
 
 lemma Radon_s_lemma:
   assumes "finite S"
     and "sum f S = (0::real)"
   shows "sum f {x\<in>S. 0 < f x} = - sum f {x\<in>S. f x < 0}"
 proof -
-  have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x"
+  have "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x"
     by auto
-  show ?thesis
-    unfolding add_eq_0_iff[symmetric] and sum.inter_filter[OF assms(1)]
-      and sum.distrib[symmetric] and *
-    using assms(2)
-    by assumption
+  then show ?thesis
+    using assms by (simp add: sum.inter_filter flip: sum.distrib add_eq_0_iff)
 qed
 
 lemma Radon_v_lemma:
@@ -2115,19 +1799,15 @@
     and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)"
   shows "(sum f {x\<in>S. 0 < g x}) = - sum f {x\<in>S. g x < 0}"
 proof -
-  have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x"
-    using assms(3) by auto
-  show ?thesis
-    unfolding eq_neg_iff_add_eq_0 and sum.inter_filter[OF assms(1)]
-      and sum.distrib[symmetric] and *
-    using assms(2)
-    apply assumption
-    done
+  have "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x"
+    using assms by auto
+  then show ?thesis
+    using assms by (simp add: sum.inter_filter eq_neg_iff_add_eq_0 flip: sum.distrib add_eq_0_iff)
 qed
 
 lemma Radon_partition:
   assumes "finite C" "affine_dependent C"
-  shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = C \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}"
+  shows "\<exists>M P. M \<inter> P = {} \<and> M \<union> P = C \<and> (convex hull M) \<inter> (convex hull P) \<noteq> {}"
 proof -
   obtain u v where uv: "sum u C = 0" "v\<in>C" "u v \<noteq> 0"  "(\<Sum>v\<in>C. u v *\<^sub>R v) = 0"
     using Radon_ex_lemma[OF assms] by auto
@@ -2139,19 +1819,12 @@
     case False
     then have "u v < 0" by auto
     then show ?thesis
-    proof (cases "\<exists>w\<in>{x \<in> C. 0 < u x}. u w > 0")
-      case True
-      then show ?thesis
-        using sum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
-    next
-      case False
-      then have "sum u C \<le> sum (\<lambda>x. if x=v then u v else 0) C"
-        by (rule_tac sum_mono, auto)
-      then show ?thesis
-        unfolding sum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto
-    qed
-  qed (insert sum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
-
+      by (smt (verit) assms(1) fin(1) mem_Collect_eq sum.neutral_const sum_mono_inv uv)
+  next
+    case True
+    with fin uv show "sum u {x \<in> C. 0 < u x} \<noteq> 0"
+      by (smt (verit) fin(1) mem_Collect_eq sum_nonneg_eq_0_iff uv)
+  qed
   then have *: "sum u {x\<in>C. u x > 0} > 0"
     unfolding less_le by (metis (no_types, lifting) mem_Collect_eq sum_nonneg)
   moreover have "sum u ({x \<in> C. 0 < u x} \<union> {x \<in> C. u x < 0}) = sum u C"
@@ -2188,97 +1861,85 @@
 
 theorem Radon:
   assumes "affine_dependent c"
-  obtains m p where "m \<subseteq> c" "p \<subseteq> c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}"
-proof -
-  from assms[unfolded affine_dependent_explicit]
-  obtain S u where
-      "finite S" "S \<subseteq> c" "sum u S = 0" "\<exists>v\<in>S. u v \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
-    by blast
-  then have *: "finite S" "affine_dependent S" and S: "S \<subseteq> c"
-    unfolding affine_dependent_explicit by auto
-  from Radon_partition[OF *]
-  obtain m p where "m \<inter> p = {}" "m \<union> p = S" "convex hull m \<inter> convex hull p \<noteq> {}"
-    by blast
-  with S show ?thesis
-    by (force intro: that[of p m])
-qed
+  obtains M P where "M \<subseteq> c" "P \<subseteq> c" "M \<inter> P = {}" "(convex hull M) \<inter> (convex hull P) \<noteq> {}"
+  by (smt (verit) Radon_partition affine_dependent_explicit affine_dependent_explicit_finite assms le_sup_iff)
 
 
 subsection \<open>Helly's theorem\<close>
 
 lemma Helly_induct:
-  fixes f :: "'a::euclidean_space set set"
-  assumes "card f = n"
+  fixes \<F> :: "'a::euclidean_space set set"
+  assumes "card \<F> = n"
     and "n \<ge> DIM('a) + 1"
-    and "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
-  shows "\<Inter>f \<noteq> {}"
+    and "\<forall>S\<in>\<F>. convex S" "\<forall>T\<subseteq>\<F>. card T = DIM('a) + 1 \<longrightarrow> \<Inter>T \<noteq> {}"
+  shows "\<Inter>\<F> \<noteq> {}"
   using assms
-proof (induction n arbitrary: f)
+proof (induction n arbitrary: \<F>)
   case 0
   then show ?case by auto
 next
   case (Suc n)
-  have "finite f"
-    using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite)
-  show "\<Inter>f \<noteq> {}"
+  have "finite \<F>"
+    using \<open>card \<F> = Suc n\<close> by (auto intro: card_ge_0_finite)
+  show "\<Inter>\<F> \<noteq> {}"
   proof (cases "n = DIM('a)")
     case True
     then show ?thesis
-      by (simp add: Suc.prems(1) Suc.prems(4))
+      by (simp add: Suc.prems)
   next
     case False
-    have "\<Inter>(f - {s}) \<noteq> {}" if "s \<in> f" for s
+    have "\<Inter>(\<F> - {S}) \<noteq> {}" if "S \<in> \<F>" for S
     proof (rule Suc.IH[rule_format])
-      show "card (f - {s}) = n"
-        by (simp add: Suc.prems(1) \<open>finite f\<close> that)
+      show "card (\<F> - {S}) = n"
+        by (simp add: Suc.prems(1) \<open>finite \<F>\<close> that)
       show "DIM('a) + 1 \<le> n"
         using False Suc.prems(2) by linarith
-      show "\<And>t. \<lbrakk>t \<subseteq> f - {s}; card t = DIM('a) + 1\<rbrakk> \<Longrightarrow> \<Inter>t \<noteq> {}"
+      show "\<And>t. \<lbrakk>t \<subseteq> \<F> - {S}; card t = DIM('a) + 1\<rbrakk> \<Longrightarrow> \<Inter>t \<noteq> {}"
         by (simp add: Suc.prems(4) subset_Diff_insert)
     qed (use Suc in auto)
-    then have "\<forall>s\<in>f. \<exists>x. x \<in> \<Inter>(f - {s})"
+    then have "\<forall>S\<in>\<F>. \<exists>x. x \<in> \<Inter>(\<F> - {S})"
       by blast
-    then obtain X where X: "\<And>s. s\<in>f \<Longrightarrow> X s \<in> \<Inter>(f - {s})"
+    then obtain X where X: "\<And>S. S\<in>\<F> \<Longrightarrow> X S \<in> \<Inter>(\<F> - {S})"
       by metis
     show ?thesis
-    proof (cases "inj_on X f")
+    proof (cases "inj_on X \<F>")
       case False
-      then obtain s t where "s\<noteq>t" and st: "s\<in>f" "t\<in>f" "X s = X t"
+      then obtain S T where "S\<noteq>T" and st: "S\<in>\<F>" "T\<in>\<F>" "X S = X T"
         unfolding inj_on_def by auto
-      then have *: "\<Inter>f = \<Inter>(f - {s}) \<inter> \<Inter>(f - {t})" by auto
+      then have *: "\<Inter>\<F> = \<Inter>(\<F> - {S}) \<inter> \<Inter>(\<F> - {T})" by auto
       show ?thesis
         by (metis "*" X disjoint_iff_not_equal st)
     next
       case True
-      then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
-        using Radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
-        unfolding card_image[OF True] and \<open>card f = Suc n\<close>
-        using Suc(3) \<open>finite f\<close> and False
+      then obtain M P where mp: "M \<inter> P = {}" "M \<union> P = X ` \<F>" "convex hull M \<inter> convex hull P \<noteq> {}"
+        using Radon_partition[of "X ` \<F>"] and affine_dependent_biggerset[of "X ` \<F>"]
+        unfolding card_image[OF True] and \<open>card \<F> = Suc n\<close>
+        using Suc(3) \<open>finite \<F>\<close> and False
         by auto
-      have "m \<subseteq> X ` f" "p \<subseteq> X ` f"
+      have "M \<subseteq> X ` \<F>" "P \<subseteq> X ` \<F>"
         using mp(2) by auto
-      then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f"
+      then obtain \<G> \<H> where gh:"M = X ` \<G>" "P = X ` \<H>" "\<G> \<subseteq> \<F>" "\<H> \<subseteq> \<F>"
         unfolding subset_image_iff by auto
-      then have "f \<union> (g \<union> h) = f" by auto
-      then have f: "f = g \<union> h"
-        using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True
+      then have "\<F> \<union> (\<G> \<union> \<H>) = \<F>" by auto
+      then have \<F>: "\<F> = \<G> \<union> \<H>"
+        using inj_on_Un_image_eq_iff[of X \<F> "\<G> \<union> \<H>"] and True
         unfolding mp(2)[unfolded image_Un[symmetric] gh]
         by auto
-      have *: "g \<inter> h = {}"
-        using gh(1) gh(2) local.mp(1) by blast
-      have "convex hull (X ` h) \<subseteq> \<Inter>g" "convex hull (X ` g) \<subseteq> \<Inter>h"
-        by (rule hull_minimal; use X * f in \<open>auto simp: Suc.prems(3) convex_Inter\<close>)+
+      have *: "\<G> \<inter> \<H> = {}"
+        using gh local.mp(1) by blast
+      have "convex hull (X ` \<H>) \<subseteq> \<Inter>\<G>" "convex hull (X ` \<G>) \<subseteq> \<Inter>\<H>"
+        by (rule hull_minimal; use X * \<F> in \<open>auto simp: Suc.prems(3) convex_Inter\<close>)+
       then show ?thesis
-        unfolding f using mp(3)[unfolded gh] by blast
+        unfolding \<F> using mp(3)[unfolded gh] by blast
     qed
   qed 
 qed
 
 theorem Helly:
-  fixes f :: "'a::euclidean_space set set"
-  assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s"
-    and "\<And>t. \<lbrakk>t\<subseteq>f; card t = DIM('a) + 1\<rbrakk> \<Longrightarrow> \<Inter>t \<noteq> {}"
-  shows "\<Inter>f \<noteq> {}"
+  fixes \<F> :: "'a::euclidean_space set set"
+  assumes "card \<F> \<ge> DIM('a) + 1" "\<forall>s\<in>\<F>. convex s"
+    and "\<And>t. \<lbrakk>t\<subseteq>\<F>; card t = DIM('a) + 1\<rbrakk> \<Longrightarrow> \<Inter>t \<noteq> {}"
+  shows "\<Inter>\<F> \<noteq> {}"
   using Helly_induct assms by blast
 
 subsection \<open>Epigraphs of convex functions\<close>
@@ -2371,25 +2032,18 @@
 
 lemma convex_set_plus:
   assumes "convex S" and "convex T" shows "convex (S + T)"
-proof -
-  have "convex (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
-    using assms by (rule convex_sums)
-  moreover have "(\<Union>x\<in> S. \<Union>y \<in> T. {x + y}) = S + T"
-    unfolding set_plus_def by auto
-  finally show "convex (S + T)" .
-qed
+  by (metis assms convex_hull_eq convex_hull_set_plus)
 
 lemma convex_set_sum:
   assumes "\<And>i. i \<in> A \<Longrightarrow> convex (B i)"
   shows "convex (\<Sum>i\<in>A. B i)"
-proof (cases "finite A")
-  case True then show ?thesis using assms
-    by induct (auto simp: convex_set_plus)
-qed auto
+  using assms
+  by (induction A rule: infinite_finite_induct) (auto simp: convex_set_plus)
 
 lemma finite_set_sum:
-  assumes "finite A" and "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)"
-  using assms by (induct set: finite, simp, simp add: finite_set_plus)
+  assumes "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)"
+  using assms
+  by (induction A rule: infinite_finite_induct) (auto simp: finite_set_plus)
 
 lemma box_eq_set_sum_Basis:
   "{x. \<forall>i\<in>Basis. x\<bullet>i \<in> B i} = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (B i))" (is "?lhs = ?rhs")
@@ -2401,7 +2055,7 @@
   have "sum f Basis \<bullet> i \<in> B i" if "i \<in> Basis" and f: "\<forall>i\<in>Basis. f i \<in> (\<lambda>x. x *\<^sub>R i) ` B i" for i f
   proof -
     have "(\<Sum>x\<in>Basis - {i}. f x \<bullet> i) = 0"
-    proof (rule sum.neutral, intro strip)
+    proof (intro strip sum.neutral)
       show "f x \<bullet> i = 0" if "x \<in> Basis - {i}" for x
         using that f \<open>i \<in> Basis\<close> inner_Basis that by fastforce
     qed
@@ -2418,10 +2072,6 @@
 
 lemma convex_hull_set_sum:
   "convex hull (\<Sum>i\<in>A. B i) = (\<Sum>i\<in>A. convex hull (B i))"
-proof (cases "finite A")
-  assume "finite A" then show ?thesis
-    by (induct set: finite, simp, simp add: convex_hull_set_plus)
-qed simp
-
+  by (induction A rule: infinite_finite_induct) (auto simp: convex_hull_set_plus)
 
 end
\ No newline at end of file
--- a/src/HOL/Analysis/Derivative.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Derivative.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -163,27 +163,27 @@
 lemma linear_imp_differentiable_on:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
   shows "linear f \<Longrightarrow> f differentiable_on S"
-by (simp add: differentiable_on_def linear_imp_differentiable)
+  by (simp add: differentiable_on_def linear_imp_differentiable)
 
 lemma differentiable_on_minus [simp, derivative_intros]:
-   "f differentiable_on S \<Longrightarrow> (\<lambda>z. -(f z)) differentiable_on S"
-by (simp add: differentiable_on_def)
+  "f differentiable_on S \<Longrightarrow> (\<lambda>z. -(f z)) differentiable_on S"
+  by (simp add: differentiable_on_def)
 
 lemma differentiable_on_add [simp, derivative_intros]:
-   "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) differentiable_on S"
-by (simp add: differentiable_on_def)
+  "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) differentiable_on S"
+  by (simp add: differentiable_on_def)
 
 lemma differentiable_on_diff [simp, derivative_intros]:
-   "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) differentiable_on S"
-by (simp add: differentiable_on_def)
+  "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) differentiable_on S"
+  by (simp add: differentiable_on_def)
 
 lemma differentiable_on_inverse [simp, derivative_intros]:
   fixes f :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_field"
   shows "f differentiable_on S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> f x \<noteq> 0) \<Longrightarrow> (\<lambda>x. inverse (f x)) differentiable_on S"
-by (simp add: differentiable_on_def)
+  by (simp add: differentiable_on_def)
 
 lemma differentiable_on_scaleR [derivative_intros, simp]:
-   "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>x. f x *\<^sub>R g x) differentiable_on S"
+  "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>x. f x *\<^sub>R g x) differentiable_on S"
   unfolding differentiable_on_def
   by (blast intro: differentiable_scaleR)
 
@@ -195,22 +195,22 @@
 lemma differentiable_sqnorm_at [derivative_intros, simp]:
   fixes a :: "'a :: {real_normed_vector,real_inner}"
   shows "(\<lambda>x. (norm x)\<^sup>2) differentiable (at a)"
-by (force simp add: differentiable_def intro: has_derivative_sqnorm_at)
+  by (force simp add: differentiable_def intro: has_derivative_sqnorm_at)
 
 lemma differentiable_on_sqnorm [derivative_intros, simp]:
   fixes S :: "'a :: {real_normed_vector,real_inner} set"
   shows "(\<lambda>x. (norm x)\<^sup>2) differentiable_on S"
-by (simp add: differentiable_at_imp_differentiable_on)
+  by (simp add: differentiable_at_imp_differentiable_on)
 
 lemma differentiable_norm_at [derivative_intros, simp]:
   fixes a :: "'a :: {real_normed_vector,real_inner}"
   shows "a \<noteq> 0 \<Longrightarrow> norm differentiable (at a)"
-using differentiableI has_derivative_norm by blast
+  using differentiableI has_derivative_norm by blast
 
 lemma differentiable_on_norm [derivative_intros, simp]:
   fixes S :: "'a :: {real_normed_vector,real_inner} set"
   shows "0 \<notin> S \<Longrightarrow> norm differentiable_on S"
-by (metis differentiable_at_imp_differentiable_on differentiable_norm_at)
+  by (metis differentiable_at_imp_differentiable_on differentiable_norm_at)
 
 
 subsection \<open>Frechet derivative and Jacobian matrix\<close>
@@ -480,12 +480,7 @@
   assumes x: "x \<in> box a b"
     and f: "(f has_derivative f' ) (at x within box a b)" "(f has_derivative f'') (at x within box a b)"
   shows "f' = f''"
-proof -
-  have "at x within box a b = at x"
-    by (metis x at_within_interior interior_open open_box)
-  with f show "f' = f''"
-    by (simp add: has_derivative_unique)
-qed
+  by (metis at_within_open assms has_derivative_unique open_box)
 
 lemma frechet_derivative_at:
   "(f has_derivative f') (at x) \<Longrightarrow> f' = frechet_derivative f (at x)"
@@ -613,7 +608,7 @@
   shows "\<exists>x\<in>{a..b}. f b - f a = f' x (b - a)"
 proof (cases "a = b")
   interpret bounded_linear "f' b"
-    using assms(2) assms(1) by auto
+    using assms by auto
   case True
   then show ?thesis
     by force
@@ -944,10 +939,12 @@
       and dn:  "\<And>z. z \<in> S \<Longrightarrow> norm (f' z) \<le> B"
       and "x \<in> S"  "y \<in> S"
     shows "norm(f x - f y) \<le> B * norm(x - y)"
-  apply (rule differentiable_bound [OF cvs])
-  apply (erule df [unfolded has_field_derivative_def])
-  apply (rule onorm_le, simp_all add: norm_mult mult_right_mono assms)
-  done
+proof (rule differentiable_bound [OF cvs])
+  show "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative (*) (f' x)) (at x within S)"
+    by (simp add: df has_field_derivative_imp_has_derivative)
+  show "\<And>x. x \<in> S \<Longrightarrow> onorm ((*) (f' x)) \<le> B"
+    by (metis (no_types, opaque_lifting) dn norm_mult onorm_le order.refl order_trans)
+qed (use assms in auto)
 
 lemma
   differentiable_bound_segment:
--- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -1090,8 +1090,7 @@
 proposition compact_eq_Bolzano_Weierstrass:
   fixes S :: "'a::metric_space set"
   shows "compact S \<longleftrightarrow> (\<forall>T. infinite T \<and> T \<subseteq> S \<longrightarrow> (\<exists>x \<in> S. x islimpt T))"
-  using Bolzano_Weierstrass_imp_seq_compact Heine_Borel_imp_Bolzano_Weierstrass compact_eq_seq_compact_metric 
-  by blast
+  by (meson Bolzano_Weierstrass_imp_seq_compact Heine_Borel_imp_Bolzano_Weierstrass seq_compact_imp_Heine_Borel)
 
 proposition Bolzano_Weierstrass_imp_bounded:
   "(\<And>T. \<lbrakk>infinite T; T \<subseteq> S\<rbrakk> \<Longrightarrow> (\<exists>x \<in> S. x islimpt T)) \<Longrightarrow> bounded S"
@@ -1511,10 +1510,8 @@
     using bounded_imp_convergent_subsequence [OF \<open>bounded (range f)\<close>] by auto
   from f have fr: "\<forall>n. (f \<circ> r) n \<in> S"
     by simp
-  have "l \<in> S" using \<open>closed S\<close> fr l
-    by (rule closed_sequentially)
-  show "\<exists>l\<in>S. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
-    using \<open>l \<in> S\<close> r l by blast
+  show "\<exists>l\<in>S. \<exists>r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
+    using assms(2) closed_sequentially fr l r by blast
 qed
 
 lemma compact_eq_bounded_closed:
@@ -2829,7 +2826,7 @@
     moreover have "(\<lambda>i. (x \<circ> r) (i + n)) \<longlonglongrightarrow> l"
       using lr(3) by (rule LIMSEQ_ignore_initial_segment)
     ultimately show "l \<in> S n"
-      by (rule closed_sequentially)
+      by (metis closed_sequentially)
   qed
   then show ?thesis 
     using that by blast
--- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -1776,9 +1776,9 @@
     by blast
   have "?thesis1 \<and> ?thesis2 \<and> ?thesis3"
     apply (rule *)
-    using continuous_disconnected_range_constant apply metis
-    apply clarify
-    apply (frule discrete_subset_disconnected; blast)
+    using continuous_disconnected_range_constant
+    apply (metis image_subset_iff_funcset)
+    apply (smt (verit, best) discrete_subset_disconnected mem_Collect_eq subsetD subsetI)
     apply (blast dest: finite_implies_discrete)
     apply (blast intro!: finite_range_constant_imp_connected)
     done
--- a/src/HOL/Analysis/Elementary_Topology.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Elementary_Topology.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -53,13 +53,14 @@
 
 lemma topological_basis:
   "topological_basis B \<longleftrightarrow> (\<forall>x. open x \<longleftrightarrow> (\<exists>B'. B' \<subseteq> B \<and> \<Union>B' = x))"
-  unfolding topological_basis_def
-  apply safe
-     apply fastforce
-    apply fastforce
-   apply (erule_tac x=x in allE, simp)
-   apply (rule_tac x="{x}" in exI, auto)
-  done
+    (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (meson local.open_Union subsetD topological_basis_def)
+  show "?rhs \<Longrightarrow> ?lhs"
+    unfolding topological_basis_def
+    by (metis cSup_singleton empty_subsetI insert_subset)
+qed
 
 lemma topological_basis_iff:
   assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
@@ -89,7 +90,7 @@
   assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
     and "\<And>O' x. open O' \<Longrightarrow> x \<in> O' \<Longrightarrow> \<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'"
   shows "topological_basis B"
-  using assms by (subst topological_basis_iff) auto
+  by (simp add: assms topological_basis_iff)
 
 lemma topological_basisE:
   fixes O'
@@ -97,17 +98,10 @@
     and "open O'"
     and "x \<in> O'"
   obtains B' where "B' \<in> B" "x \<in> B'" "B' \<subseteq> O'"
-proof atomize_elim
-  from assms have "\<And>B'. B'\<in>B \<Longrightarrow> open B'"
-    by (simp add: topological_basis_def)
-  with topological_basis_iff assms
-  show  "\<exists>B'. B' \<in> B \<and> x \<in> B' \<and> B' \<subseteq> O'"
-    using assms by (simp add: Bex_def)
-qed
+  by (metis assms topological_basis_def topological_basis_iff)
 
 lemma topological_basis_open:
-  assumes "topological_basis B"
-    and "X \<in> B"
+  assumes "topological_basis B" and "X \<in> B"
   shows "open X"
   using assms by (simp add: topological_basis_def)
 
@@ -131,17 +125,9 @@
 lemma basis_dense:
   fixes B :: "'a set set"
     and f :: "'a set \<Rightarrow> 'a"
-  assumes "topological_basis B"
-    and choosefrom_basis: "\<And>B'. B' \<noteq> {} \<Longrightarrow> f B' \<in> B'"
+  assumes "topological_basis B" and "\<And>B'. B' \<noteq> {} \<Longrightarrow> f B' \<in> B'"
   shows "\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>B' \<in> B. f B' \<in> X)"
-proof (intro allI impI)
-  fix X :: "'a set"
-  assume "open X" and "X \<noteq> {}"
-  from topological_basisE[OF \<open>topological_basis B\<close> \<open>open X\<close> choosefrom_basis[OF \<open>X \<noteq> {}\<close>]]
-  obtain B' where "B' \<in> B" "f X \<in> B'" "B' \<subseteq> X" .
-  then show "\<exists>B'\<in>B. f B' \<in> X"
-    by (auto intro!: choosefrom_basis)
-qed
+  by (metis assms equals0D in_mono topological_basisE)
 
 end
 
@@ -149,27 +135,31 @@
   assumes A: "topological_basis A"
     and B: "topological_basis B"
   shows "topological_basis ((\<lambda>(a, b). a \<times> b) ` (A \<times> B))"
-  unfolding topological_basis_def
-proof (safe, simp_all del: ex_simps add: subset_image_iff ex_simps(1)[symmetric])
-  fix S :: "('a \<times> 'b) set"
-  assume "open S"
-  then show "\<exists>X\<subseteq>A \<times> B. (\<Union>(a,b)\<in>X. a \<times> b) = S"
-  proof (safe intro!: exI[of _ "{x\<in>A \<times> B. fst x \<times> snd x \<subseteq> S}"])
-    fix x y
-    assume "(x, y) \<in> S"
-    from open_prod_elim[OF \<open>open S\<close> this]
-    obtain a b where a: "open a""x \<in> a" and b: "open b" "y \<in> b" and "a \<times> b \<subseteq> S"
-      by (metis mem_Sigma_iff)
-    moreover
-    from A a obtain A0 where "A0 \<in> A" "x \<in> A0" "A0 \<subseteq> a"
-      by (rule topological_basisE)
-    moreover
-    from B b obtain B0 where "B0 \<in> B" "y \<in> B0" "B0 \<subseteq> b"
-      by (rule topological_basisE)
-    ultimately show "(x, y) \<in> (\<Union>(a, b)\<in>{X \<in> A \<times> B. fst X \<times> snd X \<subseteq> S}. a \<times> b)"
-      by (intro UN_I[of "(A0, B0)"]) auto
-  qed auto
-qed (metis A B topological_basis_open open_Times)
+proof -
+  have "\<exists>X\<subseteq>A \<times> B. (\<Union>(a,b)\<in>X. a \<times> b) = S" if "open S" for S
+  proof -
+    have "(x, y) \<in> (\<Union>(a, b)\<in>{X \<in> A \<times> B. fst X \<times> snd X \<subseteq> S}. a \<times> b)"
+      if xy: "(x, y) \<in> S" for x y
+    proof -
+      obtain a b where a: "open a""x \<in> a" and b: "open b" "y \<in> b" and "a \<times> b \<subseteq> S"
+        by (metis open_prod_elim[OF \<open>open S\<close>] xy mem_Sigma_iff)
+      moreover obtain A0 where "A0 \<in> A" "x \<in> A0" "A0 \<subseteq> a"
+        using A a b topological_basisE by blast
+      moreover
+      from B b obtain B0 where "B0 \<in> B" "y \<in> B0" "B0 \<subseteq> b"
+        by (rule topological_basisE)
+      ultimately show ?thesis
+        by (intro UN_I[of "(A0, B0)"]) auto
+    qed
+    then have "(\<Union>(a, b)\<in>{x \<in> A \<times> B. fst x \<times> snd x \<subseteq> S}. a \<times> b) = S"
+      by force
+    then show ?thesis
+      using subset_eq by force
+  qed
+  with A B show ?thesis
+    unfolding topological_basis_def
+    by (smt (verit) SigmaE imageE image_mono open_Times case_prod_conv)
+qed
 
 
 subsection \<open>Countable Basis\<close>
@@ -189,8 +179,7 @@
 lemma open_countable_basisE:
   assumes "p X"
   obtains B' where "B' \<subseteq> B" "X = \<Union>B'"
-  using assms open_countable_basis_ex
-  by atomize_elim simp
+  using assms open_countable_basis_ex by auto
 
 lemma countable_dense_exists:
   "\<exists>D::'a set. countable D \<and> (\<forall>X. p X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>d \<in> D. d \<in> X))"
@@ -221,11 +210,10 @@
 proof -
   obtain \<A> where \<A>: "(\<forall>i::nat. x \<in> \<A> i \<and> open (\<A> i))" "(\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. \<A> i \<subseteq> S))"
     using first_countable_basis[of x] by metis
-  show thesis
-  proof 
-    show "countable (range \<A>)"
-      by simp
-  qed (use \<A> in auto)
+  moreover have "countable (range \<A>)"
+    by simp
+  ultimately show thesis
+    by (smt (verit, best) imageE rangeI that)
 qed
 
 lemma (in first_countable_topology) first_countable_basis_Int_stableE:
@@ -262,8 +250,14 @@
     fix S
     assume "open S" "x \<in> S"
     then obtain a where a: "a\<in>\<B>" "a \<subseteq> S" using \<B> by blast
-    then show "\<exists>a\<in>\<A>. a \<subseteq> S" using a \<B>
-      by (intro bexI[where x=a]) (auto simp: \<A>_def intro: image_eqI[where x="{to_nat_on \<B> a}"])
+    moreover have "a\<in>\<A>"
+      unfolding \<A>_def 
+    proof (rule image_eqI)
+      show "a = \<Inter> (from_nat_into \<B> ` {to_nat_on \<B> a})"
+        by (simp add: \<B> a)
+    qed auto
+    ultimately show "\<exists>a\<in>\<A>. a \<subseteq> S"
+      by blast 
   qed
 qed
 
@@ -279,9 +273,9 @@
     using range_from_nat_into_subset[OF \<open>\<A> \<noteq> {}\<close>] 1 by auto
 next
   fix S
-  assume "open S" "x\<in>S" from 2[OF this]
-  show "\<exists>i. from_nat_into \<A> i \<subseteq> S"
-    using subset_range_from_nat_into[OF \<open>countable \<A>\<close>] by auto
+  assume "open S" "x\<in>S" 
+  then show "\<exists>i. from_nat_into \<A> i \<subseteq> S"
+    by (metis "2" \<open>countable \<A>\<close> from_nat_into_surj)
 qed
 
 instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology
@@ -314,9 +308,8 @@
     then obtain a' b' where a'b': "open a'" "open b'" "x \<in> a' \<times> b'" "a' \<times> b' \<subseteq> S"
       by (rule open_prod_elim)
     moreover
-    from a'b' \<A>(4)[of a'] B(4)[of b']
     obtain a b where "a \<in> \<A>" "a \<subseteq> a'" "b \<in> B" "b \<subseteq> b'"
-      by auto
+      by (meson B(4) \<A>(4) a'b' mem_Times_iff)
     ultimately
     show "\<exists>a\<in>(\<lambda>(a, b). a \<times> b) ` (\<A> \<times> B). a \<subseteq> S"
       by (auto intro!: bexI[of _ "a \<times> b"] bexI[of _ a] bexI[of _ b])
@@ -396,9 +389,7 @@
     using univ_second_countable by blast
   define \<D> where "\<D> \<equiv> {S. S \<in> \<B> \<and> (\<exists>U. U \<in> \<F> \<and> S \<subseteq> U)}"
   have "countable \<D>"
-    apply (rule countable_subset [OF _ \<open>countable \<B>\<close>])
-    apply (force simp: \<D>_def)
-    done
+    by (simp add: \<D>_def \<open>countable \<B>\<close>)
   have "\<And>S. \<exists>U. S \<in> \<D> \<longrightarrow> U \<in> \<F> \<and> S \<subseteq> U"
     by (simp add: \<D>_def)
   then obtain G where G: "\<And>S. S \<in> \<D> \<longrightarrow> G S \<in> \<F> \<and> S \<subseteq> G S"
@@ -408,12 +399,10 @@
   moreover have "\<Union>\<D> \<subseteq> \<Union>\<F>"
     using \<D>_def by blast
   ultimately have eq1: "\<Union>\<F> = \<Union>\<D>" ..
-  have eq2: "\<Union>\<D> = \<Union> (G ` \<D>)"
+  moreover have "\<Union>\<D> = \<Union> (G ` \<D>)"
     using G eq1 by auto
-  show ?thesis
-    apply (rule_tac \<F>' = "G ` \<D>" in that)
-    using G \<open>countable \<D>\<close>
-    by (auto simp: eq1 eq2)
+  ultimately show ?thesis
+    by (metis G \<open>countable \<D>\<close> countable_image image_subset_iff that)
 qed
 
 lemma countable_disjoint_open_subsets:
@@ -606,11 +595,8 @@
   unfolding islimpt_UNIV_iff by (rule not_open_singleton)
 
 lemma closed_limpt: "closed S \<longleftrightarrow> (\<forall>x. x islimpt S \<longrightarrow> x \<in> S)"
-  unfolding closed_def
-  apply (subst open_subopen)
-  apply (simp add: islimpt_def subset_eq)
-  apply (metis ComplE ComplI)
-  done
+  unfolding closed_def open_subopen [of "-S"]
+  by (metis Compl_iff islimptE islimptI open_subopen subsetI)
 
 lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
   by (auto simp: islimpt_def)
@@ -625,25 +611,25 @@
 
 lemma islimpt_insert:
   fixes x :: "'a::t1_space"
-  shows "x islimpt (insert a s) \<longleftrightarrow> x islimpt s"
+  shows "x islimpt (insert a S) \<longleftrightarrow> x islimpt S"
 proof
-  assume "x islimpt (insert a s)"
-  then show "x islimpt s"
+  assume "x islimpt (insert a S)"
+  then show "x islimpt S"
     by (metis closed_limpt closed_singleton empty_iff insert_iff insert_is_Un islimpt_Un islimpt_def)
 next
-  assume "x islimpt s"
-  then show "x islimpt (insert a s)"
+  assume "x islimpt S"
+  then show "x islimpt (insert a S)"
     by (rule islimpt_subset) auto
 qed
 
 lemma islimpt_finite:
   fixes x :: "'a::t1_space"
-  shows "finite s \<Longrightarrow> \<not> x islimpt s"
+  shows "finite S \<Longrightarrow> \<not> x islimpt S"
   by (induct set: finite) (simp_all add: islimpt_insert)
 
 lemma islimpt_Un_finite:
   fixes x :: "'a::t1_space"
-  shows "finite s \<Longrightarrow> x islimpt (s \<union> t) \<longleftrightarrow> x islimpt t"
+  shows "finite S \<Longrightarrow> x islimpt (S \<union> T) \<longleftrightarrow> x islimpt T"
   by (simp add: islimpt_Un islimpt_finite)
 
 lemma islimpt_eq_acc_point:
@@ -659,10 +645,8 @@
 next
   fix T
   assume *: "\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> S)" "l \<in> T" "open T"
-  then have "infinite (T \<inter> S - {l})"
-    by auto
   then have "\<exists>x. x \<in> (T \<inter> S - {l})"
-    unfolding ex_in_conv by (intro notI) simp
+    by (metis ex_in_conv finite.emptyI infinite_remove)
   then show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> l"
     by auto
 qed
@@ -684,11 +668,9 @@
     have "infinite (A (Suc n) \<inter> range f - f`{.. i})"
       using l A by auto
     then have "\<exists>x. x \<in> A (Suc n) \<inter> range f - f`{.. i}"
-      unfolding ex_in_conv by (intro notI) simp
-    then have "\<exists>j. f j \<in> A (Suc n) \<and> j \<notin> {.. i}"
-      by auto
+      by (metis all_not_in_conv finite.emptyI)
     then have "\<exists>a. i < a \<and> f a \<in> A (Suc n)"
-      by (auto simp: not_le)
+      by (force simp: linorder_not_le)
     then have "i < s n i" "f (s n i) \<in> A (Suc n)"
       unfolding s_def by (auto intro: someI2_ex)
   }
@@ -728,17 +710,15 @@
 
 lemma sequence_unique_limpt:
   fixes f :: "nat \<Rightarrow> 'a::t2_space"
-  assumes "(f \<longlongrightarrow> l) sequentially"
+  assumes f: "(f \<longlongrightarrow> l) sequentially"
     and "l' islimpt (range f)"
   shows "l' = l"
 proof (rule ccontr)
   assume "l' \<noteq> l"
   obtain s t where "open s" "open t" "l' \<in> s" "l \<in> t" "s \<inter> t = {}"
     using hausdorff [OF \<open>l' \<noteq> l\<close>] by auto
-  have "eventually (\<lambda>n. f n \<in> t) sequentially"
-    using assms(1) \<open>open t\<close> \<open>l \<in> t\<close> by (rule topological_tendstoD)
   then obtain N where "\<forall>n\<ge>N. f n \<in> t"
-    unfolding eventually_sequentially by auto
+    by (meson f lim_explicit)
 
   have "UNIV = {..<N} \<union> {N..}"
     by auto
@@ -752,10 +732,8 @@
     using \<open>l' \<in> s\<close> \<open>open s\<close> by (rule islimptE)
   then obtain n where "N \<le> n" "f n \<in> s" "f n \<noteq> l'"
     by auto
-  with \<open>\<forall>n\<ge>N. f n \<in> t\<close> have "f n \<in> s \<inter> t"
-    by simp
-  with \<open>s \<inter> t = {}\<close> show False
-    by simp
+  with \<open>\<forall>n\<ge>N. f n \<in> t\<close> \<open>s \<inter> t = {}\<close> show False
+    by blast
 qed
 
 (*could prove directly from islimpt_sequential_inj, but only for metric spaces*)
@@ -826,18 +804,9 @@
 qed
 
 lemma islimpt_image:
-  assumes "z islimpt g -` A \<inter> B" "g z \<notin> A" "z \<in> B" "open B" "continuous_on B g"
+  assumes "z islimpt g -` A \<inter> B" "g z \<notin> A" "z \<in> B" "continuous_on B g"
   shows   "g z islimpt A"
-  unfolding islimpt_def
-proof clarify
-  fix T assume T: "g z \<in> T" "open T"
-  have "z \<in> g -` T \<inter> B"
-    using T assms by auto
-  moreover have "open (g -` T \<inter> B)"
-    using T continuous_on_open_vimage assms by blast
-  ultimately show "\<exists>y\<in>A. y \<in> T \<and> y \<noteq> g z"
-    using assms by (metis (mono_tags, lifting) IntD1 islimptE vimageE)
-qed
+  by (smt (verit, best) IntD1 assms continuous_on_topological inf_le2 islimpt_def subset_eq vimageE)
   
 
 subsection \<open>Interior of a Set\<close>
@@ -905,12 +874,12 @@
   fixes x :: "'a::perfect_space"
   assumes x: "x \<in> interior S"
   shows "x islimpt S"
-  using x islimpt_UNIV [of x]
-  unfolding interior_def islimpt_def
-  apply (clarsimp, rename_tac T T')
-  apply (drule_tac x="T \<inter> T'" in spec)
-  apply (auto simp: open_Int)
-  done
+proof -
+  obtain T where "x \<in> T" "T \<subseteq> S" "open T"
+    using interior_subset x by blast
+  with x islimpt_UNIV [of x]  show ?thesis
+    unfolding islimpt_def by (metis (full_types) Int_iff open_Int subsetD)
+qed
 
 lemma open_imp_islimpt:
   fixes x::"'a:: perfect_space"
@@ -950,12 +919,8 @@
       assume "x \<notin> interior S"
       with \<open>x \<in> R\<close> \<open>open R\<close> obtain y where "y \<in> R - S"
         unfolding interior_def by fast
-      from \<open>open R\<close> \<open>closed S\<close> have "open (R - S)"
-        by (rule open_Diff)
-      from \<open>R \<subseteq> S \<union> T\<close> have "R - S \<subseteq> T"
-        by fast
-      from \<open>y \<in> R - S\<close> \<open>open (R - S)\<close> \<open>R - S \<subseteq> T\<close> \<open>interior T = {}\<close> show False
-        unfolding interior_def by fast
+      then show False
+        by (metis Diff_subset_conv \<open>R \<subseteq> S \<union> T\<close> \<open>open R\<close> cS empty_iff iT interiorI open_Diff)
     qed
   qed
 qed
@@ -1039,7 +1004,7 @@
 subsection \<open>Closure of a Set\<close>
 
 definition\<^marker>\<open>tag important\<close> closure :: "('a::topological_space) set \<Rightarrow> 'a set" where
-"closure S = S \<union> {x . x islimpt S}"
+  "closure S = S \<union> {x . x islimpt S}"
 
 lemma interior_closure: "interior S = - (closure (- S))"
   by (auto simp: interior_def closure_def islimpt_def)
@@ -1102,18 +1067,8 @@
 proof
   fix x
   assume *: "open S" "x \<in> S \<inter> closure T"
-  have "x islimpt (S \<inter> T)" if **: "x islimpt T"
-  proof (rule islimptI)
-    fix A
-    assume "x \<in> A" "open A"
-    with * have "x \<in> A \<inter> S" "open (A \<inter> S)"
-      by (simp_all add: open_Int)
-    with ** obtain y where "y \<in> T" "y \<in> A \<inter> S" "y \<noteq> x"
-      by (rule islimptE)
-    then have "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"
-      by simp_all
-    then show "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..
-  qed
+  then have "x islimpt (S \<inter> T)" if "x islimpt T"
+    by (metis IntD1 eventually_at_in_open' inf_commute islimpt_Int_eventually that)
   with * show "x \<in> closure (S \<inter> T)"
     unfolding closure_def by blast
 qed
@@ -1148,35 +1103,10 @@
 lemma closure_open_Int_superset:
   assumes "open S" "S \<subseteq> closure T"
   shows "closure(S \<inter> T) = closure S"
-proof -
-  have "closure S \<subseteq> closure(S \<inter> T)"
-    by (metis assms closed_closure closure_minimal inf.orderE open_Int_closure_subset)
-  then show ?thesis
-    by (simp add: closure_mono dual_order.antisym)
-qed
+  by (metis Int_Un_distrib Un_Int_eq(4) assms closure_Un closure_closure open_Int_closure_subset sup.orderE)
 
-lemma closure_Int: "closure (\<Inter>I) \<le> \<Inter>{closure S |S. S \<in> I}"
-proof -
-  {
-    fix y
-    assume "y \<in> \<Inter>I"
-    then have y: "\<forall>S \<in> I. y \<in> S" by auto
-    {
-      fix S
-      assume "S \<in> I"
-      then have "y \<in> closure S"
-        using closure_subset y by auto
-    }
-    then have "y \<in> \<Inter>{closure S |S. S \<in> I}"
-      by auto
-  }
-  then have "\<Inter>I \<subseteq> \<Inter>{closure S |S. S \<in> I}"
-    by auto
-  moreover have "closed (\<Inter>{closure S |S. S \<in> I})"
-    unfolding closed_Inter closed_closure by auto
-  ultimately show ?thesis using closure_hull[of "\<Inter>I"]
-    hull_minimal[of "\<Inter>I" "\<Inter>{closure S |S. S \<in> I}" "closed"] by auto
-qed
+lemma closure_Int: "closure (\<Inter>I) \<subseteq> \<Inter>{closure S |S. S \<in> I}"
+  by (simp add: INF_greatest Inter_lower Setcompr_eq_image closure_mono)
 
 lemma islimpt_in_closure: "(x islimpt S) = (x\<in>closure(S-{x}))"
   unfolding closure_def using islimpt_punctured by blast
@@ -1202,7 +1132,7 @@
 subsection \<open>Frontier (also known as boundary)\<close>
 
 definition\<^marker>\<open>tag important\<close> frontier :: "('a::topological_space) set \<Rightarrow> 'a set" where
-"frontier S = closure S - interior S"
+  "frontier S = closure S - interior S"
 
 lemma frontier_closed [iff]: "closed (frontier S)"
   by (simp add: frontier_def closed_Diff)
@@ -1224,14 +1154,7 @@
 lemma frontier_Int_closed:
   assumes "closed S" "closed T"
   shows "frontier(S \<inter> T) = (frontier S \<inter> T) \<union> (S \<inter> frontier T)"
-proof -
-  have "closure (S \<inter> T) = T \<inter> S"
-    using assms by (simp add: Int_commute closed_Int)
-  moreover have "T \<inter> (closure S \<inter> closure (- S)) = frontier S \<inter> T"
-    by (simp add: Int_commute frontier_closures)
-  ultimately show ?thesis
-    by (simp add: Int_Un_distrib Int_assoc Int_left_commute assms frontier_closures)
-qed
+  by (smt (verit, best) Diff_Int Int_Diff assms closed_Int closure_eq frontier_def inf_commute interior_Int)
 
 lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
   by (metis frontier_def closure_closed Diff_subset)
@@ -1240,16 +1163,7 @@
   by (simp add: frontier_def)
 
 lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
-proof -
-  {
-    assume "frontier S \<subseteq> S"
-    then have "closure S \<subseteq> S"
-      using interior_subset unfolding frontier_def by auto
-    then have "closed S"
-      using closure_subset_eq by auto
-  }
-  then show ?thesis using frontier_subset_closed[of S] ..
-qed
+  by (metis Diff_subset_conv closure_subset_eq frontier_def interior_subset subset_Un_eq)
 
 lemma frontier_complement [simp]: "frontier (- S) = frontier S"
   by (auto simp: frontier_def closure_complement interior_complement)
@@ -1271,12 +1185,7 @@
   by (simp add: Diff_mono frontier_interiors interior_mono interior_subset)
 
 lemma closure_Un_frontier: "closure S = S \<union> frontier S"
-proof -
-  have "S \<union> interior S = S"
-    using interior_subset by auto
-  then show ?thesis
-    using closure_subset by (auto simp: frontier_def)
-qed
+  by (simp add: closure_def frontier_closures sup_inf_distrib1)
 
 
 subsection\<^marker>\<open>tag unimportant\<close> \<open>Filters and the ``eventually true'' quantifier\<close>
@@ -1284,22 +1193,8 @@
 text \<open>Identify Trivial limits, where we can't approach arbitrarily closely.\<close>
 
 lemma trivial_limit_within: "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
-proof
-  assume "trivial_limit (at a within S)"
-  then show "\<not> a islimpt S"
-    unfolding trivial_limit_def
-    unfolding eventually_at_topological
-    unfolding islimpt_def
-    apply (clarsimp simp add: set_eq_iff)
-    apply (rename_tac T, rule_tac x=T in exI)
-    apply (clarsimp, drule_tac x=y in bspec, simp_all)
-    done
-next
-  assume "\<not> a islimpt S"
-  then show "trivial_limit (at a within S)"
     unfolding trivial_limit_def eventually_at_topological islimpt_def
-    by metis
-qed
+    by blast
 
 lemma trivial_limit_at_iff: "trivial_limit (at a) \<longleftrightarrow> \<not> a islimpt UNIV"
   using trivial_limit_within [of a UNIV] by simp
@@ -1312,13 +1207,13 @@
   using islimpt_in_closure by (metis trivial_limit_within)
 
 lemma not_in_closure_trivial_limitI:
-  "x \<notin> closure s \<Longrightarrow> trivial_limit (at x within s)"
-  using not_trivial_limit_within[of x s]
+  "x \<notin> closure S \<Longrightarrow> trivial_limit (at x within S)"
+  using not_trivial_limit_within[of x S]
   by safe (metis Diff_empty Diff_insert0 closure_subset contra_subsetD)
 
-lemma filterlim_at_within_closure_implies_filterlim: "filterlim f l (at x within s)"
-  if "x \<in> closure s \<Longrightarrow> filterlim f l (at x within s)"
-  by (metis bot.extremum filterlim_filtercomap filterlim_mono not_in_closure_trivial_limitI that)
+lemma filterlim_at_within_closure_implies_filterlim: "filterlim f l (at x within S)"
+  if "x \<in> closure S \<Longrightarrow> filterlim f l (at x within S)"
+  by (metis bot.extremum filterlim_iff_le_filtercomap not_in_closure_trivial_limitI that)
 
 lemma at_within_eq_bot_iff: "at c within A = bot \<longleftrightarrow> c \<notin> closure (A - {c})"
   using not_trivial_limit_within[of c A] by blast
@@ -1345,23 +1240,7 @@
 lemma eventually_within_interior:
   assumes "x \<in> interior S"
   shows "eventually P (at x within S) \<longleftrightarrow> eventually P (at x)"
-  (is "?lhs = ?rhs")
-proof
-  from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S" ..
-  {
-    assume ?lhs
-    then obtain A where "open A" and "x \<in> A" and "\<forall>y\<in>A. y \<noteq> x \<longrightarrow> y \<in> S \<longrightarrow> P y"
-      by (auto simp: eventually_at_topological)
-    with T have "open (A \<inter> T)" and "x \<in> A \<inter> T" and "\<forall>y \<in> A \<inter> T. y \<noteq> x \<longrightarrow> P y"
-      by auto
-    then show ?rhs
-      by (auto simp: eventually_at_topological)
-  next
-    assume ?rhs
-    then show ?lhs
-      by (auto elim: eventually_mono simp: eventually_at_filter)
-  }
-qed
+  by (metis assms at_within_open_subset interior_subset open_interior)
 
 lemma at_within_interior: "x \<in> interior S \<Longrightarrow> at x within S = at x"
   unfolding filter_eq_iff by (intro allI eventually_within_interior)
@@ -1443,25 +1322,7 @@
 lemma closure_sequential:
   fixes l :: "'a::first_countable_topology"
   shows "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x \<longlongrightarrow> l) sequentially)"
-  (is "?lhs = ?rhs")
-proof
-  assume "?lhs"
-  moreover
-  {
-    assume "l \<in> S"
-    then have "?rhs" using tendsto_const[of l sequentially] by auto
-  }
-  moreover
-  {
-    assume "l islimpt S"
-    then have "?rhs" unfolding islimpt_sequential by auto
-  }
-  ultimately show "?rhs"
-    unfolding closure_def by auto
-next
-  assume "?rhs"
-  then show "?lhs" unfolding closure_def islimpt_sequential by auto
-qed
+  by (metis Diff_empty Diff_insert0 Un_iff closure_def islimpt_sequential mem_Collect_eq tendsto_const)
 
 lemma closed_sequential_limits:
   fixes S :: "'a::first_countable_topology set"
@@ -1469,20 +1330,20 @@
 by (metis closure_sequential closure_subset_eq subset_iff)
 
 lemma tendsto_If_within_closures:
-  assumes f: "x \<in> s \<union> (closure s \<inter> closure t) \<Longrightarrow>
-      (f \<longlongrightarrow> l x) (at x within s \<union> (closure s \<inter> closure t))"
-  assumes g: "x \<in> t \<union> (closure s \<inter> closure t) \<Longrightarrow>
-      (g \<longlongrightarrow> l x) (at x within t \<union> (closure s \<inter> closure t))"
-  assumes "x \<in> s \<union> t"
-  shows "((\<lambda>x. if x \<in> s then f x else g x) \<longlongrightarrow> l x) (at x within s \<union> t)"
+  assumes f: "x \<in> S \<union> (closure S \<inter> closure T) \<Longrightarrow>
+      (f \<longlongrightarrow> l x) (at x within S \<union> (closure S \<inter> closure T))"
+  assumes g: "x \<in> T \<union> (closure S \<inter> closure T) \<Longrightarrow>
+      (g \<longlongrightarrow> l x) (at x within T \<union> (closure S \<inter> closure T))"
+  assumes "x \<in> S \<union> T"
+  shows "((\<lambda>x. if x \<in> S then f x else g x) \<longlongrightarrow> l x) (at x within S \<union> T)"
 proof -
-  have *: "(s \<union> t) \<inter> {x. x \<in> s} = s" "(s \<union> t) \<inter> {x. x \<notin> s} = t - s"
+  have *: "(S \<union> T) \<inter> {x. x \<in> S} = S" "(S \<union> T) \<inter> {x. x \<notin> S} = T - S"
     by auto
-  have "(f \<longlongrightarrow> l x) (at x within s)"
+  have "(f \<longlongrightarrow> l x) (at x within S)"
     by (rule filterlim_at_within_closure_implies_filterlim)
        (use \<open>x \<in> _\<close> in \<open>auto simp: inf_commute closure_def intro: tendsto_within_subset[OF f]\<close>)
   moreover
-  have "(g \<longlongrightarrow> l x) (at x within t - s)"
+  have "(g \<longlongrightarrow> l x) (at x within T - S)"
     by (rule filterlim_at_within_closure_implies_filterlim)
       (use \<open>x \<in> _\<close> in
         \<open>auto intro!: tendsto_within_subset[OF g] simp: closure_def intro: islimpt_subset\<close>)
@@ -1495,76 +1356,50 @@
 
 lemma brouwer_compactness_lemma:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "compact s"
-    and "continuous_on s f"
-    and "\<not> (\<exists>x\<in>s. f x = 0)"
-  obtains d where "0 < d" and "\<forall>x\<in>s. d \<le> norm (f x)"
-proof (cases "s = {}")
+  assumes "compact S"
+    and "continuous_on S f"
+    and "\<not> (\<exists>x\<in>S. f x = 0)"
+  obtains d where "0 < d" and "\<forall>x\<in>S. d \<le> norm (f x)"
+proof (cases "S = {}")
   case True
   show thesis
     by (rule that [of 1]) (auto simp: True)
 next
   case False
-  have "continuous_on s (norm \<circ> f)"
+  have "continuous_on S (norm \<circ> f)"
     by (rule continuous_intros continuous_on_norm assms(2))+
-  with False obtain x where x: "x \<in> s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y"
+  with False obtain x where x: "x \<in> S" "\<forall>y\<in>S. (norm \<circ> f) x \<le> (norm \<circ> f) y"
     using continuous_attains_inf[OF assms(1), of "norm \<circ> f"]
     unfolding o_def
     by auto
-  have "(norm \<circ> f) x > 0"
-    using assms(3) and x(1)
-    by auto
   then show ?thesis
-    by (rule that) (insert x(2), auto simp: o_def)
+    by (metis assms(3) that comp_apply zero_less_norm_iff)
 qed
 
 subsubsection \<open>Bolzano-Weierstrass property\<close>
 
 proposition Heine_Borel_imp_Bolzano_Weierstrass:
-  assumes "compact s"
-    and "infinite t"
-    and "t \<subseteq> s"
-  shows "\<exists>x \<in> s. x islimpt t"
+  assumes "compact S"
+    and "infinite T"
+    and "T \<subseteq> S"
+  shows "\<exists>x \<in> S. x islimpt T"
 proof (rule ccontr)
-  assume "\<not> (\<exists>x \<in> s. x islimpt t)"
-  then obtain f where f: "\<forall>x\<in>s. x \<in> f x \<and> open (f x) \<and> (\<forall>y\<in>t. y \<in> f x \<longrightarrow> y = x)"
-    unfolding islimpt_def
-    using bchoice[of s "\<lambda> x T. x \<in> T \<and> open T \<and> (\<forall>y\<in>t. y \<in> T \<longrightarrow> y = x)"]
-    by auto
-  obtain g where g: "g \<subseteq> {t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
-    using assms(1)[unfolded compact_eq_Heine_Borel, THEN spec[where x="{t. \<exists>x. x\<in>s \<and> t = f x}"]]
+  assume "\<not> (\<exists>x \<in> S. x islimpt T)"
+  then obtain f where f: "\<forall>x\<in>S. x \<in> f x \<and> open (f x) \<and> (\<forall>y\<in>T. y \<in> f x \<longrightarrow> y = x)"
+    unfolding islimpt_def by metis
+  obtain g where g: "g \<subseteq> {T. \<exists>x. x \<in> S \<and> T = f x}" "finite g" "S \<subseteq> \<Union>g"
+    using assms(1)[unfolded compact_eq_Heine_Borel, THEN spec[where x="{T. \<exists>x. x\<in>S \<and> T = f x}"]]
     using f by auto
-  from g(1,3) have g':"\<forall>x\<in>g. \<exists>xa \<in> s. x = f xa"
+  then have g': "\<forall>x\<in>g. \<exists>y \<in> S. x = f y"
     by auto
-  {
-    fix x y
-    assume "x \<in> t" "y \<in> t" "f x = f y"
-    then have "x \<in> f x"  "y \<in> f x \<longrightarrow> y = x"
-      using f[THEN bspec[where x=x]] and \<open>t \<subseteq> s\<close> by auto
-    then have "x = y"
-      using \<open>f x = f y\<close> and f[THEN bspec[where x=y]] and \<open>y \<in> t\<close> and \<open>t \<subseteq> s\<close>
-      by auto
-  }
-  then have "inj_on f t"
-    unfolding inj_on_def by simp
-  then have "infinite (f ` t)"
+  have "inj_on f T"
+    by (smt (verit, best) assms(3) f inj_onI subsetD)
+  then have "infinite (f ` T)"
     using assms(2) using finite_imageD by auto
   moreover
-  {
-    fix x
-    assume "x \<in> t" "f x \<notin> g"
-    from g(3) assms(3) \<open>x \<in> t\<close> obtain h where "h \<in> g" and "x \<in> h"
-      by auto
-    then obtain y where "y \<in> s" "h = f y"
-      using g'[THEN bspec[where x=h]] by auto
-    then have "y = x"
-      using f[THEN bspec[where x=y]] and \<open>x\<in>t\<close> and \<open>x\<in>h\<close>[unfolded \<open>h = f y\<close>]
-      by auto
-    then have False
-      using \<open>f x \<notin> g\<close> \<open>h \<in> g\<close> unfolding \<open>h = f y\<close>
-      by auto
-  }
-  then have "f ` t \<subseteq> g" by auto
+    have False if "x \<in> T" "f x \<notin> g" for x
+      by (smt (verit) UnionE assms(3) f g' g(3) subsetD that)
+  then have "f ` T \<subseteq> g" by auto
   ultimately show False
     using g(2) using finite_subset by auto
 qed
@@ -1585,25 +1420,24 @@
 qed
 
 lemma Bolzano_Weierstrass_imp_closed:
-  fixes s :: "'a::{first_countable_topology,t2_space} set"
-  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
-  shows "closed s"
+  fixes S :: "'a::{first_countable_topology,t2_space} set"
+  assumes "\<forall>T. infinite T \<and> T \<subseteq> S --> (\<exists>x \<in> S. x islimpt T)"
+  shows "closed S"
 proof -
   {
     fix x l
-    assume as: "\<forall>n::nat. x n \<in> s" "(x \<longlongrightarrow> l) sequentially"
-    then have "l \<in> s"
+    assume as: "\<forall>n::nat. x n \<in> S" "(x \<longlongrightarrow> l) sequentially"
+    have "l \<in> S"
     proof (cases "\<forall>n. x n \<noteq> l")
       case False
-      then show "l\<in>s" using as(1) by auto
+      then show "l\<in>S" using as(1) by auto
     next
-      case True note cas = this
+      case True
       with as(2) have "infinite (range x)"
         using sequence_infinite_lemma[of x l] by auto
-      then obtain l' where "l'\<in>s" "l' islimpt (range x)"
-        using assms[THEN spec[where x="range x"]] as(1) by auto
-      then show "l\<in>s" using sequence_unique_limpt[of x l l']
-        using as cas by auto
+      then obtain l' where "l'\<in>S" "l' islimpt (range x)"
+        using as(1) assms by auto
+      then show "l\<in>S" using sequence_unique_limpt as True by auto
     qed
   }
   then show ?thesis
@@ -1612,38 +1446,30 @@
 
 lemma closure_insert:
   fixes x :: "'a::t1_space"
-  shows "closure (insert x s) = insert x (closure s)"
-  by (meson closed_closure closed_insert closure_minimal closure_subset dual_order.antisym insert_mono insert_subset)
+  shows "closure (insert x S) = insert x (closure S)"
+  by (metis closed_singleton closure_Un closure_closed insert_is_Un)
 
 lemma finite_not_islimpt_in_compact:
   assumes "compact A" "\<And>z. z \<in> A \<Longrightarrow> \<not>z islimpt B"
   shows   "finite (A \<inter> B)"
-proof (rule ccontr)
-  assume "infinite (A \<inter> B)"
-  have "\<exists>z\<in>A. z islimpt A \<inter> B"
-    by (rule Heine_Borel_imp_Bolzano_Weierstrass) (use assms \<open>infinite _\<close> in auto)
-  hence "\<exists>z\<in>A. z islimpt B"
-    using islimpt_subset by blast
-  thus False using assms(2)
-    by simp
-qed
+  by (meson Heine_Borel_imp_Bolzano_Weierstrass assms inf_le1 inf_le2 islimpt_subset)
 
 
 text\<open>In particular, some common special cases.\<close>
 
 lemma compact_Un [intro]:
-  assumes "compact s"
-    and "compact t"
-  shows " compact (s \<union> t)"
+  assumes "compact S"
+    and "compact T"
+  shows " compact (S \<union> T)"
 proof (rule compactI)
   fix f
-  assume *: "Ball f open" "s \<union> t \<subseteq> \<Union>f"
-  from * \<open>compact s\<close> obtain s' where "s' \<subseteq> f \<and> finite s' \<and> s \<subseteq> \<Union>s'"
+  assume *: "Ball f open" "S \<union> T \<subseteq> \<Union>f"
+  from * \<open>compact S\<close> obtain s' where "s' \<subseteq> f \<and> finite s' \<and> S \<subseteq> \<Union>s'"
     unfolding compact_eq_Heine_Borel by (auto elim!: allE[of _ f])
   moreover
-  from * \<open>compact t\<close> obtain t' where "t' \<subseteq> f \<and> finite t' \<and> t \<subseteq> \<Union>t'"
+  from * \<open>compact T\<close> obtain t' where "t' \<subseteq> f \<and> finite t' \<and> T \<subseteq> \<Union>t'"
     unfolding compact_eq_Heine_Borel by (auto elim!: allE[of _ f])
-  ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<union> t \<subseteq> \<Union>f'"
+  ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> S \<union> T \<subseteq> \<Union>f'"
     by (auto intro!: exI[of _ "s' \<union> t'"])
 qed
 
@@ -1652,40 +1478,34 @@
 
 lemma compact_UN [intro]:
   "finite A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> compact (B x)) \<Longrightarrow> compact (\<Union>x\<in>A. B x)"
-  by (rule compact_Union) auto
+  by blast
 
 lemma closed_Int_compact [intro]:
-  assumes "closed s"
-    and "compact t"
-  shows "compact (s \<inter> t)"
-  using compact_Int_closed [of t s] assms
+  assumes "closed S" and "compact T"
+  shows "compact (S \<inter> T)"
+  using compact_Int_closed [of T S] assms
   by (simp add: Int_commute)
 
 lemma compact_Int [intro]:
-  fixes s t :: "'a :: t2_space set"
-  assumes "compact s"
-    and "compact t"
-  shows "compact (s \<inter> t)"
+  fixes S T :: "'a :: t2_space set"
+  assumes "compact S" and "compact T"
+  shows "compact (S \<inter> T)"
   using assms by (intro compact_Int_closed compact_imp_closed)
 
 lemma compact_sing [simp]: "compact {a}"
   unfolding compact_eq_Heine_Borel by auto
 
 lemma compact_insert [simp]:
-  assumes "compact s"
-  shows "compact (insert x s)"
-proof -
-  have "compact ({x} \<union> s)"
-    using compact_sing assms by (rule compact_Un)
-  then show ?thesis by simp
-qed
+  assumes "compact S"
+  shows "compact (insert x S)"
+  by (metis assms compact_Un compact_sing insert_is_Un)
 
-lemma finite_imp_compact: "finite s \<Longrightarrow> compact s"
+lemma finite_imp_compact: "finite S \<Longrightarrow> compact S"
   by (induct set: finite) simp_all
 
 lemma open_delete:
-  fixes s :: "'a::t1_space set"
-  shows "open s \<Longrightarrow> open (s - {x})"
+  fixes S :: "'a::t1_space set"
+  shows "open S \<Longrightarrow> open (S - {x})"
   by (simp add: open_Diff)
 
 
@@ -1696,20 +1516,17 @@
 proof safe
   assume x: "x \<in> closure X"
   fix S A
-  assume "open S" "x \<in> S" "X \<inter> A = {}" "S \<subseteq> A"
+  assume \<section>: "open S" "x \<in> S" "X \<inter> A = {}" "S \<subseteq> A"
   then have "x \<notin> closure (-S)"
-    by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI)
+    by (simp add: closed_open)
   with x have "x \<in> closure X - closure (-S)"
     by auto
-  also have "\<dots> \<subseteq> closure (X \<inter> S)"
-    using \<open>open S\<close> open_Int_closure_subset[of S X] by (simp add: closed_Compl ac_simps)
-  finally have "X \<inter> S \<noteq> {}" by auto
-  then show False using \<open>X \<inter> A = {}\<close> \<open>S \<subseteq> A\<close> by auto
+  with \<section> show False
+    by (metis Compl_iff Diff_iff closure_mono disjoint_iff subsetD subsetI)
 next
   assume "\<forall>A S. S \<subseteq> A \<longrightarrow> open S \<longrightarrow> x \<in> S \<longrightarrow> X \<inter> A \<noteq> {}"
-  from this[THEN spec, of "- X", THEN spec, of "- closure X"]
-  show "x \<in> closure X"
-    by (simp add: closure_subset open_Compl)
+  then show "x \<in> closure X"
+    by (metis ComplI Compl_disjoint closure_interior interior_subset open_interior)
 qed
 
 lemma compact_filter:
@@ -1745,7 +1562,7 @@
   proof safe
     fix P Q R S
     assume "eventually R F" "open S" "x \<in> S"
-    with open_Int_closure_eq_empty[of S "{x. R x}"] x[of "closure {x. R x}"]
+    with open_Int_closure_eq_empty[of S "{x. R x}"] x
     have "S \<inter> {x. R x} \<noteq> {}" by (auto simp: Z_def)
     moreover assume "Ball S Q" "\<forall>x. Q x \<and> R x \<longrightarrow> bot x"
     ultimately show False by (auto simp: set_eq_iff)
@@ -1871,9 +1688,8 @@
 subsubsection\<open>Sequential compactness\<close>
 
 definition\<^marker>\<open>tag important\<close> seq_compact :: "'a::topological_space set \<Rightarrow> bool" where
-"seq_compact S \<longleftrightarrow>
-  (\<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) \<longlongrightarrow> l) sequentially))"
+  "seq_compact S \<longleftrightarrow>
+    (\<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"
@@ -1886,38 +1702,30 @@
   using assms unfolding seq_compact_def by fast
 
 lemma closed_sequentially: (* TODO: move upwards *)
-  assumes "closed s" and "\<forall>n. f n \<in> s" and "f \<longlonglongrightarrow> l"
-  shows "l \<in> s"
-proof (rule ccontr)
-  assume "l \<notin> s"
-  with \<open>closed s\<close> and \<open>f \<longlonglongrightarrow> l\<close> have "eventually (\<lambda>n. f n \<in> - s) sequentially"
-    by (fast intro: topological_tendstoD)
-  with \<open>\<forall>n. f n \<in> s\<close> show "False"
-    by simp
-qed
+  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)"
+  assumes "seq_compact S" and "closed T"
+  shows "seq_compact (S \<inter> T)"
 proof (rule seq_compactI)
-  fix f assume "\<forall>n::nat. f n \<in> s \<inter> t"
-  hence "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"
+  fix f assume "\<forall>n::nat. f n \<in> S \<inter> T"
+  hence "\<forall>n. f n \<in> S" and "\<forall>n. f n \<in> T"
     by simp_all
-  from \<open>seq_compact s\<close> and \<open>\<forall>n. f n \<in> s\<close>
-  obtain l r where "l \<in> s" and r: "strict_mono r" and l: "(f \<circ> r) \<longlonglongrightarrow> l"
+  from \<open>seq_compact S\<close> and \<open>\<forall>n. f n \<in> S\<close>
+  obtain l r where "l \<in> S" and r: "strict_mono r" and l: "(f \<circ> r) \<longlonglongrightarrow> l"
     by (rule seq_compactE)
-  from \<open>\<forall>n. f n \<in> t\<close> have "\<forall>n. (f \<circ> r) n \<in> t"
+  from \<open>\<forall>n. f n \<in> T\<close> have "\<forall>n. (f \<circ> r) n \<in> T"
     by simp
-  from \<open>closed t\<close> and this and l have "l \<in> t"
-    by (rule closed_sequentially)
-  with \<open>l \<in> s\<close> and r and l show "\<exists>l\<in>s \<inter> t. \<exists>r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
-    by fast
+  with \<open>l \<in> S\<close> and r and l show "\<exists>l\<in>S \<inter> T. \<exists>r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
+    by (metis Int_iff \<open>closed T\<close> closed_sequentially)
 qed
 
 lemma seq_compact_closed_subset:
-  assumes "closed s" and "s \<subseteq> t" and "seq_compact t"
-  shows "seq_compact s"
-  using assms seq_compact_Int_closed [of t s] by (simp add: Int_absorb1)
+  assumes "closed S" and "S \<subseteq> T" and "seq_compact T"
+  shows "seq_compact S"
+  using assms seq_compact_Int_closed [of T S] by (simp add: Int_absorb1)
 
 lemma seq_compact_imp_countably_compact:
   fixes U :: "'a :: first_countable_topology set"
@@ -1951,9 +1759,9 @@
         by auto
       from \<open>x\<in>U\<close> \<open>U \<subseteq> \<Union>A\<close> from_nat_into_surj[OF \<open>countable A\<close>]
       obtain n where "x \<in> from_nat_into A n" by auto
-      with r(2) A(1) from_nat_into[OF \<open>A \<noteq> {}\<close>, of n]
+      with r(2) A(1) from_nat_into[OF \<open>A \<noteq> {}\<close>]
       have "eventually (\<lambda>i. X (r i) \<in> from_nat_into A n) sequentially"
-        unfolding tendsto_def by (auto simp: comp_def)
+        unfolding tendsto_def by fastforce
       then obtain N where "\<And>i. N \<le> i \<Longrightarrow> X (r i) \<in> from_nat_into A n"
         by (auto simp: eventually_sequentially)
       moreover from X have "\<And>i. n \<le> r i \<Longrightarrow> X (r i) \<notin> from_nat_into A n"
@@ -2037,95 +1845,81 @@
 qed
 
 lemma countably_compact_imp_acc_point:
-  assumes "countably_compact s"
-    and "countable t"
-    and "infinite t"
-    and "t \<subseteq> s"
-  shows "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t)"
+  assumes "countably_compact S"
+    and "countable T"
+    and "infinite T"
+    and "T \<subseteq> S"
+  shows "\<exists>x\<in>S. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> T)"
 proof (rule ccontr)
-  define C where "C = (\<lambda>F. interior (F \<union> (- t))) ` {F. finite F \<and> F \<subseteq> t }"
-  note \<open>countably_compact s\<close>
-  moreover have "\<forall>t\<in>C. open t"
+  define C where "C = (\<lambda>F. interior (F \<union> (- T))) ` {F. finite F \<and> F \<subseteq> T }"
+  note \<open>countably_compact S\<close>
+  moreover have "\<forall>T\<in>C. open T"
     by (auto simp: C_def)
   moreover
-  assume "\<not> (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t))"
-  then have s: "\<And>x. x \<in> s \<Longrightarrow> \<exists>U. x\<in>U \<and> open U \<and> finite (U \<inter> t)" by metis
-  have "s \<subseteq> \<Union>C"
-    using \<open>t \<subseteq> s\<close>
+  assume "\<not> (\<exists>x\<in>S. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> T))"
+  then have S: "\<And>x. x \<in> S \<Longrightarrow> \<exists>U. x\<in>U \<and> open U \<and> finite (U \<inter> T)" by metis
+  have "S \<subseteq> \<Union>C"
+    using \<open>T \<subseteq> S\<close>
     unfolding C_def
-    apply (safe dest!: s)
-    apply (rule_tac a="U \<inter> t" in UN_I)
+    apply (safe dest!: S)
+    apply (rule_tac a="U \<inter> T" in UN_I)
     apply (auto intro!: interiorI simp add: finite_subset)
     done
   moreover
-  from \<open>countable t\<close> have "countable C"
+  from \<open>countable T\<close> have "countable C"
     unfolding C_def by (auto intro: countable_Collect_finite_subset)
   ultimately
-  obtain D where "D \<subseteq> C" "finite D" "s \<subseteq> \<Union>D"
+  obtain D where "D \<subseteq> C" "finite D" "S \<subseteq> \<Union>D"
     by (rule countably_compactE)
-  then obtain E where E: "E \<subseteq> {F. finite F \<and> F \<subseteq> t }" "finite E"
-    and s: "s \<subseteq> (\<Union>F\<in>E. interior (F \<union> (- t)))"
+  then obtain E where E: "E \<subseteq> {F. finite F \<and> F \<subseteq> T }" "finite E"
+    and S: "S \<subseteq> (\<Union>F\<in>E. interior (F \<union> (- T)))"
     by (metis (lifting) finite_subset_image C_def)
-  from s \<open>t \<subseteq> s\<close> have "t \<subseteq> \<Union>E"
+  from S \<open>T \<subseteq> S\<close> have "T \<subseteq> \<Union>E"
     using interior_subset by blast
   moreover have "finite (\<Union>E)"
     using E by auto
-  ultimately show False using \<open>infinite t\<close>
+  ultimately show False using \<open>infinite T\<close>
     by (auto simp: finite_subset)
 qed
 
 lemma countable_acc_point_imp_seq_compact:
-  fixes s :: "'a::first_countable_topology set"
-  assumes "\<forall>t. infinite t \<and> countable t \<and> t \<subseteq> s \<longrightarrow>
-    (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t))"
-  shows "seq_compact s"
-proof -
-  {
-    fix f :: "nat \<Rightarrow> 'a"
-    assume f: "\<forall>n. f n \<in> s"
-    have "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
-    proof (cases "finite (range f)")
-      case True
-      obtain l where "infinite {n. f n = f l}"
-        using pigeonhole_infinite[OF _ True] by auto
-      then obtain r :: "nat \<Rightarrow> nat" where "strict_mono  r" and fr: "\<forall>n. f (r n) = f l"
-        using infinite_enumerate by blast
-      then have "strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> f l"
-        by (simp add: fr o_def)
-      with f show "\<exists>l\<in>s. \<exists>r. strict_mono  r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
-        by auto
-    next
-      case False
-      with f assms have "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)"
-        by auto
-      then obtain l where "l \<in> s" "\<forall>U. l\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)" ..
-      from this(2) have "\<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
-        using acc_point_range_imp_convergent_subsequence[of l f] by auto
-      with \<open>l \<in> s\<close> show "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" ..
-    qed
-  }
-  then show ?thesis
-    unfolding seq_compact_def by auto
+  fixes S :: "'a::first_countable_topology set"
+  assumes "\<And>T. \<lbrakk>infinite T; countable T; T \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> T)"
+  shows "seq_compact S"
+  unfolding seq_compact_def
+proof (intro strip)
+  fix f :: "nat \<Rightarrow> 'a"
+  assume f: "\<forall>n. f n \<in> S"
+  show "\<exists>l\<in>S. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
+  proof (cases "finite (range f)")
+    case True
+    obtain l where "infinite {n. f n = f l}"
+      using pigeonhole_infinite[OF _ True] by auto
+    then obtain r :: "nat \<Rightarrow> nat" where "strict_mono  r" and fr: "\<forall>n. f (r n) = f l"
+      using infinite_enumerate by blast
+    then have "strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> f l"
+      by (simp add: fr o_def)
+    with f show "\<exists>l\<in>S. \<exists>r. strict_mono  r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
+      by auto
+  next
+    case False
+    with f assms obtain l where "l \<in> S" "\<forall>U. l\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)"
+      by (metis image_subset_iff uncountable_def) 
+    with \<open>l \<in> S\<close> show "\<exists>l\<in>S. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
+      by (meson acc_point_range_imp_convergent_subsequence) 
+  qed
 qed
 
 lemma seq_compact_eq_countably_compact:
   fixes U :: "'a :: first_countable_topology set"
   shows "seq_compact U \<longleftrightarrow> countably_compact U"
-  using
-    countable_acc_point_imp_seq_compact
-    countably_compact_imp_acc_point
-    seq_compact_imp_countably_compact
-  by metis
+  by (metis countable_acc_point_imp_seq_compact countably_compact_imp_acc_point seq_compact_imp_countably_compact)
 
 lemma seq_compact_eq_acc_point:
-  fixes s :: "'a :: first_countable_topology set"
-  shows "seq_compact s \<longleftrightarrow>
-    (\<forall>t. infinite t \<and> countable t \<and> t \<subseteq> s --> (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t)))"
-  using
-    countable_acc_point_imp_seq_compact[of s]
-    countably_compact_imp_acc_point[of s]
-    seq_compact_imp_countably_compact[of s]
-  by metis
+  fixes S :: "'a :: first_countable_topology set"
+  shows "seq_compact S \<longleftrightarrow>
+    (\<forall>T. infinite T \<and> countable T \<and> T \<subseteq> S --> (\<exists>x\<in>S. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> T)))"
+  by (metis countable_acc_point_imp_seq_compact countably_compact_imp_acc_point seq_compact_imp_countably_compact)
 
 lemma seq_compact_eq_compact:
   fixes U :: "'a :: second_countable_topology set"
@@ -2133,74 +1927,72 @@
   using seq_compact_eq_countably_compact countably_compact_eq_compact by blast
 
 proposition Bolzano_Weierstrass_imp_seq_compact:
-  fixes s :: "'a::{t1_space, first_countable_topology} set"
-  shows "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> seq_compact s"
+  fixes S :: "'a::{t1_space, first_countable_topology} set"
+  shows "(\<And>T. \<lbrakk>infinite T; T \<subseteq> S\<rbrakk> \<Longrightarrow>\<exists>x \<in> S. x islimpt T) \<Longrightarrow> seq_compact S"
   by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point)
 
 
 subsection\<^marker>\<open>tag unimportant\<close> \<open>Cartesian products\<close>
 
-lemma seq_compact_Times: "seq_compact s \<Longrightarrow> seq_compact t \<Longrightarrow> seq_compact (s \<times> t)"
+lemma seq_compact_Times:
+  assumes "seq_compact S" "seq_compact T"
+  shows "seq_compact (S \<times> T)"
   unfolding seq_compact_def
-  apply clarify
-  apply (drule_tac x="fst \<circ> f" in spec)
-  apply (drule mp, simp add: mem_Times_iff)
-  apply (clarify, rename_tac l1 r1)
-  apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)
-  apply (drule mp, simp add: mem_Times_iff)
-  apply (clarify, rename_tac l2 r2)
-  apply (rule_tac x="(l1, l2)" in rev_bexI, simp)
-  apply (rule_tac x="r1 \<circ> r2" in exI)
-  apply (rule conjI, simp add: strict_mono_def)
-  apply (drule_tac f=r2 in LIMSEQ_subseq_LIMSEQ, assumption)
-  apply (drule (1) tendsto_Pair) back
-  apply (simp add: o_def)
-  done
+proof clarify
+  fix h :: "nat \<Rightarrow> 'a \<times> 'b"
+  assume "\<forall>n. h n \<in> S \<times> T"
+  then have *: "\<And>n. (fst \<circ> h) n \<in> S" "\<And>n. (snd \<circ> h) n \<in> T"
+    by (simp_all add: mem_Times_iff)
+  then obtain lS and rS :: "nat\<Rightarrow>nat"
+    where "lS\<in>S" "strict_mono rS" and lS: "(fst \<circ> h \<circ> rS) \<longlonglongrightarrow> lS"
+    using assms seq_compact_def by metis
+  then obtain lT and rT :: "nat\<Rightarrow>nat" 
+    where  "lT\<in>T" "strict_mono rT" and lT: "(snd \<circ> h \<circ> rS \<circ> rT) \<longlonglongrightarrow> lT"
+    using assms seq_compact_def *
+    by (metis (mono_tags, lifting) comp_apply) 
+  have "strict_mono (rS \<circ> rT)"
+    by (simp add: \<open>strict_mono rS\<close> \<open>strict_mono rT\<close> strict_mono_o)
+  moreover have "(h \<circ> (rS \<circ> rT)) \<longlonglongrightarrow> (lS,lT)"
+    using tendsto_Pair [OF LIMSEQ_subseq_LIMSEQ [OF lS \<open>strict_mono rT\<close>] lT]
+    by (simp add: o_def)
+  ultimately show "\<exists>l\<in>S \<times> T. \<exists>r. strict_mono r \<and> (h \<circ> r) \<longlonglongrightarrow> l"
+    using \<open>lS \<in> S\<close> \<open>lT \<in> T\<close> by blast
+qed
 
 lemma compact_Times:
-  assumes "compact s" "compact t"
-  shows "compact (s \<times> t)"
+  assumes "compact S" "compact T"
+  shows "compact (S \<times> T)"
 proof (rule compactI)
-  fix C
-  assume C: "\<forall>t\<in>C. open t" "s \<times> t \<subseteq> \<Union>C"
-  have "\<forall>x\<in>s. \<exists>a. open a \<and> x \<in> a \<and> (\<exists>d\<subseteq>C. finite d \<and> a \<times> t \<subseteq> \<Union>d)"
+  fix \<C>
+  assume C: "\<forall>T\<in>\<C>. open T" "S \<times> T \<subseteq> \<Union>\<C>"
+  have "\<forall>x\<in>S. \<exists>A. open A \<and> x \<in> A \<and> (\<exists>D\<subseteq>\<C>. finite D \<and> A \<times> T \<subseteq> \<Union>D)"
   proof
     fix x
-    assume "x \<in> s"
-    have "\<forall>y\<in>t. \<exists>a b c. c \<in> C \<and> open a \<and> open b \<and> x \<in> a \<and> y \<in> b \<and> a \<times> b \<subseteq> c" (is "\<forall>y\<in>t. ?P y")
-    proof
-      fix y
-      assume "y \<in> t"
-      with \<open>x \<in> s\<close> C obtain c where "c \<in> C" "(x, y) \<in> c" "open c" by auto
-      then show "?P y" by (auto elim!: open_prod_elim)
-    qed
-    then obtain a b c where b: "\<And>y. y \<in> t \<Longrightarrow> open (b y)"
-      and c: "\<And>y. y \<in> t \<Longrightarrow> c y \<in> C \<and> open (a y) \<and> open (b y) \<and> x \<in> a y \<and> y \<in> b y \<and> a y \<times> b y \<subseteq> c y"
+    assume "x \<in> S"
+    have "\<forall>y\<in>T. \<exists>A B C. C \<in> \<C> \<and> open A \<and> open B \<and> x \<in> A \<and> y \<in> B \<and> A \<times> B \<subseteq> C"
+      by (smt (verit, ccfv_threshold) C UnionE \<open>x \<in> S\<close> mem_Sigma_iff open_prod_def subsetD)
+    then obtain a b c where b: "\<And>y. y \<in> T \<Longrightarrow> open (b y)"
+      and c: "\<And>y. y \<in> T \<Longrightarrow> c y \<in> \<C> \<and> open (a y) \<and> open (b y) \<and> x \<in> a y \<and> y \<in> b y \<and> a y \<times> b y \<subseteq> c y"
       by metis
-    then have "\<forall>y\<in>t. open (b y)" "t \<subseteq> (\<Union>y\<in>t. b y)" by auto
-    with compactE_image[OF \<open>compact t\<close>] obtain D where D: "D \<subseteq> t" "finite D" "t \<subseteq> (\<Union>y\<in>D. b y)"
+    then have "\<forall>y\<in>T. open (b y)" "T \<subseteq> (\<Union>y\<in>T. b y)" by auto
+    with compactE_image[OF \<open>compact T\<close>] obtain D where D: "D \<subseteq> T" "finite D" "T \<subseteq> (\<Union>y\<in>D. b y)"
       by metis
-    moreover from D c have "(\<Inter>y\<in>D. a y) \<times> t \<subseteq> (\<Union>y\<in>D. c y)"
+    moreover from D c have "(\<Inter>y\<in>D. a y) \<times> T \<subseteq> (\<Union>y\<in>D. c y)"
       by (fastforce simp: subset_eq)
-    ultimately show "\<exists>a. open a \<and> x \<in> a \<and> (\<exists>d\<subseteq>C. finite d \<and> a \<times> t \<subseteq> \<Union>d)"
+    ultimately show "\<exists>a. open a \<and> x \<in> a \<and> (\<exists>d\<subseteq>\<C>. finite d \<and> a \<times> T \<subseteq> \<Union>d)"
       using c by (intro exI[of _ "c`D"] exI[of _ "\<Inter>(a`D)"] conjI) (auto intro!: open_INT)
   qed
-  then obtain a d where a: "\<And>x. x\<in>s \<Longrightarrow> open (a x)" "s \<subseteq> (\<Union>x\<in>s. a x)"
-    and d: "\<And>x. x \<in> s \<Longrightarrow> d x \<subseteq> C \<and> finite (d x) \<and> a x \<times> t \<subseteq> \<Union>(d x)"
+  then obtain a d where a: "\<And>x. x\<in>S \<Longrightarrow> open (a x)" "S \<subseteq> (\<Union>x\<in>S. a x)"
+    and d: "\<And>x. x \<in> S \<Longrightarrow> d x \<subseteq> \<C> \<and> finite (d x) \<and> a x \<times> T \<subseteq> \<Union>(d x)"
     unfolding subset_eq UN_iff by metis
   moreover
-  from compactE_image[OF \<open>compact s\<close> a]
-  obtain e where e: "e \<subseteq> s" "finite e" and s: "s \<subseteq> (\<Union>x\<in>e. a x)"
+  from compactE_image[OF \<open>compact S\<close> a]
+  obtain e where e: "e \<subseteq> S" "finite e" and S: "S \<subseteq> (\<Union>x\<in>e. a x)"
     by auto
   moreover
-  {
-    from s have "s \<times> t \<subseteq> (\<Union>x\<in>e. a x \<times> t)"
-      by auto
-    also have "\<dots> \<subseteq> (\<Union>x\<in>e. \<Union>(d x))"
-      using d \<open>e \<subseteq> s\<close> by (intro UN_mono) auto
-    finally have "s \<times> t \<subseteq> (\<Union>x\<in>e. \<Union>(d x))" .
-  }
-  ultimately show "\<exists>C'\<subseteq>C. finite C' \<and> s \<times> t \<subseteq> \<Union>C'"
+  have "S \<times> T \<subseteq> (\<Union>x\<in>e. \<Union>(d x))"
+    by (smt (verit, del_insts) S SigmaE UN_iff d e(1) mem_Sigma_iff subset_eq)
+  ultimately show "\<exists>C'\<subseteq>\<C>. finite C' \<and> S \<times> T \<subseteq> \<Union>C'"
     by (intro exI[of _ "(\<Union>x\<in>e. d x)"]) (auto simp: subset_eq)
 qed
 
@@ -2248,9 +2040,9 @@
   have "open {..<e}" by simp
   have "continuous_on (U \<times> C) psi"
     by (auto intro!: continuous_intros simp: psi_def split_beta')
-  from this[unfolded continuous_on_open_invariant, rule_format, OF \<open>open {..<e}\<close>]
-  obtain W where W: "open W" "W \<inter> U \<times> C = W0 \<inter> U \<times> C"
-    unfolding W0_eq by blast
+  then obtain W where W: "open W" "W \<inter> U \<times> C = W0 \<inter> U \<times> C"
+    unfolding W0_eq
+    by (metis \<open>open {..<e}\<close> continuous_on_open_invariant inf_right_idem) 
   have "{x0} \<times> C \<subseteq> W \<inter> U \<times> C"
     unfolding W
     by (auto simp: W0_def psi_def \<open>0 < e\<close>)
@@ -2265,20 +2057,9 @@
     fix t assume t: "t \<in> C"
     have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)"
       by (auto simp: psi_def)
-    also
-    {
-      have "(x, t) \<in> X0 \<times> C"
-        using t x
-        by auto
-      also note \<open>\<dots> \<subseteq> W\<close>
-      finally have "(x, t) \<in> W" .
-      with t x have "(x, t) \<in> W \<inter> U \<times> C"
-        by blast
-      also note \<open>W \<inter> U \<times> C = W0 \<inter> U \<times> C\<close>
-      finally  have "psi (x, t) < e"
-        by (auto simp: W0_def)
-    }
-    finally show "dist (fx (x, t)) (fx (x0, t)) \<le> e" by simp
+    also have "psi (x, t) < e"
+      using W(2) W0_def X0(3) t x by fastforce
+    finally show "dist (fx (x,t)) (fx (x0,t)) \<le> e" by simp
   qed
   from X0(1,2) this show ?thesis ..
 qed
@@ -2296,15 +2077,15 @@
 lemmas continuous_on = continuous_on_def \<comment> \<open>legacy theorem name\<close>
 
 lemma continuous_within_subset:
-  "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s \<Longrightarrow> continuous (at x within t) f"
+  "continuous (at x within S) f \<Longrightarrow> t \<subseteq> S \<Longrightarrow> continuous (at x within t) f"
   unfolding continuous_within by(metis tendsto_within_subset)
 
 lemma continuous_on_interior:
-  "continuous_on s f \<Longrightarrow> x \<in> interior s \<Longrightarrow> continuous (at x) f"
+  "continuous_on S f \<Longrightarrow> x \<in> interior S \<Longrightarrow> continuous (at x) f"
   by (metis continuous_on_eq_continuous_at continuous_on_subset interiorE)
 
 lemma continuous_on_eq:
-  "\<lbrakk>continuous_on s f; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> continuous_on s g"
+  "\<lbrakk>continuous_on S f; \<And>x. x \<in> S \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> continuous_on S g"
   unfolding continuous_on_def tendsto_def eventually_at_topological
   by simp
 
@@ -2312,39 +2093,38 @@
 
 lemma continuous_within_sequentiallyI:
   fixes f :: "'a::{first_countable_topology, t2_space} \<Rightarrow> 'b::topological_space"
-  assumes "\<And>u::nat \<Rightarrow> 'a. u \<longlonglongrightarrow> a \<Longrightarrow> (\<forall>n. u n \<in> s) \<Longrightarrow> (\<lambda>n. f (u n)) \<longlonglongrightarrow> f a"
-  shows "continuous (at a within s) f"
+  assumes "\<And>u::nat \<Rightarrow> 'a. u \<longlonglongrightarrow> a \<Longrightarrow> (\<forall>n. u n \<in> S) \<Longrightarrow> (\<lambda>n. f (u n)) \<longlonglongrightarrow> f a"
+  shows "continuous (at a within S) f"
   using assms unfolding continuous_within tendsto_def[where l = "f a"]
   by (auto intro!: sequentially_imp_eventually_within)
 
 lemma continuous_within_tendsto_compose:
   fixes f::"'a::t2_space \<Rightarrow> 'b::topological_space"
-  assumes "continuous (at a within s) f"
-          "eventually (\<lambda>n. x n \<in> s) F"
-          "(x \<longlongrightarrow> a) F "
+  assumes f: "continuous (at a within S) f"
+         and  "eventually (\<lambda>n. x n \<in> S) F" "(x \<longlongrightarrow> a) F "
   shows "((\<lambda>n. f (x n)) \<longlongrightarrow> f a) F"
 proof -
-  have *: "filterlim x (inf (nhds a) (principal s)) F"
-    using assms(2) assms(3) unfolding at_within_def filterlim_inf by (auto simp: filterlim_principal eventually_mono)
+  have *: "filterlim x (inf (nhds a) (principal S)) F"
+    by (simp add: assms filterlim_inf filterlim_principal)
   show ?thesis
-    by (auto simp: assms(1) continuous_within[symmetric] tendsto_at_within_iff_tendsto_nhds[symmetric] intro!: filterlim_compose[OF _ *])
+    using "*" f continuous_within filterlim_compose tendsto_at_within_iff_tendsto_nhds by blast
 qed
 
 lemma continuous_within_tendsto_compose':
   fixes f::"'a::t2_space \<Rightarrow> 'b::topological_space"
-  assumes "continuous (at a within s) f"
-    "\<And>n. x n \<in> s"
+  assumes "continuous (at a within S) f"
+    "\<And>n. x n \<in> S"
     "(x \<longlongrightarrow> a) F "
   shows "((\<lambda>n. f (x n)) \<longlongrightarrow> f a) F"
-  by (auto intro!: continuous_within_tendsto_compose[OF assms(1)] simp add: assms)
+  using always_eventually assms continuous_within_tendsto_compose by blast
 
 lemma continuous_within_sequentially:
   fixes f :: "'a::{first_countable_topology, t2_space} \<Rightarrow> 'b::topological_space"
-  shows "continuous (at a within s) f \<longleftrightarrow>
-    (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x \<longlongrightarrow> a) sequentially
+  shows "continuous (at a within S) f \<longleftrightarrow>
+    (\<forall>x. (\<forall>n::nat. x n \<in> S) \<and> (x \<longlongrightarrow> a) sequentially
          \<longrightarrow> ((f \<circ> x) \<longlongrightarrow> f a) sequentially)"
-  using continuous_within_tendsto_compose'[of a s f _ sequentially]
-    continuous_within_sequentiallyI[of a s f]
+  using continuous_within_tendsto_compose'[of a S f _ sequentially]
+    continuous_within_sequentiallyI[of a S f]
   by (auto simp: o_def)
 
 lemma continuous_at_sequentiallyI:
@@ -2361,42 +2141,26 @@
 
 lemma continuous_on_sequentiallyI:
   fixes f :: "'a::{first_countable_topology, t2_space} \<Rightarrow> 'b::topological_space"
-  assumes "\<And>u a. (\<forall>n. u n \<in> s) \<Longrightarrow> a \<in> s \<Longrightarrow> u \<longlonglongrightarrow> a \<Longrightarrow> (\<lambda>n. f (u n)) \<longlonglongrightarrow> f a"
-  shows "continuous_on s f"
+  assumes "\<And>u a. (\<forall>n. u n \<in> S) \<Longrightarrow> a \<in> S \<Longrightarrow> u \<longlonglongrightarrow> a \<Longrightarrow> (\<lambda>n. f (u n)) \<longlonglongrightarrow> f a"
+  shows "continuous_on S f"
   using assms unfolding continuous_on_eq_continuous_within
-  using continuous_within_sequentiallyI[of _ s f] by auto
+  using continuous_within_sequentiallyI[of _ S f] by auto
 
 lemma continuous_on_sequentially:
   fixes f :: "'a::{first_countable_topology, t2_space} \<Rightarrow> 'b::topological_space"
-  shows "continuous_on s f \<longleftrightarrow>
-    (\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x \<longlongrightarrow> a) sequentially
-      --> ((f \<circ> x) \<longlongrightarrow> f a) sequentially)"
-    (is "?lhs = ?rhs")
-proof
-  assume ?rhs
-  then show ?lhs
-    using continuous_within_sequentially[of _ s f]
-    unfolding continuous_on_eq_continuous_within
-    by auto
-next
-  assume ?lhs
-  then show ?rhs
-    unfolding continuous_on_eq_continuous_within
-    using continuous_within_sequentially[of _ s f]
-    by auto
-qed
+  shows "continuous_on S f \<longleftrightarrow>
+    (\<forall>x. \<forall>a \<in> S. (\<forall>n. x(n) \<in> S) \<and> (x \<longlongrightarrow> a) sequentially
+      \<longrightarrow> ((f \<circ> x) \<longlongrightarrow> f a) sequentially)"
+  by (meson continuous_on_eq_continuous_within continuous_within_sequentially)
 
 text \<open>Continuity in terms of open preimages.\<close>
 
 lemma continuous_at_open:
-  "continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t --> (\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x' \<in> s. (f x') \<in> t)))"
-  unfolding continuous_within_topological [of x UNIV f]
-  unfolding imp_conjL
-  by (intro all_cong imp_cong ex_cong conj_cong refl) auto
+  "continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t \<longrightarrow> (\<exists>S. open S \<and> x \<in> S \<and> (\<forall>x' \<in> S. (f x') \<in> t)))"
+  by (metis UNIV_I continuous_within_topological)
 
 lemma continuous_imp_tendsto:
-  assumes "continuous (at x0) f"
-    and "x \<longlonglongrightarrow> x0"
+  assumes "continuous (at x0) f" and "x \<longlonglongrightarrow> x0"
   shows "(f \<circ> x) \<longlonglongrightarrow> (f x0)"
 proof (rule topological_tendstoI)
   fix S
@@ -2411,20 +2175,20 @@
 
 subsection \<open>Homeomorphisms\<close>
 
-definition\<^marker>\<open>tag important\<close> "homeomorphism s t f g \<longleftrightarrow>
-  (\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and>
-  (\<forall>y\<in>t. (f(g y) = y)) \<and> (g ` t = s) \<and> continuous_on t g"
+definition\<^marker>\<open>tag important\<close> "homeomorphism S T f g \<longleftrightarrow>
+  (\<forall>x\<in>S. (g(f x) = x)) \<and> (f ` S = T) \<and> continuous_on S f \<and>
+  (\<forall>y\<in>T. (f(g y) = y)) \<and> (g ` T = S) \<and> continuous_on T g"
 
 lemma homeomorphismI [intro?]:
   assumes "continuous_on S f" "continuous_on T g"
-          "f ` S \<subseteq> T" "g ` T \<subseteq> S" "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y"
-    shows "homeomorphism S T f g"
+    "f ` S \<subseteq> T" "g ` T \<subseteq> S" "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y"
+  shows "homeomorphism S T f g"
   using assms by (force simp: homeomorphism_def)
 
 lemma homeomorphism_translation:
   fixes a :: "'a :: real_normed_vector"
   shows "homeomorphism ((+) a ` S) S ((+) (- a)) ((+) a)"
-unfolding homeomorphism_def by (auto simp: algebra_simps continuous_intros)
+  unfolding homeomorphism_def by (auto simp: algebra_simps continuous_intros)
 
 lemma homeomorphism_ident: "homeomorphism T T (\<lambda>a. a) (\<lambda>a. a)"
   by (rule homeomorphismI) auto
@@ -2469,43 +2233,25 @@
      "S homeomorphic {} \<longleftrightarrow> S = {}" "{} homeomorphic S \<longleftrightarrow> S = {}"
   by (auto simp: homeomorphic_def homeomorphism_def)
 
-lemma homeomorphic_refl: "s homeomorphic s"
-  unfolding homeomorphic_def homeomorphism_def
-  using continuous_on_id
-  apply (rule_tac x = "(\<lambda>x. x)" in exI)
-  apply (rule_tac x = "(\<lambda>x. x)" in exI)
-  apply blast
-  done
+lemma homeomorphic_refl: "S homeomorphic S"
+  using homeomorphic_def homeomorphism_ident by fastforce
 
-lemma homeomorphic_sym: "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
+lemma homeomorphic_sym: "S homeomorphic T \<longleftrightarrow> T homeomorphic S"
   unfolding homeomorphic_def homeomorphism_def
   by blast
 
 lemma homeomorphic_trans [trans]:
-  assumes "S homeomorphic T"
-      and "T homeomorphic U"
-    shows "S homeomorphic U"
-  using assms
-  unfolding homeomorphic_def
-by (metis homeomorphism_compose)
+  assumes "S homeomorphic T" and "T homeomorphic U"
+  shows "S homeomorphic U"
+  using assms unfolding homeomorphic_def
+  by (metis homeomorphism_compose)
 
 lemma homeomorphic_minimal:
-  "s homeomorphic t \<longleftrightarrow>
-    (\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and>
-           (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
-           continuous_on s f \<and> continuous_on t g)"
-   (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs
-    by (fastforce simp: homeomorphic_def homeomorphism_def)
-next
-  assume ?rhs
-  then show ?lhs
-    apply clarify
-    unfolding homeomorphic_def homeomorphism_def
-    by (metis equalityI image_subset_iff subsetI)
- qed
+  "S homeomorphic T \<longleftrightarrow>
+    (\<exists>f g. (\<forall>x\<in>S. f(x) \<in> T \<and> (g(f(x)) = x)) \<and>
+           (\<forall>y\<in>T. g(y) \<in> S \<and> (f(g(y)) = y)) \<and>
+           continuous_on S f \<and> continuous_on T g)"
+  by (smt (verit, ccfv_threshold) homeomorphic_def homeomorphismI homeomorphism_def image_eqI image_subset_iff)
 
 lemma homeomorphicI [intro?]:
    "\<lbrakk>f ` S = T; g ` T = S;
@@ -2517,8 +2263,7 @@
 lemma homeomorphism_of_subsets:
    "\<lbrakk>homeomorphism S T f g; S' \<subseteq> S; T'' \<subseteq> T; f ` S' = T'\<rbrakk>
     \<Longrightarrow> homeomorphism S' T' f g"
-apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
-by (metis subsetD imageI)
+  by (smt (verit, del_insts) continuous_on_subset homeomorphismI homeomorphism_def imageE subset_eq)
 
 lemma homeomorphism_apply1: "\<lbrakk>homeomorphism S T f g; x \<in> S\<rbrakk> \<Longrightarrow> g(f x) = x"
   by (simp add: homeomorphism_def)
@@ -2555,77 +2300,37 @@
 proof
   assume "S homeomorphic T"
   with assms show ?rhs
-    apply (auto simp: homeomorphic_def homeomorphism_def)
-     apply (metis finite_imageI)
-    by (metis card_image_le finite_imageI le_antisym)
+    by (metis (full_types) card_image_le finite_imageI homeomorphic_def homeomorphism_def le_antisym)
 next
   assume R: ?rhs
   with finite_same_card_bij obtain h where "bij_betw h S T"
     by auto
   with R show ?lhs
-    apply (auto simp: homeomorphic_def homeomorphism_def continuous_on_finite)
-    apply (rule_tac x=h in exI)
-    apply (rule_tac x="inv_into S h" in exI)
-    apply (auto simp:  bij_betw_inv_into_left bij_betw_inv_into_right bij_betw_imp_surj_on inv_into_into bij_betwE)
-    apply (metis bij_betw_def bij_betw_inv_into)
-    done
+    apply (simp only: homeomorphic_def homeomorphism_def continuous_on_finite)
+    by (smt (verit, ccfv_SIG) bij_betw_imp_surj_on bij_betw_inv_into bij_betw_inv_into_left bij_betw_inv_into_right)
 qed
 
 text \<open>Relatively weak hypotheses if a set is compact.\<close>
 
 lemma homeomorphism_compact:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
-  assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
-  shows "\<exists>g. homeomorphism s t f g"
+  assumes "compact S" "continuous_on S f"  "f ` S = T"  "inj_on f S"
+  shows "\<exists>g. homeomorphism S T f g"
 proof -
-  define g where "g x = (SOME y. y\<in>s \<and> f y = x)" for x
-  have g: "\<forall>x\<in>s. g (f x) = x"
-    using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto
-  {
-    fix y
-    assume "y \<in> t"
-    then obtain x where x:"f x = y" "x\<in>s"
-      using assms(3) by auto
-    then have "g (f x) = x" using g by auto
-    then have "f (g y) = y" unfolding x(1)[symmetric] by auto
-  }
-  then have g':"\<forall>x\<in>t. f (g x) = x" by auto
-  moreover
-  {
-    fix x
-    have "x\<in>s \<Longrightarrow> x \<in> g ` t"
-      using g[THEN bspec[where x=x]]
-      unfolding image_iff
-      using assms(3)
-      by (auto intro!: bexI[where x="f x"])
-    moreover
-    {
-      assume "x\<in>g ` t"
-      then obtain y where y:"y\<in>t" "g y = x" by auto
-      then obtain x' where x':"x'\<in>s" "f x' = y"
-        using assms(3) by auto
-      then have "x \<in> s"
-        unfolding g_def
-        using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"]
-        unfolding y(2)[symmetric] and g_def
-        by auto
-    }
-    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" ..
-  }
-  then have "g ` t = s" by auto
-  ultimately show ?thesis
-    unfolding homeomorphism_def homeomorphic_def
-    using assms continuous_on_inv by fastforce
+  obtain g where g: "\<forall>x\<in>S. g (f x) = x" "\<forall>x\<in>T. f (g x) = x" "g ` T = S"
+    using assms the_inv_into_f_f by fastforce 
+  with assms show ?thesis
+    unfolding homeomorphism_def homeomorphic_def by (metis continuous_on_inv)
 qed
 
 lemma homeomorphic_compact:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
-  shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s \<Longrightarrow> s homeomorphic t"
+  shows "compact S \<Longrightarrow> continuous_on S f \<Longrightarrow> (f ` S = T) \<Longrightarrow> inj_on f S \<Longrightarrow> S homeomorphic T"
   unfolding homeomorphic_def by (metis homeomorphism_compact)
 
 text\<open>Preservation of topological properties.\<close>
 
-lemma homeomorphic_compactness: "s homeomorphic t \<Longrightarrow> (compact s \<longleftrightarrow> compact t)"
+lemma homeomorphic_compactness: "S homeomorphic T \<Longrightarrow> (compact S \<longleftrightarrow> compact T)"
   unfolding homeomorphic_def homeomorphism_def
   by (metis compact_continuous_image)
 
@@ -2639,8 +2344,8 @@
 proof (rule islimptI)
   fix T
   assume "open T" "a \<in> T"
-  from open_right[OF this \<open>a < b\<close>]
-  obtain c where c: "a < c" "{a..<c} \<subseteq> T" by auto
+  then obtain c where c: "a < c" "{a..<c} \<subseteq> T"
+    by (meson assms open_right)
   with assms dense[of a "min c b"]
   show "\<exists>y\<in>{a<..<b}. y \<in> T \<and> y \<noteq> a"
     by (metis atLeastLessThan_iff greaterThanLessThan_iff min_less_iff_conj
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -133,16 +133,13 @@
 proof -
   from rational_boxes[OF assms, of c] obtain a b where ab: "c \<in> box a b" "box a b \<subseteq> ball c r"
     by auto
-  from ab have "0 < content (box a b)"
+  then have "0 < content (box a b)"
     by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def)
   have "emeasure lborel (box a b) \<le> emeasure lborel (ball c r)"
     using ab by (intro emeasure_mono) auto
-  also have "emeasure lborel (box a b) = ennreal (content (box a b))"
-    using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto
-  also have "emeasure lborel (ball c r) = ennreal (content (ball c r))"
-    using emeasure_lborel_ball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto
-  finally show ?thesis
-    using \<open>content (box a b) > 0\<close> by simp
+  then show ?thesis
+    using \<open>content (box a b) > 0\<close>
+    by (smt (verit, best) Sigma_Algebra.measure_def emeasure_lborel_ball_finite enn2real_mono infinity_ennreal_def)
 qed
 
 lemma content_cball_pos:
@@ -151,12 +148,10 @@
 proof -
   from rational_boxes[OF assms, of c] obtain a b where ab: "c \<in> box a b" "box a b \<subseteq> ball c r"
     by auto
-  from ab have "0 < content (box a b)"
+  then have "0 < content (box a b)"
     by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def)
-  have "emeasure lborel (box a b) \<le> emeasure lborel (ball c r)"
+  have "emeasure lborel (box a b) \<le> emeasure lborel (cball c r)"
     using ab by (intro emeasure_mono) auto
-  also have "\<dots> \<le> emeasure lborel (cball c r)"
-    by (intro emeasure_mono) auto
   also have "emeasure lborel (box a b) = ennreal (content (box a b))"
     using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto
   also have "emeasure lborel (cball c r) = ennreal (content (cball c r))"
@@ -209,7 +204,7 @@
   assumes "content (cbox a b) = 0"
     and "p tagged_division_of (cbox a b)"
   shows "(\<Sum>(x,K)\<in>p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)"
-proof (rule sum.neutral, rule)
+proof (intro sum.neutral strip)
   fix y
   assume y: "y \<in> p"
   obtain x K where xk: "y = (x, K)"
@@ -257,7 +252,7 @@
 qed
 
 lemma content_real_eq_0: "content {a..b::real} = 0 \<longleftrightarrow> a \<ge> b"
-  by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0)
+  by simp
 
 lemma property_empty_interval: "\<forall>a b. content (cbox a b) = 0 \<longrightarrow> P (cbox a b) \<Longrightarrow> P {}"
   using content_empty unfolding empty_as_interval by auto
@@ -332,7 +327,7 @@
 definition "integral i f = (SOME y. (f has_integral y) i \<or> \<not> f integrable_on i \<and> y=0)"
 
 lemma integrable_integral[intro]: "f integrable_on i \<Longrightarrow> (f has_integral (integral i f)) i"
-  unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex)
+  unfolding integrable_on_def integral_def by (metis (mono_tags, lifting))
 
 lemma not_integrable_integral: "\<not> f integrable_on i \<Longrightarrow> integral i f = 0"
   unfolding integrable_on_def integral_def by blast
@@ -351,7 +346,7 @@
 lemma has_integral_unique_cbox:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
   shows "(f has_integral k1) (cbox a b) \<Longrightarrow> (f has_integral k2) (cbox a b) \<Longrightarrow> k1 = k2"
-    by (auto simp: has_integral_cbox intro: tendsto_unique[OF division_filter_not_empty])    
+  by (meson division_filter_not_empty has_integral_cbox tendsto_unique)
 
 lemma has_integral_unique:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
@@ -400,9 +395,7 @@
 
 lemma eq_integralD: "integral k f = y \<Longrightarrow> (f has_integral y) k \<or> \<not> f integrable_on k \<and> y=0"
   unfolding integral_def integrable_on_def
-  apply (erule subst)
-  apply (rule someI_ex)
-  by blast
+  by (metis (mono_tags, lifting))
 
 lemma has_integral_const [intro]:
   fixes a b :: "'a::euclidean_space"
@@ -423,12 +416,12 @@
 lemma integral_const [simp]:
   fixes a b :: "'a::euclidean_space"
   shows "integral (cbox a b) (\<lambda>x. c) = content (cbox a b) *\<^sub>R c"
-  by (rule integral_unique) (rule has_integral_const)
+  by blast
 
 lemma integral_const_real [simp]:
   fixes a b :: real
   shows "integral {a..b} (\<lambda>x. c) = content {a..b} *\<^sub>R c"
-  by (metis box_real(2) integral_const)
+  by blast
 
 lemma has_integral_is_0_cbox:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
@@ -685,7 +678,7 @@
 
 lemma integral_0 [simp]:
   "integral S (\<lambda>x::'n::euclidean_space. 0::'m::real_normed_vector) = 0"
-  by (rule integral_unique has_integral_0)+
+  by auto
 
 lemma integral_add: "f integrable_on S \<Longrightarrow> g integrable_on S \<Longrightarrow>
     integral S (\<lambda>x. f x + g x) = integral S f + integral S g"
@@ -704,17 +697,10 @@
 lemma integral_mult:
   fixes K::real
   shows "f integrable_on X \<Longrightarrow> K * integral X f = integral X (\<lambda>x. K * f x)"
-  unfolding real_scaleR_def[symmetric] integral_cmul ..
+  by simp
 
 lemma integral_neg [simp]: "integral S (\<lambda>x. - f x) = - integral S f"
-proof (cases "f integrable_on S")
-  case True then show ?thesis
-    by (simp add: has_integral_neg integrable_integral integral_unique)
-next
-  case False then have "\<not> (\<lambda>x. - f x) integrable_on S"
-    using has_integral_neg [of "(\<lambda>x. - f x)" _ S ] by auto
-  with False show ?thesis by (simp add: not_integrable_integral)
-qed
+  by (metis eq_integralD equation_minus_iff has_integral_iff has_integral_neg_iff neg_equal_0_iff_equal)
 
 lemma integral_diff: "f integrable_on S \<Longrightarrow> g integrable_on S \<Longrightarrow>
     integral S (\<lambda>x. f x - g x) = integral S f - integral S g"
@@ -787,15 +773,8 @@
   assumes "finite T"
     and "\<And>a. a \<in> T \<Longrightarrow> ((f a) has_integral (i a)) S"
   shows "((\<lambda>x. sum (\<lambda>a. f a x) T) has_integral (sum i T)) S"
-  using assms(1) subset_refl[of T]
-proof (induct rule: finite_subset_induct)
-  case empty
-  then show ?case by auto
-next
-  case (insert x F)
-  with assms show ?case
-    by (simp add: has_integral_add)
-qed
+  using \<open>finite T\<close> subset_refl[of T]
+  by (induct rule: finite_subset_induct) (use assms in \<open>auto simp: has_integral_add\<close>)
 
 lemma integral_sum:
   "\<lbrakk>finite I;  \<And>a. a \<in> I \<Longrightarrow> f a integrable_on S\<rbrakk> \<Longrightarrow>
@@ -808,11 +787,11 @@
 
 lemma has_integral_eq:
   assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
-    and "(f has_integral k) s"
+    and f: "(f has_integral k) s"
   shows "(g has_integral k) s"
-  using has_integral_diff[OF assms(2), of "\<lambda>x. f x - g x" 0]
+  using has_integral_diff[OF f, of "\<lambda>x. f x - g x" 0]
   using has_integral_is_0[of s "\<lambda>x. f x - g x"]
-  using assms(1)
+  using assms
   by auto
 
 lemma integrable_eq: "\<lbrakk>f integrable_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g integrable_on s"
@@ -822,8 +801,7 @@
 lemma has_integral_cong:
   assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
   shows "(f has_integral i) s = (g has_integral i) s"
-  using has_integral_eq[of s f g] has_integral_eq[of s g f] assms
-  by auto
+  by (metis assms has_integral_eq)
 
 lemma integrable_cong:
   assumes "\<And>x. x \<in> A \<Longrightarrow> f x = g x"
@@ -855,25 +833,19 @@
   fixes f :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}"
   assumes "f integrable_on s"
   shows "(\<lambda>x. f x * of_real c) integrable_on s"
-using integrable_on_cmult_left [OF assms] by (simp add: mult.commute)
+  using integrable_on_cmult_left [OF assms] by (simp add: mult.commute)
 
 lemma integrable_on_cmult_right_iff [simp]:
   fixes f :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}"
   assumes "c \<noteq> 0"
   shows "(\<lambda>x. f x * of_real c) integrable_on s \<longleftrightarrow> f integrable_on s"
-using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute)
-
-lemma integrable_on_cdivide:
-  fixes f :: "_ \<Rightarrow> 'b :: real_normed_field"
-  assumes "f integrable_on s"
-  shows "(\<lambda>x. f x / of_real c) integrable_on s"
-by (simp add: integrable_on_cmult_right divide_inverse assms flip: of_real_inverse)
+  using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute)
 
 lemma integrable_on_cdivide_iff [simp]:
   fixes f :: "_ \<Rightarrow> 'b :: real_normed_field"
   assumes "c \<noteq> 0"
   shows "(\<lambda>x. f x / of_real c) integrable_on s \<longleftrightarrow> f integrable_on s"
-by (simp add: divide_inverse assms flip: of_real_inverse)
+  by (simp add: divide_inverse assms flip: of_real_inverse)
 
 lemma has_integral_null [intro]: "content(cbox a b) = 0 \<Longrightarrow> (f has_integral 0) (cbox a b)"
   unfolding has_integral_cbox
@@ -902,7 +874,7 @@
   unfolding integrable_on_def by auto
 
 lemma integral_empty[simp]: "integral {} f = 0"
-  by (rule integral_unique) (rule has_integral_empty)
+  by blast
 
 lemma has_integral_refl[intro]:
   fixes a :: "'a::euclidean_space"
@@ -919,7 +891,7 @@
   unfolding integrable_on_def by auto
 
 lemma integral_refl [simp]: "integral (cbox a a) f = 0"
-  by (rule integral_unique) auto
+  by auto
 
 lemma integral_singleton [simp]: "integral {a} f = 0"
   by auto
@@ -1000,7 +972,6 @@
   by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def)
 
 
-
 subsection \<open>Cauchy-type criterion for integrability\<close>
 
 proposition integrable_Cauchy:
@@ -1634,10 +1605,10 @@
 corollary integrable_bound:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
   assumes "0 \<le> B"
-      and "f integrable_on (cbox a b)"
-      and "\<And>x. x\<in>cbox a b \<Longrightarrow> norm (f x) \<le> B"
-    shows "norm (integral (cbox a b) f) \<le> B * content (cbox a b)"
-by (metis integrable_integral has_integral_bound assms)
+    and "f integrable_on (cbox a b)"
+    and "\<And>x. x\<in>cbox a b \<Longrightarrow> norm (f x) \<le> B"
+  shows "norm (integral (cbox a b) f) \<le> B * content (cbox a b)"
+  by (metis integrable_integral has_integral_bound assms)
 
 
 subsection \<open>Similar theorems about relationship among components\<close>
@@ -1752,21 +1723,14 @@
     and "(f has_integral i) S"
     and "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> (f x)\<bullet>k"
   shows "0 \<le> i\<bullet>k"
-  using has_integral_component_le[OF assms(1) has_integral_0 assms(2)]
-  using assms(3-)
-  by auto
+  by (metis (no_types, lifting) assms euclidean_all_zero_iff has_integral_0 has_integral_component_le)
 
 lemma integral_component_nonneg:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "k \<in> Basis"
     and  "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> (f x)\<bullet>k"
   shows "0 \<le> (integral S f)\<bullet>k"
-proof (cases "f integrable_on S")
-  case True show ?thesis
-    using True assms has_integral_component_nonneg by blast
-next
-  case False then show ?thesis by (simp add: not_integrable_integral)
-qed
+  by (smt (verit, ccfv_threshold) assms eq_integralD euclidean_all_zero_iff has_integral_component_nonneg)
 
 lemma has_integral_component_neg:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -1774,8 +1738,7 @@
     and "(f has_integral i) S"
     and "\<And>x. x \<in> S \<Longrightarrow> (f x)\<bullet>k \<le> 0"
   shows "i\<bullet>k \<le> 0"
-  using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-)
-  by auto
+  by (metis (no_types, lifting) assms has_integral_0 has_integral_component_le inner_zero_left)
 
 lemma has_integral_component_lbound:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -1808,8 +1771,7 @@
     and "\<forall>x\<in>{a..b}. B \<le> f(x)\<bullet>k"
     and "k \<in> Basis"
   shows "B * content {a..b} \<le> (integral {a..b} f)\<bullet>k"
-  using assms
-  by (metis box_real(2) integral_component_lbound)
+  using assms by (metis box_real(2) integral_component_lbound)
 
 lemma integral_component_ubound:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -1825,8 +1787,7 @@
     and "\<forall>x\<in>{a..b}. f x\<bullet>k \<le> B"
     and "k \<in> Basis"
   shows "(integral {a..b} f)\<bullet>k \<le> B * content {a..b}"
-  using assms
-  by (metis box_real(2) integral_component_ubound)
+  using assms by (metis box_real(2) integral_component_ubound)
 
 subsection \<open>Uniform limit of integrable functions is integrable\<close>
 
@@ -1863,7 +1824,7 @@
     then obtain M where "M \<noteq> 0" and M: "1 / (real M) < e/4 / content (cbox a b)"
       by (metis inverse_eq_divide real_arch_inverse)
     show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (h m) (h n) < e"
-    proof (rule exI [where x=M], clarify)
+    proof (intro exI strip)
       fix m n
       assume m: "M \<le> m" and n: "M \<le> n"
       have "e/4>0" using \<open>e>0\<close> by auto
@@ -2611,8 +2572,7 @@
   fixes f :: "real \<Rightarrow> 'a::banach"
   assumes "continuous_on (closed_segment a b) f"
   shows "f integrable_on (closed_segment a b)"
-  using assms
-  by (auto intro!: integrable_continuous_interval simp: closed_segment_eq_real_ivl)
+  by (metis assms closed_segment_eq_real_ivl integrable_continuous_interval)
 
 
 subsection \<open>Specialization of additivity to one dimension\<close>
@@ -2773,7 +2733,7 @@
 lemma ident_integrable_on:
   fixes a::real
   shows "(\<lambda>x. x) integrable_on {a..b}"
-by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral)
+  using continuous_on_id integrable_continuous_real by blast
 
 lemma integral_sin [simp]:
   fixes a::real
@@ -2862,7 +2822,7 @@
   assumes f'[derivative_intros]:
     "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
   assumes line_in: "\<And>t. t \<in> {0..1} \<Longrightarrow> x + t *\<^sub>R y \<in> S"
-  shows "f (x + y) - f x = integral {0..1} (\<lambda>t. f' (x + t *\<^sub>R y) y)" (is ?th1)
+  shows "f (x + y) - f x = integral {0..1} (\<lambda>t. f' (x + t *\<^sub>R y) y)"
 proof -
   from assms have subset: "(\<lambda>xa. x + xa *\<^sub>R y) ` {0..1} \<subseteq> S" by auto
   note [derivative_intros] =
@@ -2879,7 +2839,7 @@
         linear_cmul[OF has_derivative_linear[OF f'], symmetric]
       intro!: derivative_eq_intros)
   from fundamental_theorem_of_calculus[rule_format, OF _ this]
-  show ?th1
+  show ?thesis
     by (auto intro!: integral_unique[symmetric])
 qed
 
@@ -2969,26 +2929,20 @@
     using atLeastatMost_empty'[simp del]
     by (simp add: i_def g_def Dg_def)
   also
-  have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)"
-    and "{..<p} \<inter> {i. p = Suc i} = {p - 1}"
-    for p'
-    using \<open>p > 0\<close>
-    by (auto simp: power_mult_distrib[symmetric])
+  have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" "{..<p} \<inter> {i. p = Suc i} = {p - 1}" for p'
+    using \<open>p > 0\<close> by (auto simp: power_mult_distrib)
   then have "?sum b = f b"
     using Suc_pred'[OF \<open>p > 0\<close>]
     by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib
         if_distribR sum.If_cases f0)
   also
-  have "{..<p} = (\<lambda>x. p - x - 1) ` {..<p}"
-  proof safe
-    fix x
-    assume "x < p"
-    thus "x \<in> (\<lambda>x. p - x - 1) ` {..<p}"
-      by (auto intro!: image_eqI[where x = "p - x - 1"])
-  qed simp
-  from _ this
   have "?sum a = (\<Sum>i<p. ((b-a) ^ i / fact i) *\<^sub>R Df i a)"
-    by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one)
+  proof (rule sum.reindex_cong)
+    have "\<And>i. i < p \<Longrightarrow> \<exists>j<p. i = p - Suc j"
+      by (metis Suc_diff_Suc \<open>p>0\<close> diff_Suc_less diff_diff_cancel less_or_eq_imp_le)
+    then show "{..<p} = (\<lambda>x. p - x - 1) ` {..<p}"
+      by force
+  qed (auto simp add: inj_on_def Dg_def one)
   finally show c: ?case .
   case 2 show ?case using c integral_unique
     by (metis (lifting) add.commute diff_eq_eq integral_unique)
@@ -3112,15 +3066,7 @@
   interpret operative conj True "\<lambda>i. f integrable_on i"
     using order_refl by (rule operative_integrableI)
   show ?thesis
-  proof (cases "cbox c d = {}")
-    case True
-    then show ?thesis
-      using division [symmetric] f by (auto simp: comm_monoid_set_F_and)
-  next
-    case False
-    then show ?thesis
-      by (metis cd comm_monoid_set_F_and division division_of_finite f partial_division_extend_1)
-  qed
+    by (metis cd division division_of_finite empty f partial_division_extend_1 remove)
 qed
 
 lemma integrable_subinterval_real:
@@ -3151,8 +3097,7 @@
     by (auto simp: split: if_split_asm)
   then have "f integrable_on cbox a b"
     using ac cb by (auto split: if_split_asm)
-  with *
-  show ?thesis
+  with * show ?thesis
     using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm)
 qed
 
@@ -3168,10 +3113,8 @@
   moreover
   have "(f has_integral integral {c..b} f) {c..b}"
     using ab \<open>a \<le> c\<close> integrable_subinterval_real by fastforce
-  ultimately have "(f has_integral integral {a..c} f + integral {c..b} f) {a..b}"
-    using \<open>a \<le> c\<close> \<open>c \<le> b\<close> has_integral_combine by blast
-  then show ?thesis
-    by (simp add: has_integral_integrable_integral)
+  ultimately show ?thesis
+    by (smt (verit, best) assms has_integral_combine integral_unique)
 qed
 
 lemma integrable_combine:
@@ -3181,9 +3124,7 @@
     and "f integrable_on {a..c}"
     and "f integrable_on {c..b}"
   shows "f integrable_on {a..b}"
-  using assms
-  unfolding integrable_on_def
-  by (auto intro!: has_integral_combine)
+  using assms has_integral_combine by blast
 
 lemma integral_minus_sets:
   fixes f::"real \<Rightarrow> 'a::banach"
@@ -3304,7 +3245,7 @@
   assumes "continuous_on {a..b} f"
     and "x \<in> {a..b}"
   shows "((\<lambda>u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})"
-using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real]
+  using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real]
   by (fastforce simp: continuous_on_eq_continuous_within)
 
 lemma integral_has_real_derivative:
@@ -3396,29 +3337,29 @@
         show "d fine (\<lambda>(x, k). (g x, g ` k)) ` p"
           using finep unfolding fine_def d'_def by auto
       next
-        fix x k
-        assume xk: "(x, k) \<in> p"
-        show "g x \<in> g ` k"
+        fix x K
+        assume xk: "(x, K) \<in> p"
+        show "g x \<in> g ` K"
           using p(2)[OF xk] by auto
-        show "\<exists>u v. g ` k = cbox u v"
+        show "\<exists>u v. g ` K = cbox u v"
           using p(4)[OF xk] using assms(5-6) by auto
         fix x' K' u
-        assume xk': "(x', K') \<in> p" and u: "u \<in> interior (g ` k)" "u \<in> interior (g ` K')"
-        have "interior k \<inter> interior K' \<noteq> {}"
+        assume xk': "(x', K') \<in> p" and u: "u \<in> interior (g ` K)" "u \<in> interior (g ` K')"
+        have "interior K \<inter> interior K' \<noteq> {}"
         proof 
-          assume "interior k \<inter> interior K' = {}"
-          moreover have "u \<in> g ` (interior k \<inter> interior K')"
+          assume "interior K \<inter> interior K' = {}"
+          moreover have "u \<in> g ` (interior K \<inter> interior K')"
             using interior_image_subset[OF \<open>inj g\<close> contg] u
             unfolding image_Int[OF inj(1)] by blast
           ultimately show False by blast
         qed
-        then have same: "(x, k) = (x', K')"
+        then have same: "(x, K) = (x', K')"
           using ptag xk' xk by blast
         then show "g x = g x'"
           by auto
-        show "g u \<in> g ` K'"if "u \<in> k" for u
+        show "g u \<in> g ` K'"if "u \<in> K" for u
           using that same by auto
-        show "g u \<in> g ` k"if "u \<in> K'" for u
+        show "g u \<in> g ` K"if "u \<in> K'" for u
           using that same by auto
       next
         fix x
@@ -3528,8 +3469,7 @@
     moreover from True have *: "\<And>i. (m *\<^sub>R b + c) \<bullet> i - (m *\<^sub>R a + c) \<bullet> i = m *\<^sub>R (b-a) \<bullet> i"
       by (simp add: inner_simps field_simps)
     ultimately show ?thesis
-      by (simp add: image_affinity_cbox True content_cbox'
-        prod.distrib inner_diff_left)
+      by (simp add: image_affinity_cbox True content_cbox' prod.distrib inner_diff_left)
   next
     case False
     with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \<noteq> {}"
@@ -3578,7 +3518,8 @@
   assumes "c \<noteq> 0"
   shows   "((\<lambda>x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \<longleftrightarrow> (f has_integral I) A"
   using assms has_integral_cmul[of f I A c]
-        has_integral_cmul[of "\<lambda>x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps)
+        has_integral_cmul[of "\<lambda>x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] 
+  by (auto simp: field_simps)
 
 lemma has_integral_cmul_iff':
   assumes "c \<noteq> 0"
@@ -3607,7 +3548,8 @@
     by (subst image_smult_cbox) simp_all
   also have "(\<lambda>x. - ((1 / m) *\<^sub>R c) + x) ` \<dots> = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)"
     by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps)
-  finally show ?thesis using \<open>m > 0\<close> by (simp add: field_simps)
+  finally show ?thesis using \<open>m > 0\<close> 
+    by (simp add: field_simps)
 qed
 
 lemma has_integral_affinity_iff:
@@ -3619,7 +3561,7 @@
 proof
   assume ?lhs
   from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \<open>m > 0\<close>
-    show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps)
+  show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps)
 next
   assume ?rhs
   from has_integral_affinity'[OF this, of m c] and \<open>m > 0\<close>
@@ -3708,12 +3650,7 @@
   shows "uminus ` cbox a b = cbox (-b) (-a)"
 proof -
   have "x \<in> uminus ` cbox a b" if "x \<in> cbox (- b) (- a)" for x
-  proof -
-    have "-x \<in> cbox a b"
-      using that by (auto simp: mem_box)
-    then show ?thesis
-      by force
-  qed
+    by (smt (verit) add.inverse_inverse image_iff inner_minus_left mem_box(2) that)
   then show ?thesis
     by (auto simp: mem_box)
 qed
@@ -3727,9 +3664,7 @@
 lemma has_integral_reflect_lemma_real[intro]:
   assumes "(f has_integral i) {a..b::real}"
   shows "((\<lambda>x. f(-x)) has_integral i) {-b .. -a}"
-  using assms
-  unfolding box_real[symmetric]
-  by (rule has_integral_reflect_lemma)
+  by (metis has_integral_reflect_lemma interval_cbox assms)
 
 lemma has_integral_reflect[simp]:
   "((\<lambda>x. f (-x)) has_integral i) (cbox (-b) (-a)) \<longleftrightarrow> (f has_integral i) (cbox a b)"
@@ -3832,9 +3767,8 @@
         proof (rule add_mono)
           have "norm ((c - a) *\<^sub>R f' a) \<le> norm (l *\<^sub>R f' a)"
             by (auto intro: mult_right_mono [OF lel])
-          also have "... \<le> e * (b-a) / 8"
-            by (rule l)
-          finally show "norm ((c - a) *\<^sub>R f' a) \<le> e * (b-a) / 8" .
+          with l show "norm ((c - a) *\<^sub>R f' a) \<le> e * (b-a) / 8"
+            by linarith 
         next
           have "norm (f c - f a) < e * (b-a) / 8"
           proof (cases "a = c")
@@ -3959,9 +3893,8 @@
         by (auto intro: sum_norm_le)
       also have "... \<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k)/2)"
         using non by (fastforce intro: sum_mono)
-      finally have I: "norm (\<Sum>(x, k)\<in>p - ?A.
-                  content k *\<^sub>R f' x - (f (Sup k) - f (Inf k)))
-             \<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))/2"
+      finally have I: "norm (\<Sum>(x, k)\<in>p - ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k)))
+                     \<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))/2"
         by (simp add: sum_divide_distrib)
       have II: "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) -
              (\<Sum>n\<in>p \<inter> ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))
@@ -3971,7 +3904,7 @@
         proof -
           obtain u v where uv: "k = cbox u v"
             by (meson Int_iff xkp p(4))
-          with p(2) that uv have "cbox u v \<noteq> {}"
+          with p that have "cbox u v \<noteq> {}"
             by blast
           then show "0 \<le> e * ((Sup k) - (Inf k))"
             unfolding uv using e by (auto simp add: field_simps)
@@ -3987,10 +3920,8 @@
           proof -
             have xk: "(x,K) \<in> p" and k0: "content K = 0"
               using that by auto
-            then obtain u v where uv: "K = cbox u v"
-              using p(4) by blast
-            then have "u = v"
-              using xk k0 p(2) by force
+            then obtain u v where uv: "K = cbox u v" "u = v"
+              using xk k0 p by fastforce
             then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0"
               using xk unfolding uv by auto
           qed
@@ -4021,7 +3952,7 @@
             also have "... \<le> e * (b - a) / 4 + e * (b - a) / 4"
             proof (rule norm_triangle_le [OF add_mono])
               have pa: "\<exists>v. k = cbox a v \<and> a \<le> v" if "(a, k) \<in> p" for k
-                using p(2) p(3) p(4) that by fastforce
+                using p that by fastforce
               show "norm (\<Sum>(x,K) \<in> ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (b - a) / 4"
               proof (intro norm_le; clarsimp)
                 fix K K'
@@ -4059,7 +3990,7 @@
               qed (use ab e in auto)
             next
               have pb: "\<exists>v. k = cbox v b \<and> b \<ge> v" if "(b, k) \<in> p" for k
-                using p(2) p(3) p(4) that by fastforce
+                using p that by fastforce
               show "norm (\<Sum>(x,K) \<in> ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (b - a) / 4"
               proof (intro norm_le; clarsimp)
                 fix K K'
@@ -4071,7 +4002,7 @@
                   unfolding v v' by (auto simp: mem_box)
                 then have "interior (box (max v v') b) \<subseteq> interior K \<inter> interior K'"
                   using interior_Int interior_mono by blast
-                moreover have " ((b + ?v)/2) \<in> box ?v b"
+                moreover have "((b + ?v)/2) \<in> box ?v b"
                   using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box)
                 ultimately have "((b + ?v)/2) \<in> interior K \<inter> interior K'"
                   unfolding interior_open[OF open_box] by auto
@@ -4181,18 +4112,11 @@
     hence e3: "0 < e/3 / norm (f c)" using \<open>e>0\<close> by simp
     moreover have "norm (f c) * norm (c - t) < e/3" 
       if "t < c" and "c - e/3 / norm (f c) < t" for t
-    proof -
-      have "norm (c - t) < e/3 / norm (f c)"
-        using that by auto
-      then show "norm (f c) * norm (c - t) < e/3"
-        by (metis e3 mult.commute norm_not_less_zero pos_less_divide_eq zero_less_divide_iff)
-    qed
+      unfolding real_norm_def
+      by (smt (verit) False divide_right_mono nonzero_mult_div_cancel_left norm_eq_zero norm_ge_zero that)
     ultimately show ?thesis
       using that by auto
-  next
-    case True then show ?thesis
-      using \<open>e > 0\<close> that by auto
-  qed
+  qed (use \<open>e > 0\<close> in auto)
 
   let ?SUM = "\<lambda>p. (\<Sum>(x,K) \<in> p. content K *\<^sub>R f x)"
   have e3: "e/3 > 0"
@@ -4288,7 +4212,7 @@
       ultimately have cwt: "c - w < t"
         by (auto simp add: field_simps)
       have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) -
-             integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c"
+                integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c"
         by auto
       have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3"
         unfolding eq
@@ -4373,15 +4297,8 @@
         show "0 < min d1 d2"
           using \<open>0 < d1\<close> \<open>0 < d2\<close> by simp
         show "dist (integral {a..y} f) (integral {a..x} f) < e"
-             if "y \<in> {a..b}" "dist y x < min d1 d2" for y
-        proof (cases "y < x")
-          case True
-          with that d1 show ?thesis by (auto simp: dist_commute dist_norm)
-        next
-          case False
-          with that d2 show ?thesis
-            by (auto simp: dist_commute dist_norm)
-        qed
+          if "dist y x < min d1 d2" for y
+          by (smt (verit) d1 d2 dist_norm dist_real_def norm_minus_commute that)
       qed
     qed
   qed
@@ -4562,18 +4479,7 @@
   assumes "connected S" "open S" "finite K" "continuous_on S f"
       and "\<forall>x\<in>S-K. (f has_derivative (\<lambda>h. 0)) (at x within S)"
   shows   "f constant_on S"
-proof (cases "S = {}")
-  case True
-  then show ?thesis
-    by (simp add: constant_on_def)
-next
-  case False
-  then obtain c where "c \<in> S"
-    by (metis equals0I)
-  then show ?thesis
-    unfolding constant_on_def
-    by (metis has_derivative_zero_unique_strong_connected assms )
-qed
+  by (smt (verit, best) assms constant_on_def has_derivative_zero_unique_strong_connected)
 
 lemma DERIV_zero_connected_constant_on:
   fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
@@ -4594,13 +4500,8 @@
   fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
   assumes "\<And>z. z \<in> S \<Longrightarrow> (f has_field_derivative 0) (at z)" and S: "connected S" "open S"
   shows   "f constant_on S"
-proof -
-  have *: "continuous_on S f"
-    using assms(1) by (intro DERIV_continuous_on[of _ _ "\<lambda>_. 0"])
-                      (use assms in \<open>auto simp: at_within_open[of _ S]\<close>)
-  show ?thesis
-    using DERIV_zero_connected_constant_on[OF S finite.emptyI *] assms(1) by blast
-qed
+  using DERIV_zero_connected_constant_on [where K="Basis"]
+  by (metis DERIV_isCont Diff_iff assms continuous_at_imp_continuous_on eucl.finite_Basis)
 
 
 subsection \<open>Integrating characteristic function of an interval\<close>
@@ -4653,19 +4554,11 @@
       using integrable_spike_interior[where f=f]
       by (meson g_def has_integral_integrable intf)
     moreover have "integral (cbox c d) g = i"
-    proof (rule has_integral_unique[OF has_integral_spike_interior intf])
-      show "\<And>x. x \<in> box c d \<Longrightarrow> f x = g x"
-        by (auto simp: g_def)
-      show "(g has_integral integral (cbox c d) g) (cbox c d)"
-        by (rule integrable_integral[OF intg])
-    qed
+      by (meson g_def has_integral_iff has_integral_spike_interior intf)
     ultimately have "F (\<lambda>A. if g integrable_on A then Some (integral A g) else None) p = Some i"
       by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral)
-    then
-    have "(g has_integral i) (cbox a b)"
-      by (metis integrable_on_def integral_unique operat option.inject option.simps(3))
     with False show ?thesis
-      by blast
+      by (metis integrable_integral not_None_eq operat option.inject)
   qed
 qed
 
@@ -4799,17 +4692,13 @@
     and "g integrable_on S"
     and "\<And>x. x \<in> S \<Longrightarrow> f x \<le> g x"
   shows "integral S f \<le> integral S g"
-  by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)])
+  by (meson assms has_integral_le integrable_integral)
 
 lemma has_integral_nonneg:
   fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "(f has_integral i) S"
-    and "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x"
+  assumes "(f has_integral i) S" and "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x"
   shows "0 \<le> i"
-  using has_integral_component_nonneg[of 1 f i S]
-  unfolding o_def
-  using assms
-  by auto
+  using assms has_integral_0 has_integral_le by blast
 
 lemma integral_nonneg:
   fixes f :: "'n::euclidean_space \<Rightarrow> real"
@@ -4862,16 +4751,11 @@
 
 lemma has_integral_on_superset:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes f: "(f has_integral i) S"
+  assumes "(f has_integral i) S"
       and "\<And>x. x \<notin> S \<Longrightarrow> f x = 0"
       and "S \<subseteq> T"
     shows "(f has_integral i) T"
-proof -
-  have "(\<lambda>x. if x \<in> S then f x else 0) = (\<lambda>x. if x \<in> T then f x else 0)"
-    using assms by fastforce
-  with f show ?thesis
-    by (simp only: has_integral_restrict_UNIV [symmetric, of f])
-qed
+  by (smt (verit, ccfv_SIG) assms has_integral_cong has_integral_restrict)
 
 lemma integrable_on_superset:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
@@ -4906,14 +4790,14 @@
 lemma has_integral_subset_component_le:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
   assumes k: "k \<in> Basis"
-    and as: "S \<subseteq> T" "(f has_integral i) S" "(f has_integral j) T" "\<And>x. x\<in>T \<Longrightarrow> 0 \<le> f(x)\<bullet>k"
+    and "S \<subseteq> T" "(f has_integral i) S" "(f has_integral j) T" "\<And>x. x\<in>T \<Longrightarrow> 0 \<le> f(x)\<bullet>k"
   shows "i\<bullet>k \<le> j\<bullet>k"
 proof -
   have \<section>: "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) UNIV"
           "((\<lambda>x. if x \<in> T then f x else 0) has_integral j) UNIV"
     by (simp_all add: assms)
   show ?thesis
-    using as by (force intro!: has_integral_component_le[OF k \<section>])
+    using assms by (force intro!: has_integral_component_le[OF k \<section>])
 qed
 
 subsection\<open>Integrals on set differences\<close>
@@ -4961,13 +4845,7 @@
   assume R: ?r
   show ?l
     unfolding negligible_def
-  proof safe
-    fix a b
-    have "negligible (s \<inter> cbox a b)"
-      by (simp add: R)
-    then show "(indicator s has_integral 0) (cbox a b)"
-      by (meson Diff_iff Int_iff has_integral_negligible indicator_simps(2))
-  qed
+    by (metis Diff_iff Int_iff R has_integral_negligible indicator_simps(2))
 qed (simp add: negligible_Int)
 
 lemma negligible_translation:
@@ -4994,8 +4872,8 @@
 
 lemma negligible_translation_rev:
   assumes "negligible ((+) c ` S)"
-    shows "negligible S"
-by (metis negligible_translation [OF assms, of "-c"] translation_galois)
+  shows "negligible S"
+  by (metis negligible_translation [OF assms, of "-c"] translation_galois)
 
 lemma negligible_atLeastAtMostI: "b \<le> a \<Longrightarrow> negligible {a..(b::real)}"
   using negligible_insert by fastforce
@@ -5077,12 +4955,7 @@
 lemma has_integral_Icc_iff_Ioo:
   fixes f :: "real \<Rightarrow> 'a :: banach"
   shows "(f has_integral I) {a..b} \<longleftrightarrow> (f has_integral I) {a<..<b}"
-proof (rule has_integral_spike_set_eq)
-  show "negligible {x \<in> {a..b} - {a<..<b}. f x \<noteq> 0}"
-    by (rule negligible_subset [of "{a,b}"]) auto
-  show "negligible {x \<in> {a<..<b} - {a..b}. f x \<noteq> 0}"
-    by (rule negligible_subset [of "{}"]) auto
-qed
+  by (metis box_real(1) cbox_interval has_integral_open_interval)
 
 lemma integrable_on_Icc_iff_Ioo:
   fixes f :: "real \<Rightarrow> 'a :: banach"
@@ -5099,8 +4972,7 @@
     and "(f has_integral j) t"
     and "\<forall>x\<in>t. 0 \<le> f x"
   shows "i \<le> j"
-  using has_integral_subset_component_le[OF _ assms(1), of 1 f i j]
-  using assms
+  using assms has_integral_subset_component_le[OF _ assms(1), of 1 f i j]
   by auto
 
 lemma integral_subset_component_le:
@@ -5166,7 +5038,6 @@
       show "cbox a b \<subseteq> cbox ?a ?b"
         by (force simp: mem_box)
     qed
-  
     fix e :: real
     assume "e > 0"
     with lhs show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
@@ -5659,12 +5530,7 @@
   shows "(f has_integral (sum (\<lambda>i. integral i f) \<D>)) K"
 proof -
   have "f integrable_on L" if "L \<in> \<D>" for L
-  proof -
-    have "L \<subseteq> S"
-      using \<open>K \<subseteq> S\<close> \<D> that by blast
-    then show "f integrable_on L"
-      using that by (metis (no_types) f \<D> division_ofD(4) integrable_on_subcbox)
-  qed
+    by (smt (verit, best) assms cbox_division_memE f integrable_on_subcbox subset_trans that)
   then show ?thesis
     by (meson \<D> has_integral_combine_division has_integral_integrable_integral)
 qed
@@ -5691,12 +5557,7 @@
   shows "f integrable_on i"
 proof -
   have "f integrable_on i" if "i \<in> \<D>" for i
-proof -
-  have "i \<subseteq> S"
-    using assms that by auto
-  then show "f integrable_on i"
-    using that by (metis (no_types) \<D> f division_ofD(4) integrable_on_subcbox)
-qed
+    by (smt (verit, best) assms cbox_division_memE f integrable_on_subcbox order_trans that)
   then show ?thesis
     using \<D> integrable_combine_division by blast
 qed
@@ -5711,12 +5572,7 @@
   shows "(f has_integral (\<Sum>(x,k)\<in>p. i k)) S"
 proof -
   have *: "(f has_integral (\<Sum>k\<in>snd`p. integral k f)) S"
-  proof -
-    have "snd ` p division_of S"
-      by (simp add: assms(1) division_of_tagged_division)
-    with assms show ?thesis
-      by (metis (mono_tags, lifting) has_integral_combine_division has_integral_integrable_integral imageE prod.collapse)
-  qed
+    by (smt (verit, del_insts) assms division_of_tagged_division has_integral_combine_division has_integral_iff imageE prod.collapse)
   also have "(\<Sum>k\<in>snd`p. integral k f) = (\<Sum>(x, k)\<in>p. integral k f)"
     by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null)
        (simp add: content_eq_0_interior)
@@ -5850,11 +5706,9 @@
       have "L \<in> snd ` p" 
         using \<open>(x,L) \<in> p\<close> image_iff by fastforce 
       then have "L \<in> q" "K \<in> q" "L \<noteq> K"
-        using that(1,3) q(1) unfolding r_def by auto
-      with q'(5) have "interior L = {}"
-        using interior_mono[OF \<open>L \<subseteq> K\<close>] by blast
-      then show "content L *\<^sub>R f x = 0"
-        unfolding uv content_eq_0_interior[symmetric] by auto
+        using that q(1) unfolding r_def by auto
+      with q'(5) show "content L *\<^sub>R f x = 0"
+        by (metis \<open>L \<subseteq> K\<close> content_eq_0_interior inf.orderE interior_Int scaleR_eq_0_iff uv)
     qed
     show "finite (\<Union>(qq ` r))"
       by (meson finite_UN qq \<open>finite r\<close> tagged_division_of_finite)
@@ -5868,11 +5722,9 @@
       using \<open>(x, M) \<in> qq L\<close> \<open>L \<in> r\<close> kl(2) by blast
     have empty: "interior (K \<inter> L) = {}"
       by (metis DiffD1 interior_Int q'(5) r_def KL r)
-    have "interior M = {}"
-      by (metis (no_types, lifting) Int_assoc empty inf.absorb_iff2 interior_Int kl(1) subset_empty x r)
-    then show "content M *\<^sub>R f x = 0"
-      unfolding uv content_eq_0_interior[symmetric]
-      by auto
+    with that kl show "content M *\<^sub>R f x = 0"
+      by (metis content_eq_0_interior dual_order.refl inf.orderE inf_mono interior_mono 
+                scaleR_eq_0_iff subset_empty uv x)
   qed 
   ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e"
     apply (subst (asm) sum.Union_comp)
@@ -5886,12 +5738,9 @@
   have norm_le: "norm (cp - ip) \<le> e + k"
                   if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i"
                   for ir ip i cr cp::'a
-  proof -
-    from that show ?thesis
-      using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"]
-      unfolding that(3)[symmetric] norm_minus_cancel
-      by (auto simp add: algebra_simps)
-  qed
+    using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] that
+    unfolding that(3)[symmetric] norm_minus_cancel
+    by (auto simp add: algebra_simps)
 
   have "?lhs =  norm (?SUM p - (\<Sum>(x, k)\<in>p. integral k f))"
     unfolding split_def sum_subtractf ..
@@ -5911,11 +5760,9 @@
     proof -
       obtain u v where uv: "l = cbox u v"
         using inp p'(4) by blast
-      have "content (cbox u v) = 0"
-        unfolding content_eq_0_interior using that p(1) uv
-        by (auto dest: tagged_partial_division_ofD)
       then show ?thesis
-        using uv by blast
+        using uv that p
+        by (metis content_eq_0_interior dual_order.refl inf.orderE integral_null ne tagged_partial_division_ofD(5))
     qed
     then have "(\<Sum>(x, K)\<in>p. integral K f) = (\<Sum>K\<in>snd ` p. integral K f)"
       apply (subst sum.reindex_nontrivial [OF \<open>finite p\<close>])
@@ -6526,14 +6373,10 @@
 lemma has_integral_norm_bound_integral_component:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
   fixes g :: "'n \<Rightarrow> 'b::euclidean_space"
-  assumes f: "(f has_integral i) S"
-    and g: "(g has_integral j) S"
+  assumes "(f has_integral i) S" and "(g has_integral j) S"
     and "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> (g x)\<bullet>k"
   shows "norm i \<le> j\<bullet>k"
-  using integral_norm_bound_integral_component[of f S g k] 
-  unfolding integral_unique[OF f] integral_unique[OF g]
-  using assms
-  by auto
+  by (metis assms has_integral_integrable integral_norm_bound_integral_component integral_unique)
 
 
 lemma uniformly_convergent_improper_integral:
@@ -6631,7 +6474,7 @@
 
   show ?thesis
     unfolding continuous_on_def
-  proof (safe intro!: tendstoI)
+  proof (intro strip tendstoI)
     fix e'::real and x
     assume "e' > 0"
     define e where "e = e' / (content (cbox a b) + 1)"
@@ -6712,7 +6555,7 @@
       case (elim x)
       from elim have "0 < norm (x - x0)" by simp
       have "closed_segment x0 x \<subseteq> U"
-        by (rule \<open>convex U\<close>[unfolded convex_contains_segment, rule_format, OF \<open>x0 \<in> U\<close> \<open>x \<in> U\<close>])
+        by (simp add: assms closed_segment_subset elim(4))
       from elim have [intro]: "x \<in> U" by auto
       have "?F x - ?F x0 - ?dF (x - x0) =
         integral (cbox a b) (\<lambda>y. f x y - f x0 y - fx x0 y (x - x0))"
@@ -6724,16 +6567,14 @@
       also
       {
         fix t assume t: "t \<in> (cbox a b)"
+        then have deriv:
+          "((\<lambda>x. f x t) has_derivative (fx y t)) (at y within X0 \<inter> U)"
+          if "y \<in> X0 \<inter> U" for y
+          using fx has_derivative_subset that by fastforce
         have seg: "\<And>t. t \<in> {0..1} \<Longrightarrow> x0 + t *\<^sub>R (x - x0) \<in> X0 \<inter> U"
           using \<open>closed_segment x0 x \<subseteq> U\<close>
             \<open>closed_segment x0 x \<subseteq> X0\<close>
           by (force simp: closed_segment_def algebra_simps)
-        from t have deriv:
-          "((\<lambda>x. f x t) has_derivative (fx y t)) (at y within X0 \<inter> U)"
-          if "y \<in> X0 \<inter> U" for y
-          unfolding has_vector_derivative_def[symmetric]
-          using that \<open>x \<in> X0\<close>
-          by (intro has_derivative_subset[OF fx]) auto
         have "\<And>x. x \<in> X0 \<inter> U \<Longrightarrow> onorm (blinfun_apply (fx x t) - (fx x0 t)) \<le> e"
           using fx_bound t
           by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric])
@@ -6848,7 +6689,7 @@
   have fi[simp]: "f n integrable_on (cbox a b)" for n
     by (auto intro!: integrable_continuous assms)
   then obtain I where I: "\<And>n. (f n has_integral I n) (cbox a b)"
-    by atomize_elim (auto simp: integrable_on_def intro!: choice)
+    unfolding integrable_on_def by metis
 
   moreover
   have gi[simp]: "g integrable_on (cbox a b)"
@@ -6885,8 +6726,7 @@
           using elim
           by (intro integral_norm_bound_integral) (auto intro!: integrable_diff)
         also have "\<dots> < e"
-          using \<open>0 < e\<close>
-          by (simp add: e'_def)
+          using \<open>0 < e\<close> by (simp add: e'_def)
         finally show ?case .
       qed
     qed
@@ -6935,7 +6775,7 @@
           "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
   assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
   shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
-  by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
+  by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all)
 
 lemma integration_by_parts:
   fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
@@ -6945,7 +6785,7 @@
           "\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
   assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
   shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
-  by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all)
+  by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (use assms in simp_all)
 
 lemma integrable_by_parts_interior_strong:
   fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
@@ -6973,7 +6813,7 @@
           "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
   assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
   shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
-  by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
+  by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all)
 
 lemma integrable_by_parts:
   fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
@@ -6983,7 +6823,7 @@
           "\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
   assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
   shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
-  by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
+  by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all)
 
 
 subsection \<open>Integration by substitution\<close>
@@ -7433,7 +7273,7 @@
              (\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> integral {c..} (\<lambda>x. exp (-a*x))"
   proof (intro monotone_convergence_increasing allI ballI)
     fix k ::nat
-    have "(\<lambda>x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P)
+    have "(\<lambda>x. exp (-a*x)) integrable_on {c..of_real k}" 
       unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real)
     hence  "(f k) integrable_on {c..of_real k}"
       by (rule integrable_eq) (simp add: f_def)
@@ -7444,7 +7284,8 @@
     have "sequentially \<le> principal {nat \<lceil>x\<rceil>..}" unfolding at_top_def by (simp add: Inf_lower)
     also have "{nat \<lceil>x\<rceil>..} \<subseteq> {k. x \<le> real k}" by auto
     also have "inf (principal \<dots>) (principal {k. \<not>x \<le> real k}) =
-                 principal ({k. x \<le> real k} \<inter> {k. \<not>x \<le> real k})" by simp
+                 principal ({k. x \<le> real k} \<inter> {k. \<not>x \<le> real k})" 
+      by simp
     also have "{k. x \<le> real k} \<inter> {k. \<not>x \<le> real k} = {}" by blast
     finally have "inf sequentially (principal {k. \<not>x \<le> real k}) = bot"
       by (simp add: inf.coboundedI1 bot_unique)
@@ -7468,7 +7309,7 @@
   have "(\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a) \<longlonglongrightarrow> exp (-a*c)/a - 0/a"
     by (intro tendsto_intros filterlim_compose[OF exp_at_bot]
         filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+
-      (insert a, simp_all)
+      (use a in simp_all)
   moreover
   from eventually_gt_at_top[of "nat \<lceil>c\<rceil>"] have "eventually (\<lambda>k. of_nat k > c) sequentially"
     by eventually_elim linarith
@@ -7493,30 +7334,24 @@
   define f where "f = (\<lambda>k x. if x \<in> {inverse (of_nat (Suc k))..c} then x powr a else 0)"
   define F where "F = (\<lambda>k. if inverse (of_nat (Suc k)) \<le> c then
                              c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)"
-  {
-    fix k :: nat
-    have "(f k has_integral F k) {0..c}"
-    proof (cases "inverse (of_nat (Suc k)) \<le> c")
-      case True
-      {
-        fix x assume x: "x \<ge> inverse (1 + real k)"
-        have "0 < inverse (1 + real k)" by simp
-        also note x
-        finally have "x > 0" .
-      } note x = this
-      hence "((\<lambda>x. x powr a) has_integral c powr (a + 1) / (a + 1) -
+  have has_integral_f: "(f k has_integral F k) {0..c}" for k::nat
+  proof (cases "inverse (of_nat (Suc k)) \<le> c")
+    case True
+    have x: "x > 0" if "x \<ge> inverse (1 + real k)" for x
+      by (smt (verit) that inverse_Suc of_nat_Suc)
+    hence "((\<lambda>x. x powr a) has_integral c powr (a + 1) / (a + 1) -
                inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}"
-        using True a by (intro fundamental_theorem_of_calculus)
-           (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const
-              simp: has_real_derivative_iff_has_vector_derivative [symmetric])
-      with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all
-    next
-      case False
-      thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto
-    qed
-  } note has_integral_f = this
-  have integral_f: "integral {0..c} (f k) = F k" for k
-    using has_integral_f[of k] by (rule integral_unique)
+      using True a by (intro fundamental_theorem_of_calculus)
+        (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const
+          simp: has_real_derivative_iff_has_vector_derivative [symmetric])
+    with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all
+  next
+    case False
+    thus ?thesis unfolding f_def F_def 
+      by (subst has_integral_restrict) auto
+  qed
+  then have integral_f: "integral {0..c} (f k) = F k" for k
+    by blast
 
   have A: "(\<lambda>x. x powr a) integrable_on {0..c} \<and>
            (\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> integral {0..c} (\<lambda>x. x powr a)"
@@ -7527,9 +7362,8 @@
     fix k :: nat and x :: real
     {
       assume x: "inverse (real (Suc k)) \<le> x"
-      have "inverse (real (Suc (Suc k))) \<le> inverse (real (Suc k))" by (simp add: field_simps)
-      also note x
-      finally have "inverse (real (Suc (Suc k))) \<le> x" .
+      then have "inverse (real (Suc (Suc k))) \<le> x"
+        using dual_order.trans by fastforce 
     }
     thus "f k x \<le> f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc)
   next
@@ -7607,7 +7441,7 @@
     case False
     have "(f n has_integral 0) {a}" by (rule has_integral_refl)
     hence "(f n has_integral 0) {a..}"
-      by (rule has_integral_on_superset) (insert False, simp_all add: f_def)
+      using False f_def by force
     with False show ?thesis by (simp add: integral_unique)
   qed
 
@@ -7639,7 +7473,7 @@
       with assms have "a powr (e + 1) \<ge> n powr (e + 1)"
         by (intro powr_mono2') simp_all
       with assms show ?thesis by (auto simp: divide_simps F_def integral_f)
-    qed (insert assms, simp add: integral_f F_def field_split_simps)
+    qed (use assms in \<open>simp add: integral_f F_def field_split_simps\<close>)
     thus "bounded (range(\<lambda>k. integral {a..} (f k)))"
       unfolding bounded_iff by (intro exI[of _ "-F a"]) auto
   qed
@@ -7654,9 +7488,9 @@
           filterlim_ident filterlim_real_sequentially | simp)+)
   hence "(\<lambda>n. F n - F a) \<longlonglongrightarrow> -F a" by simp
   ultimately have "(\<lambda>n. integral {a..} (f n)) \<longlonglongrightarrow> -F a" by (blast intro: Lim_transform_eventually)
-  from conjunct2[OF *] and this
-    have "integral {a..} (\<lambda>x. x powr e) = -F a" by (rule LIMSEQ_unique)
-  with conjunct1[OF *] show ?thesis
+  then have "integral {a..} (\<lambda>x. x powr e) = -F a"
+    using "*" LIMSEQ_unique by blast
+  with * show ?thesis
     by (simp add: has_integral_integral F_def)
 qed
 
--- a/src/HOL/Analysis/Homeomorphism.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Homeomorphism.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -210,7 +210,7 @@
     ultimately have *: "continuous (at x within affine hull S) (\<lambda>x. surf (proj x))"
       by (simp add: continuous_within Lim_transform_within_set continuous_on_eq_continuous_within)
     show ?thesis
-      by (intro continuous_within_subset [where s = "affine hull S", OF _ Int_lower2] continuous_intros *)
+      by (intro continuous_within_subset [where S = "affine hull S", OF _ Int_lower2] continuous_intros *)
   qed
   have cont_nosp2: "continuous_on ?CBALL (\<lambda>x. norm x *\<^sub>R ((surf o proj) x))"
     by (simp add: continuous_on_eq_continuous_within cont_nosp)
--- a/src/HOL/Analysis/Homotopy.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Homotopy.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -712,7 +712,7 @@
 
 proposition homotopic_loops_subset:
    "\<lbrakk>homotopic_loops S p q; S \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_loops t p q"
-  by (fastforce simp add: homotopic_loops)
+  by (fastforce simp: homotopic_loops)
 
 proposition homotopic_loops_eq:
    "\<lbrakk>path p; path_image p \<subseteq> S; pathfinish p = pathstart p; \<And>t. t \<in> {0..1} \<Longrightarrow> p(t) = q(t)\<rbrakk>
@@ -1040,17 +1040,15 @@
   have "homotopic_paths S (subpath u w g +++ subpath w v g) 
                           ((subpath u v g +++ subpath v w g) +++ subpath w v g)"
     by (simp add: hom homotopic_paths_join homotopic_paths_sym wvg)
-  also have "homotopic_paths S ((subpath u v g +++ subpath v w g) +++ subpath w v g) 
-                               (subpath u v g +++ subpath v w g +++ subpath w v g)"
+  also have "homotopic_paths S \<dots>  (subpath u v g +++ subpath v w g +++ subpath w v g)"
     using wvg vug \<open>path g\<close>
     by (metis homotopic_paths_assoc homotopic_paths_sym path_image_subpath_commute path_subpath
         pathfinish_subpath pathstart_subpath u v w)
-  also have "homotopic_paths S (subpath u v g +++ subpath v w g +++ subpath w v g)
-                               (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))"
+  also have "homotopic_paths S \<dots> (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))"
     using wvg vug \<open>path g\<close>
     by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl path_image_subpath_commute
         path_subpath pathfinish_subpath pathstart_join pathstart_subpath reversepath_subpath u v)
-  also have "homotopic_paths S (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)"
+  also have "homotopic_paths S \<dots> (subpath u v g)"
     using vug \<open>path g\<close> by (metis homotopic_paths_rid path_image_subpath_commute path_subpath u v)
   finally have "homotopic_paths S (subpath u w g +++ subpath w v g) (subpath u v g)" .
   then show ?thesis
@@ -1077,11 +1075,11 @@
 qed
 
 lemma homotopic_loops_imp_path_component_value:
-   "\<lbrakk>homotopic_loops S p q; 0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> path_component S (p t) (q t)"
-apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs)
-apply (rule_tac x="h \<circ> (\<lambda>u. (u, t))" in exI)
-apply (fastforce elim!: continuous_on_subset intro!: continuous_intros)
-done
+  "\<lbrakk>homotopic_loops S p q; 0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> path_component S (p t) (q t)"
+  apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs)
+  apply (rule_tac x="h \<circ> (\<lambda>u. (u, t))" in exI)
+  apply (fastforce elim!: continuous_on_subset intro!: continuous_intros)
+  done
 
 lemma homotopic_points_eq_path_component:
    "homotopic_loops S (linepath a a) (linepath b b) \<longleftrightarrow> path_component S a b"
@@ -1109,12 +1107,12 @@
 lemma simply_connected_imp_path_connected:
   fixes S :: "_::real_normed_vector set"
   shows "simply_connected S \<Longrightarrow> path_connected S"
-by (simp add: simply_connected_def path_connected_eq_homotopic_points)
+  by (simp add: simply_connected_def path_connected_eq_homotopic_points)
 
 lemma simply_connected_imp_connected:
   fixes S :: "_::real_normed_vector set"
   shows "simply_connected S \<Longrightarrow> connected S"
-by (simp add: path_connected_imp_connected simply_connected_imp_path_connected)
+  by (simp add: path_connected_imp_connected simply_connected_imp_path_connected)
 
 lemma simply_connected_eq_contractible_loop_any:
   fixes S :: "_::real_normed_vector set"
@@ -1123,13 +1121,10 @@
                   \<longrightarrow> homotopic_loops S p (linepath a a))"
         (is "?lhs = ?rhs")
 proof
-  assume ?lhs then show ?rhs
-    unfolding simply_connected_def by force
-next
   assume ?rhs then show ?lhs
     unfolding simply_connected_def
     by (metis pathfinish_in_path_image subsetD  homotopic_loops_trans homotopic_loops_sym)
-qed
+qed (force simp: simply_connected_def)
 
 lemma simply_connected_eq_contractible_loop_some:
   fixes S :: "_::real_normed_vector set"
@@ -1154,15 +1149,7 @@
          S = {} \<or>
          (\<exists>a \<in> S. \<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p
                 \<longrightarrow> homotopic_loops S p (linepath a a))"
-        (is "?lhs = ?rhs")
-proof (cases "S = {}")
-  case True then show ?thesis by force
-next
-  case False
-  then obtain a where "a \<in> S" by blast
-  then show ?thesis
-    by (meson False homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any)
-qed
+  by (meson ex_in_conv homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any)
 
 lemma simply_connected_eq_contractible_path:
   fixes S :: "_::real_normed_vector set"
@@ -1207,14 +1194,12 @@
       using that
       by (smt (verit, best) homotopic_paths_join homotopic_paths_linv homotopic_paths_rid homotopic_paths_sym 
           homotopic_paths_trans pathstart_linepath)
-    also have "homotopic_paths S (p +++ reversepath q +++ q)
-                                 ((p +++ reversepath q) +++ q)"
+    also have "homotopic_paths S \<dots> ((p +++ reversepath q) +++ q)"
       by (simp add: that homotopic_paths_assoc)
-    also have "homotopic_paths S ((p +++ reversepath q) +++ q)
-                                 (linepath (pathstart q) (pathstart q) +++ q)"
+    also have "homotopic_paths S \<dots> (linepath (pathstart q) (pathstart q) +++ q)"
       using * [of "p +++ reversepath q"] that
       by (simp add: homotopic_paths_assoc homotopic_paths_join path_image_join)
-    also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q"
+    also have "homotopic_paths S \<dots> q"
       using that homotopic_paths_lid by blast
     finally show ?thesis .
   qed
@@ -1238,7 +1223,7 @@
     have "path (fst \<circ> p)"
       by (simp add: continuous_on_fst Path_Connected.path_continuous_image [OF \<open>path p\<close>])
     moreover have "path_image (fst \<circ> p) \<subseteq> S"
-      using that by (force simp add: path_image_def)
+      using that by (force simp: path_image_def)
     ultimately have p1: "homotopic_loops S (fst \<circ> p) (linepath a a)"
       using S that
       by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def)
@@ -1291,12 +1276,12 @@
 corollary contractible_imp_connected:
   fixes S :: "_::real_normed_vector set"
   shows "contractible S \<Longrightarrow> connected S"
-by (simp add: contractible_imp_simply_connected simply_connected_imp_connected)
+  by (simp add: contractible_imp_simply_connected simply_connected_imp_connected)
 
 lemma contractible_imp_path_connected:
   fixes S :: "_::real_normed_vector set"
   shows "contractible S \<Longrightarrow> path_connected S"
-by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected)
+  by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected)
 
 lemma nullhomotopic_through_contractible:
   fixes S :: "_::topological_space set"
@@ -1351,7 +1336,7 @@
     with \<open>path_connected U\<close> show ?thesis by blast
   qed
   then have "homotopic_with_canon (\<lambda>h. True) S U (\<lambda>x. c2) (\<lambda>x. c1)"
-    by (auto simp add: path_component homotopic_constant_maps)
+    by (auto simp: path_component homotopic_constant_maps)
   then show ?thesis
     using c1 c2 homotopic_with_symD homotopic_with_trans by blast
 qed
@@ -1419,7 +1404,7 @@
     using \<open>a \<in> S\<close>
     unfolding homotopic_with_def
     apply (rule_tac x="\<lambda>y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI)
-    apply (force simp add: P intro: continuous_intros)
+    apply (force simp: P intro: continuous_intros)
     done
   then show ?thesis
     using that by blast
@@ -1516,18 +1501,17 @@
 where
  "locally P S \<equiv>
         \<forall>w x. openin (top_of_set S) w \<and> x \<in> w
-              \<longrightarrow> (\<exists>u v. openin (top_of_set S) u \<and> P v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w)"
+              \<longrightarrow> (\<exists>U V. openin (top_of_set S) U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> w)"
 
 lemma locallyI:
   assumes "\<And>w x. \<lbrakk>openin (top_of_set S) w; x \<in> w\<rbrakk>
-                  \<Longrightarrow> \<exists>u v. openin (top_of_set S) u \<and> P v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w"
+                  \<Longrightarrow> \<exists>U V. openin (top_of_set S) U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> w"
     shows "locally P S"
 using assms by (force simp: locally_def)
 
 lemma locallyE:
   assumes "locally P S" "openin (top_of_set S) w" "x \<in> w"
-  obtains u v where "openin (top_of_set S) u"
-                    "P v" "x \<in> u" "u \<subseteq> v" "v \<subseteq> w"
+  obtains U V where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> w"
   using assms unfolding locally_def by meson
 
 lemma locally_mono:
@@ -1636,28 +1620,26 @@
     using \<open>f ` S = T\<close> f \<open>V \<subseteq> S\<close> by auto
   have contvf: "continuous_on V f"
     using \<open>V \<subseteq> S\<close> continuous_on_subset f(3) by blast
-  have "f ` V \<subseteq> W"
-    using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto
-  then have contvg: "continuous_on (f ` V) g"
-    using \<open>W \<subseteq> T\<close> continuous_on_subset [OF g(3)] by blast
-  have "V \<subseteq> g ` f ` V"
-    by (metis \<open>V \<subseteq> S\<close> hom homeomorphism_def homeomorphism_of_subsets order_refl)
-  then have homv: "homeomorphism V (f ` V) f g"
-    using \<open>V \<subseteq> S\<close> f by (auto simp: homeomorphism_def contvf contvg)
   have "openin (top_of_set (g ` T)) U"
     using \<open>g ` T = S\<close> by (simp add: osu)
   then have "openin (top_of_set T) (T \<inter> g -` U)"
     using \<open>continuous_on T g\<close> continuous_on_open [THEN iffD1] by blast
   moreover have "\<exists>V. Q V \<and> y \<in> (T \<inter> g -` U) \<and> (T \<inter> g -` U) \<subseteq> V \<and> V \<subseteq> W"
   proof (intro exI conjI)
+    show "f ` V \<subseteq> W"
+      using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto
+    then have contvg: "continuous_on (f ` V) g"
+      using \<open>W \<subseteq> T\<close> continuous_on_subset [OF g(3)] by blast
+    have "V \<subseteq> g ` f ` V"
+      by (metis \<open>V \<subseteq> S\<close> hom homeomorphism_def homeomorphism_of_subsets order_refl)
+    then have homv: "homeomorphism V (f ` V) f g"
+      using \<open>V \<subseteq> S\<close> f by (auto simp: homeomorphism_def contvf contvg)
     show "Q (f ` V)"
       using Q homv \<open>P V\<close> by blast
     show "y \<in> T \<inter> g -` U"
       using T(2) \<open>y \<in> W\<close> \<open>g y \<in> U\<close> by blast
     show "T \<inter> g -` U \<subseteq> f ` V"
       using g(1) image_iff uv(3) by fastforce
-    show "f ` V \<subseteq> W"
-      using \<open>f ` V \<subseteq> W\<close> by blast
   qed
   ultimately show "\<exists>U. openin (top_of_set T) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)"
     by meson
@@ -1692,8 +1674,7 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes f: "linear f" "inj f" and iff: "\<And>S. P (f ` S) \<longleftrightarrow> Q S"
   shows "locally P (f ` S) \<longleftrightarrow> locally Q S"
-  using homeomorphism_locally [of "f`S" _ _ f] linear_homeomorphism_image [OF f]
-  by (metis (no_types, lifting) homeomorphism_image2 iff)
+  by (smt (verit) f homeomorphism_image2 homeomorphism_locally iff linear_homeomorphism_image)
 
 lemma locally_open_map_image:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
@@ -1782,8 +1763,8 @@
     shows "R a b"
 proof -
   have "\<And>a b c. \<lbrakk>a \<in> S; b \<in> S; c \<in> S; R a b\<rbrakk> \<Longrightarrow> R a c"
-    apply (rule connected_induction_simple [OF \<open>connected S\<close>], simp_all)
-    by (meson local.sym local.trans opI openin_imp_subset subsetCE)
+    by (smt (verit, ccfv_threshold) connected_induction_simple [OF \<open>connected S\<close>] 
+            assms openin_imp_subset subset_eq)
   then show ?thesis by (metis etc opI)
 qed
 
@@ -1833,13 +1814,13 @@
               compact v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<inter> T"
           if "open T" "x \<in> S" "x \<in> T" for x T
   proof -
-    obtain u v where uv: "x \<in> u" "u \<subseteq> v" "v \<subseteq> S" "compact v" "openin (top_of_set S) u"
+    obtain U V where uv: "x \<in> U" "U \<subseteq> V" "V \<subseteq> S" "compact V" "openin (top_of_set S) U"
       using r [OF \<open>x \<in> S\<close>] by auto
     obtain e where "e>0" and e: "cball x e \<subseteq> T"
       using open_contains_cball \<open>open T\<close> \<open>x \<in> T\<close> by blast
     show ?thesis
-      apply (rule_tac x="(S \<inter> ball x e) \<inter> u" in exI)
-      apply (rule_tac x="cball x e \<inter> v" in exI)
+      apply (rule_tac x="(S \<inter> ball x e) \<inter> U" in exI)
+      apply (rule_tac x="cball x e \<inter> V" in exI)
       using that \<open>e > 0\<close> e uv
       apply auto
       done
@@ -1860,16 +1841,8 @@
   shows "locally compact S \<longleftrightarrow>
          (\<forall>x \<in> S. \<exists>U. x \<in> U \<and>
                     openin (top_of_set S) U \<and> compact(closure U) \<and> closure U \<subseteq> S)"
-        (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs
-    by (meson bounded_subset closure_minimal compact_closure compact_imp_bounded 
-              compact_imp_closed dual_order.trans locally_compactE)
-next
-  assume ?rhs then show ?lhs
-    by (meson closure_subset locally_compact)
-qed
+  by (smt (verit, ccfv_threshold) bounded_subset closure_closed closure_mono closure_subset 
+      compact_closure compact_imp_closed order.trans locally_compact)
 
 lemma locally_compact_Int_cball:
   fixes S :: "'a :: heine_borel set"
@@ -2008,8 +1981,9 @@
   then show ?rhs
     unfolding locally_def
     apply (elim all_forward imp_forward asm_rl exE)
-    apply (rule_tac x = "u \<inter> ball x 1" in exI)
-    apply (rule_tac x = "v \<inter> cball x 1" in exI)
+    apply (rename_tac U V)
+    apply (rule_tac x = "U \<inter> ball x 1" in exI)
+    apply (rule_tac x = "V \<inter> cball x 1" in exI)
     apply (force intro: openin_trans)
     done
 next
@@ -2019,7 +1993,7 @@
 
 lemma locally_compact_openin_Un:
   fixes S :: "'a::euclidean_space set"
-  assumes LCS: "locally compact S" and LCT:"locally compact T"
+  assumes LCS: "locally compact S" and LCT: "locally compact T"
       and opS: "openin (top_of_set (S \<union> T)) S"
       and opT: "openin (top_of_set (S \<union> T)) T"
     shows "locally compact (S \<union> T)"
@@ -2453,7 +2427,7 @@
     "\<And>v x. \<lbrakk>openin (top_of_set S) v; x \<in> v\<rbrakk>
               \<Longrightarrow> \<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v"
    shows "locally path_connected S"
-  by (force simp add: locally_def dest: assms)
+  by (force simp: locally_def dest: assms)
 
 lemma locally_path_connected_2:
   assumes "locally path_connected S"
@@ -2525,7 +2499,7 @@
 proof
   assume ?lhs
   then show ?rhs
-    by (fastforce simp add: locally_connected)
+    by (fastforce simp: locally_connected)
 next
   assume ?rhs
   have *: "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> c"
@@ -3577,7 +3551,7 @@
   show thesis
     using homotopic_with_compose_continuous_map_right
            [OF homotopic_with_compose_continuous_map_left [OF b g] f]
-    by (force simp add: that)
+    by (force simp: that)
 qed
 
 lemma nullhomotopic_into_contractible_space:
@@ -3901,7 +3875,7 @@
     by (rule homotopic_with_compose_continuous_left [where Y=T])
        (use f in \<open>auto simp: hom homotopic_with_symD\<close>)
   ultimately show ?thesis
-    using that homotopic_with_trans by (fastforce simp add: o_def)
+    using that homotopic_with_trans by (fastforce simp: o_def)
 qed
 
 lemma homotopy_eqv_cohomotopic_triviality_null:
@@ -3942,7 +3916,7 @@
     by (rule homotopic_with_compose_continuous_right [where X=T])
        (use f in \<open>auto simp: hom homotopic_with_symD\<close>)
   ultimately show ?thesis
-    using homotopic_with_trans by (fastforce simp add: o_def)
+    using homotopic_with_trans by (fastforce simp: o_def)
 qed
 
 lemma homotopy_eqv_homotopic_triviality_null:
@@ -5002,7 +4976,7 @@
       using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def)
   qed (use affine_hull_open assms that in auto)
   then show ?thesis
-    using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp add: intro!: that)
+    using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp: intro!: that)
 next
   case False
   with DIM_positive have "DIM('a) = 1"
@@ -5089,7 +5063,7 @@
       using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def)
   qed
   then show ?thesis
-    using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp add: intro!: that)
+    using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp: intro!: that)
 next
   case False
   with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith
@@ -5173,9 +5147,9 @@
             using \<open>T \<subseteq> affine hull S\<close> h by auto
         qed
         show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x"
-          using h j hom homeomorphism_apply1 by (fastforce simp add: f'_def g'_def)
+          using h j hom homeomorphism_apply1 by (fastforce simp: f'_def g'_def)
         show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y"
-          using h j hom homeomorphism_apply2 by (fastforce simp add: f'_def g'_def)
+          using h j hom homeomorphism_apply2 by (fastforce simp: f'_def g'_def)
       qed
     next
       have \<section>: "\<And>x y. \<lbrakk>x \<in> affine hull S; h x = h y; y \<in> S\<rbrakk> \<Longrightarrow> x \<in> S"
@@ -5241,7 +5215,7 @@
            and d: "\<And>x'. \<lbrakk>dist x' a \<le> r; x' \<noteq> a; dist x' x < d\<rbrakk> \<Longrightarrow>
                                  dist (g x') (g x) < e"
           using contg False x \<open>e>0\<close>
-          unfolding continuous_on_iff by (fastforce simp add: dist_commute intro: that)
+          unfolding continuous_on_iff by (fastforce simp: dist_commute intro: that)
         show ?thesis
           using \<open>d > 0\<close> \<open>x \<noteq> a\<close>
           by (rule_tac x="min d (norm(x - a))" in exI)
@@ -5287,7 +5261,7 @@
   next
     assume ?rhs
     then show ?lhs
-      using equal continuous_on_const by (force simp add: homotopic_with)
+      using equal continuous_on_const by (force simp: homotopic_with)
   qed
 next
   case greater
--- a/src/HOL/Analysis/Polytope.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Polytope.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -1355,7 +1355,7 @@
               case False
               have "x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c) /\<^sub>R ux"
                 by (metis \<open>ux \<noteq> 0\<close> uxx mult.commute right_inverse scaleR_one scaleR_scaleR)
-              also have "... = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c"
+              also have "\<dots> = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c"
                 using \<open>ux \<noteq> 0\<close> equx apply (auto simp: field_split_simps)
                 by (metis add.commute add_diff_eq add_divide_distrib diff_add_cancel scaleR_add_left)
               finally have "x = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" .
@@ -1501,7 +1501,7 @@
 proof
   have "?lhs \<subseteq> convex hull {x. x extreme_point_of S}"
     using Krein_Milman_Minkowski assms by blast
-  also have "... \<subseteq> ?rhs"
+  also have "\<dots> \<subseteq> ?rhs"
   proof (rule hull_mono)
     show "{x. x extreme_point_of S} \<subseteq> frontier S"
       using closure_subset
@@ -1511,7 +1511,7 @@
 next
   have "?rhs \<subseteq> convex hull S"
     by (metis Diff_subset \<open>compact S\<close> closure_closed compact_eq_bounded_closed frontier_def hull_mono)
-  also have "... \<subseteq> ?lhs"
+  also have "\<dots> \<subseteq> ?lhs"
     by (simp add: \<open>convex S\<close> hull_same)
   finally show "?rhs \<subseteq> ?lhs" .
 qed
@@ -1626,12 +1626,11 @@
   done
 
 lemma polyhedron_UNIV [iff]: "polyhedron UNIV"
-  unfolding polyhedron_def
-  by (rule_tac x="{}" in exI) auto
+  using polyhedron_def by auto
 
 lemma polyhedron_Inter [intro,simp]:
-   "\<lbrakk>finite F; \<And>S. S \<in> F \<Longrightarrow> polyhedron S\<rbrakk> \<Longrightarrow> polyhedron(\<Inter>F)"
-by (induction F rule: finite_induct) auto
+  "\<lbrakk>finite F; \<And>S. S \<in> F \<Longrightarrow> polyhedron S\<rbrakk> \<Longrightarrow> polyhedron(\<Inter>F)"
+  by (induction F rule: finite_induct) auto
 
 
 lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)"
@@ -1640,10 +1639,7 @@
   have "\<exists>a. a \<noteq> 0 \<and> (\<exists>b. {x. i \<bullet> x \<le> -1} = {x. a \<bullet> x \<le> b})"
     by (rule_tac x="i" in exI) (force simp: i_def SOME_Basis nonzero_Basis)
   moreover have "\<exists>a b. a \<noteq> 0 \<and> {x. -i \<bullet> x \<le> - 1} = {x. a \<bullet> x \<le> b}"
-      apply (rule_tac x="-i" in exI)
-      apply (rule_tac x="-1" in exI)
-      apply (simp add: i_def SOME_Basis nonzero_Basis)
-      done
+    by (metis Basis_zero SOME_Basis i_def neg_0_equal_iff_equal)
   ultimately show ?thesis
     unfolding polyhedron_def
     by (rule_tac x="{{x. i \<bullet> x \<le> -1}, {x. -i \<bullet> x \<le> -1}}" in exI) force
@@ -1664,7 +1660,7 @@
 lemma polyhedron_halfspace_ge:
   fixes a :: "'a :: euclidean_space"
   shows "polyhedron {x. a \<bullet> x \<ge> b}"
-using polyhedron_halfspace_le [of "-a" "-b"] by simp
+  using polyhedron_halfspace_le [of "-a" "-b"] by simp
 
 lemma polyhedron_hyperplane:
   fixes a :: "'a :: euclidean_space"
@@ -1679,7 +1675,7 @@
 lemma affine_imp_polyhedron:
   fixes S :: "'a :: euclidean_space set"
   shows "affine S \<Longrightarrow> polyhedron S"
-by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S])
+  by (metis affine_hull_finite_intersection_hyperplanes hull_same polyhedron_Inter polyhedron_hyperplane)
 
 lemma polyhedron_imp_closed:
   fixes S :: "'a :: euclidean_space set"
@@ -1694,7 +1690,7 @@
 lemma polyhedron_affine_hull:
   fixes S :: "'a :: euclidean_space set"
   shows "polyhedron(affine hull S)"
-by (simp add: affine_imp_polyhedron)
+  by (simp add: affine_imp_polyhedron)
 
 
 subsection\<open>Canonical polyhedron representation making facial structure explicit\<close>
@@ -1704,14 +1700,7 @@
   shows "polyhedron S \<longleftrightarrow>
            (\<exists>F. finite F \<and> S = (affine hull S) \<inter> \<Inter>F \<and>
                 (\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}))"
-        (is "?lhs = ?rhs")
-proof
-  assume ?lhs then show ?rhs
-    using hull_subset polyhedron_def by fastforce
-next
-  assume ?rhs then show ?lhs
-    by (metis polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le)
-qed
+  by (metis hull_subset inf.absorb_iff2 polyhedron_Int polyhedron_affine_hull polyhedron_def)
 
 proposition rel_interior_polyhedron_explicit:
   assumes "finite F"
@@ -1744,13 +1733,13 @@
     define \<xi> where "\<xi> = min (1/2) (e / 2 / norm(z - x))"
     have "norm (\<xi> *\<^sub>R x - \<xi> *\<^sub>R z) = norm (\<xi> *\<^sub>R (x - z))"
       by (simp add: \<xi>_def algebra_simps norm_mult)
-    also have "... = \<xi> * norm (x - z)"
+    also have "\<dots> = \<xi> * norm (x - z)"
       using \<open>e > 0\<close> by (simp add: \<xi>_def)
-    also have "... < e"
+    also have "\<dots> < e"
       using \<open>z \<noteq> x\<close> \<open>e > 0\<close> by (simp add: \<xi>_def min_def field_split_simps norm_minus_commute)
     finally have lte: "norm (\<xi> *\<^sub>R x - \<xi> *\<^sub>R z) < e" .
     have \<xi>_aff: "\<xi> *\<^sub>R z + (1 - \<xi>) *\<^sub>R x \<in> affine hull S"
-      by (metis \<open>x \<in> S\<close> add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff)
+      by (simp add: \<open>x \<in> S\<close> hull_inc mem_affine zaff)
     have "\<xi> *\<^sub>R z + (1 - \<xi>) *\<^sub>R x \<in> S"
       using ins [OF _ \<xi>_aff] by (simp add: algebra_simps lte)
     then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \<in> S"
@@ -1809,7 +1798,7 @@
       have "a h \<noteq> 0" and "h = {x. a h \<bullet> x \<le> b h}" "h \<inter> \<Inter>F = \<Inter>F"
         using \<open>h \<in> F\<close> ab by auto
       then have "(affine hull S) \<inter> {x. a h \<bullet> x \<le> b h} \<noteq> {}"
-        by (metis (no_types) affine_hull_eq_empty inf.absorb_iff2 inf_assoc inf_bot_right inf_commute seq that(2))
+        by (metis affine_hull_eq_empty inf.absorb_iff1 inf_assoc inf_bot_left seq that(2))
       moreover have "\<not> (affine hull S \<subseteq> {x. a h \<bullet> x \<le> b h})"
         using \<open>h = {x. a h \<bullet> x \<le> b h}\<close> that(2) by blast
       ultimately show ?thesis
@@ -1822,14 +1811,14 @@
                   affine hull S \<inter> {x. a h \<bullet> x \<le> b h} = affine hull S \<inter> h \<and>
                   (\<forall>w \<in> affine hull S. (w + a h) \<in> affine hull S)"
       by metis
-    have seq2: "S = affine hull S \<inter> (\<Inter>h\<in>{h \<in> F. \<not> affine hull S \<subseteq> h}. {x. a h \<bullet> x \<le> b h})"
-      by (subst seq) (auto simp: ab INT_extend_simps)
+    let ?F = "(\<lambda>h. {x. a h \<bullet> x \<le> b h}) ` {h \<in> F. \<not> affine hull S \<subseteq> h}"
     show ?thesis
-      apply (rule_tac x="(\<lambda>h. {x. a h \<bullet> x \<le> b h}) ` {h. h \<in> F \<and> \<not>(affine hull S \<subseteq> h)}" in exI)
-      apply (intro conjI seq2)
-        using \<open>finite F\<close> apply force
-       using ab apply blast
-       done
+    proof (intro exI conjI)
+      show "finite ?F"
+        using \<open>finite F\<close> by force
+      show "S = affine hull S \<inter> \<Inter> ?F"
+        by (subst seq) (auto simp: ab INT_extend_simps)
+    qed (use ab in blast)
   qed
 next
   assume ?rhs then show ?lhs
@@ -1887,12 +1876,7 @@
          (\<exists>F. finite F \<and> S = (affine hull S) \<inter> \<Inter>F \<and>
               (\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}) \<and>
               (\<forall>F'. F' \<subset> F \<longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'))"
-     (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs
-    by (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward)
-qed (auto simp: polyhedron_Int_affine elim!: ex_forward)
+  by (metis polyhedron_Int_affine polyhedron_Int_affine_parallel_minimal)
 
 proposition facet_of_polyhedron_explicit:
   assumes "finite F"
@@ -1928,7 +1912,7 @@
     have "x \<in> h" using that xint by auto
     then have able: "a h \<bullet> x \<le> b h"
       using faceq that by blast
-    also have "... < a h \<bullet> z" using \<open>z \<notin> h\<close> faceq [OF that] xint by auto
+    also have "\<dots> < a h \<bullet> z" using \<open>z \<notin> h\<close> faceq [OF that] xint by auto
     finally have xltz: "a h \<bullet> x < a h \<bullet> z" .
     define l where "l = (b h - a h \<bullet> x) / (a h \<bullet> z - a h \<bullet> x)"
     define w where "w = (1 - l) *\<^sub>R x + l *\<^sub>R z"
@@ -1942,14 +1926,15 @@
       moreover have "l * (a i \<bullet> z) \<le> l * b i"
       proof (rule mult_left_mono)
         show "a i \<bullet> z \<le> b i"
-          by (metis Diff_insert_absorb Inter_iff Set.set_insert \<open>h \<in> F\<close> faceq insertE mem_Collect_eq that zint)
+          by (metis DiffI Inter_iff empty_iff faceq insertE mem_Collect_eq that zint)
       qed (use \<open>0 < l\<close> in auto)
       ultimately show ?thesis by (simp add: w_def algebra_simps)
     qed
     have weq: "a h \<bullet> w = b h"
       using xltz unfolding w_def l_def
       by (simp add: algebra_simps) (simp add: field_simps)
-    have faceS: "S \<inter> {x. a h \<bullet> x = b h} face_of S"
+    let ?F = "{x. a h \<bullet> x = b h}"
+    have faceS: "S \<inter> ?F face_of S"
     proof (rule face_of_Int_supporting_hyperplane_le)
       show "\<And>x. x \<in> S \<Longrightarrow> a h \<bullet> x \<le> b h"
         using faceq seq that by fastforce
@@ -1960,15 +1945,18 @@
       using \<open>a h \<bullet> w = b h\<close> awlt faceq less_eq_real_def by blast
     ultimately have "w \<in> S"
       using seq by blast
-    with weq have ne: "S \<inter> {x. a h \<bullet> x = b h} \<noteq> {}" by blast
-    moreover have "affine hull (S \<inter> {x. a h \<bullet> x = b h}) = (affine hull S) \<inter> {x. a h \<bullet> x = b h}"
+    with weq have ne: "S \<inter> ?F \<noteq> {}" by blast
+    moreover have "affine hull (S \<inter> ?F) = (affine hull S) \<inter> ?F"
     proof
-      show "affine hull (S \<inter> {x. a h \<bullet> x = b h}) \<subseteq> affine hull S \<inter> {x. a h \<bullet> x = b h}"
-        apply (intro Int_greatest hull_mono Int_lower1)
-        apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2)
-        done
+      show "affine hull (S \<inter> ?F) \<subseteq> affine hull S \<inter> ?F"
+      proof -
+        have "affine hull (S \<inter> ?F) \<subseteq> affine hull S"
+          by (simp add: hull_mono)
+        then show ?thesis
+          by (simp add: affine_hyperplane subset_hull)
+      qed
     next
-      show "affine hull S \<inter> {x. a h \<bullet> x = b h} \<subseteq> affine hull (S \<inter> {x. a h \<bullet> x = b h})"
+      show "affine hull S \<inter> ?F \<subseteq> affine hull (S \<inter> ?F)"
       proof
         fix y
         assume yaff: "y \<in> affine hull S \<inter> {y. a h \<bullet> y = b h}"
@@ -2014,7 +2002,7 @@
             case False
             with \<open>0 < inff\<close> have "inff * (a j \<bullet> y - a j \<bullet> w) \<le> 0"
               by (simp add: mult_le_0_iff)
-            also have "... < b j - a j \<bullet> w"
+            also have "\<dots> < b j - a j \<bullet> w"
               by (simp add: awlt that)
             finally show ?thesis by simp
           qed
@@ -2050,7 +2038,7 @@
           by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff])
       qed
     qed
-    ultimately have "aff_dim (affine hull (S \<inter> {x. a h \<bullet> x = b h})) = aff_dim S - 1"
+    ultimately have "aff_dim (affine hull (S \<inter> ?F)) = aff_dim S - 1"
       using \<open>b h < a h \<bullet> z\<close> zaff by (force simp: aff_dim_affine_Int_hyperplane)
     then show ?thesis
       by (simp add: ne faceS facet_of_def)
@@ -2755,7 +2743,7 @@
       by (meson \<open>0 < a\<close> divide_less_eq floor_eq_iff)
     have "?n * a \<le> a + x"
       using \<open>a>0\<close> by (simp add: distrib_right floor_divide_lower)
-    also have "... < y"
+    also have "\<dots> < y"
       by (rule 1)
     finally have "?n * a < y" .
     with x show ?thesis
@@ -2767,7 +2755,7 @@
       by (meson \<open>0 < a\<close> divide_less_eq floor_eq_iff)
     have "?n * a \<le> a + y"
       using \<open>a>0\<close> by (simp add: distrib_right floor_divide_lower)
-    also have "... < x"
+    also have "\<dots> < x"
       by (rule 2)
     finally have "?n * a < x" .
     then show ?thesis
@@ -2891,7 +2879,7 @@
             using B \<open>X \<in> \<F>'\<close> eq that by blast+
           have "norm (x - y) \<le> (\<Sum>b\<in>Basis. \<bar>(x-y) \<bullet> b\<bar>)"
             by (rule norm_le_l1)
-          also have "... \<le> of_nat (DIM('a)) * (e / 2 / DIM('a))"
+          also have "\<dots> \<le> of_nat (DIM('a)) * (e / 2 / DIM('a))"
           proof (rule sum_bounded_above)
             fix i::'a
             assume "i \<in> Basis"
@@ -2930,12 +2918,12 @@
                 using  I' [OF \<open>n \<in> C\<close> refl] n  by auto
             qed
           qed
-          also have "... = e / 2"
+          also have "\<dots> = e / 2"
             by simp
           finally show ?thesis .
         qed
       qed (use \<open>0 < e\<close> in force)
-      also have "... < e"
+      also have "\<dots> < e"
         by (simp add: \<open>0 < e\<close>)
       finally show ?thesis .
     qed
@@ -3240,7 +3228,7 @@
           using \<open>K \<in> \<U>\<close> C\<U> by blast
         have "K \<le> rel_frontier C"
           by (simp add: \<open>K \<subseteq> rel_frontier C\<close>)
-        also have "... \<le> C"
+        also have "\<dots> \<le> C"
           by (simp add: \<open>closed C\<close> rel_frontier_def subset_iff)
         finally have "K \<subseteq> C" .
         have "L \<inter> C face_of C"
@@ -3292,7 +3280,7 @@
               by (auto simp: \<open>\<not> affine_dependent I\<close> aff_independent_finite finite_imp_compact)
             moreover have "F face_of convex hull insert ?z I"
               by (metis S \<open>F face_of S\<close> \<open>K = convex hull I\<close> convex_hull_eq_empty convex_hull_insert_segments hull_hull)
-            ultimately obtain J where "J \<subseteq> insert ?z I" "F = convex hull J"
+            ultimately obtain J where J: "J \<subseteq> insert ?z I" "F = convex hull J"
               using face_of_convex_hull_subset [of "insert ?z I" F] by auto
             show ?thesis
             proof (cases "?z \<in> J")
@@ -3320,7 +3308,7 @@
               case False
               then have "F \<in> \<U>"
                 using face_of_convex_hull_affine_independent [OF \<open>\<not> affine_dependent I\<close>]
-                by (metis Int_absorb2 Int_insert_right_if0 \<open>F = convex hull J\<close> \<open>J \<subseteq> insert ?z I\<close> \<open>K = convex hull I\<close> face\<U> inf_le2 \<open>K \<in> \<U>\<close>)
+                by (metis J \<open>K = convex hull I\<close> face\<U> subset_insert \<open>K \<in> \<U>\<close>)
               then show "F \<in> \<U> \<union> ?\<T>"
                 by blast
             qed
@@ -3366,7 +3354,7 @@
             using ahK_C_disjoint \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>K \<subseteq> rel_frontier C\<close> affine_hull_convex_hull z by blast
           have "X \<inter> K face_of K"
             by (simp add: XY(1) \<open>K \<in> \<U>\<close> faceI\<U> inf_commute)
-          also have "... face_of convex hull insert ?z K"
+          also have "\<dots> face_of convex hull insert ?z K"
             by (metis I Keq \<open>?z \<notin> affine hull I\<close> aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert)
           finally have "X \<inter> K face_of convex hull insert ?z K" .
           then show ?thesis
@@ -3413,7 +3401,7 @@
                 using \<open>L \<subseteq> rel_frontier D\<close> by auto
               have "convex hull insert (SOME z. z \<in> rel_interior C) (K \<inter> L) face_of
                     convex hull insert (SOME z. z \<in> rel_interior C) K"
-                by (metis face_of_polytope_insert2 "*" IntI \<open>C \<in> \<N>\<close> aff_independent_finite ahK_C_disjoint empty_iff faceI\<U> polytope_def z \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close>\<open>K \<subseteq> rel_frontier C\<close>)
+                by (metis IntI \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>K \<subseteq> rel_frontier C\<close> \<open>L \<in> \<U>\<close> ahK_C_disjoint empty_iff faceI\<U> face_of_polytope_insert2 simpl\<U> simplex_imp_polytope z)
               then show ?thesis
                 using True X Y \<open>K \<subseteq> rel_frontier C\<close> \<open>L \<subseteq> rel_frontier C\<close> \<open>convex C\<close> \<open>convex K\<close> \<open>convex L\<close> convex_hull_insert_Int_eq z by force
             next
@@ -3442,7 +3430,7 @@
               qed
               finally have CD: "C \<inter> (rel_interior D) = {}" .
               have zKC: "(convex hull insert ?z K) \<subseteq> C"
-                by (metis DiffE \<open>C \<in> \<N>\<close> \<open>K \<subseteq> rel_frontier C\<close> closed\<N> closure_closed convex\<N> hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z)
+                by (metis \<open>K \<subseteq> C\<close> \<open>convex C\<close> in_mono insert_subsetI rel_interior_subset subset_hull z)
               have "disjnt (convex hull insert (SOME z. z \<in> rel_interior C) K) (rel_interior D)"
                 using zKC CD by (force simp: disjnt_def)
               then have eq: "convex hull (insert ?z K) \<inter> convex hull (insert ?w L) =
@@ -3454,7 +3442,7 @@
                 by (simp add: \<open>C \<in> \<N>\<close> convex\<N>)
               have "convex hull (insert ?z K) \<inter> L = L \<inter> convex hull (insert ?z K)"
                 by blast
-              also have "... = convex hull K \<inter> L"
+              also have "\<dots> = convex hull K \<inter> L"
               proof (subst Int_convex_hull_insert_rel_exterior [OF \<open>convex C\<close> \<open>K \<subseteq> C\<close> z])
                 have "(C \<inter> D) \<inter> rel_interior C = {}"
                 proof (rule face_of_disjoint_rel_interior)
@@ -3482,26 +3470,26 @@
               finally have chKL: "convex hull (insert ?z K) \<inter> L = convex hull K \<inter> L" .
               have "convex hull insert ?z K \<inter> convex hull L face_of K"
                 by (simp add: \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close> ch_id chKL faceI\<U>)
-              also have "... face_of convex hull insert ?z K"
+              also have "\<dots> face_of convex hull insert ?z K"
               proof -
                 obtain I where I: "\<not> affine_dependent I" "K = convex hull I"
                   using * [OF \<open>K \<in> \<U>\<close>] by auto
                 then have "\<And>a. a \<notin> rel_interior C \<or> a \<notin> affine hull I"
                   using ahK_C_disjoint \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>K \<subseteq> rel_frontier C\<close> affine_hull_convex_hull by blast
                 then show ?thesis
-                  by (metis I affine_independent_insert face_of_convex_hull_affine_independent hull_insert subset_insertI z)
+                  by (metis I \<open>convex K\<close> aff_independent_finite face_of_convex_hull_insert_eq face_of_refl hull_insert z)
               qed
               finally have 1: "convex hull insert ?z K \<inter> convex hull L face_of convex hull insert ?z K" .
               have "convex hull insert ?z K \<inter> convex hull L face_of L"
                 by (metis \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close> chKL ch_id faceI\<U> inf_commute)
-              also have "... face_of convex hull insert ?w L"
+              also have "\<dots> face_of convex hull insert ?w L"
               proof -
                 obtain I where I: "\<not> affine_dependent I" "L = convex hull I"
                   using * [OF \<open>L \<in> \<U>\<close>] by auto
                 then have "\<And>a. a \<notin> rel_interior D \<or> a \<notin> affine hull I"
                   using \<open>D \<in> \<N>\<close> \<open>L \<in> \<U>\<close> \<open>L \<subseteq> rel_frontier D\<close> affine_hull_convex_hull ahK_C_disjoint by blast
                 then show ?thesis
-                  by (metis I aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert w)
+                  by (metis I \<open>convex L\<close> aff_independent_finite face_of_convex_hull_insert face_of_refl hull_insert w)
               qed
               finally have 2: "convex hull insert ?z K \<inter> convex hull L face_of convex hull insert ?w L" .
               show ?thesis
@@ -3767,18 +3755,13 @@
           if "T \<in> {U - C |C. C \<in> \<S> \<and> aff_dim C < aff_dim U}" for T
           using that dense_complement_convex_closed \<open>closed U\<close> \<open>convex U\<close> by auto
       qed
-      also have "... \<subseteq> closure ?lhs"
+      also have "\<dots> \<subseteq> closure ?lhs"
       proof -
         obtain C where "C \<in> \<S>" "aff_dim C < aff_dim U"
           by (metis False Sup_upper aff_dim_subset eq eq_iff not_le)
-        have "\<exists>X. X \<in> \<S> \<and> aff_dim X = aff_dim U \<and> x \<in> X"
+        then have "\<exists>X. X \<in> \<S> \<and> aff_dim X = aff_dim U \<and> x \<in> X"
           if "\<And>V. (\<exists>C. V = U - C \<and> C \<in> \<S> \<and> aff_dim C < aff_dim U) \<Longrightarrow> x \<in> V" for x
-        proof -
-          have "x \<in> U \<and> x \<in> \<Union>\<S>"
-            using \<open>C \<in> \<S>\<close> \<open>aff_dim C < aff_dim U\<close> eq that by blast
-          then show ?thesis
-            by (metis Diff_iff Sup_upper Union_iff aff_dim_subset dual_order.order_iff_strict eq that)
-        qed
+          by (metis Diff_iff Sup_upper UnionE aff_dim_subset eq order_less_le that)
         then show ?thesis
           by (auto intro!: closure_mono)
       qed
--- a/src/HOL/Analysis/Smooth_Paths.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Smooth_Paths.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -16,11 +16,8 @@
 proof -
   have "path_image \<gamma> homeomorphic {0..1::real}"
     by (simp add: assms homeomorphic_arc_image_interval)
-  then
-  show ?thesis
-    apply (rule path_connected_complement_homeomorphic_convex_compact)
-      apply (auto simp: assms)
-    done
+  then show ?thesis
+    by (intro path_connected_complement_homeomorphic_convex_compact) (auto simp: assms)
 qed
 
 lemma connected_arc_complement:
@@ -39,15 +36,11 @@
     using assms connected_arc_image connected_convex_1_gen inside_convex by blast
 next
   case False
+  then have "connected (- path_image \<gamma>)"
+      by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym assms connected_arc_complement not_less_eq_eq)
+    then
   show ?thesis
-  proof (rule inside_bounded_complement_connected_empty)
-    show "connected (- path_image \<gamma>)"
-      apply (rule connected_arc_complement [OF assms])
-      using False
-      by (metis DIM_ge_Suc0 One_nat_def Suc_1 not_less_eq_eq order_class.order.antisym)
-    show "bounded (path_image \<gamma>)"
-      by (simp add: assms bounded_arc_image)
-  qed
+    by (simp add: assms bounded_arc_image inside_bounded_complement_connected_empty)
 qed
 
 lemma inside_simple_curve_imp_closed:
@@ -59,18 +52,18 @@
 subsection \<open>Piecewise differentiability of paths\<close>
 
 lemma continuous_on_joinpaths_D1:
-    "continuous_on {0..1} (g1 +++ g2) \<Longrightarrow> continuous_on {0..1} g1"
-  apply (rule continuous_on_eq [of _ "(g1 +++ g2) \<circ> ((*)(inverse 2))"])
-  apply (rule continuous_intros | simp)+
-  apply (auto elim!: continuous_on_subset simp: joinpaths_def)
-  done
+  assumes "continuous_on {0..1} (g1 +++ g2)"
+  shows "continuous_on {0..1} g1"
+proof (rule continuous_on_eq)
+  have "continuous_on {0..1/2} (g1 +++ g2)"
+    using assms continuous_on_subset split_01 by auto
+  then show "continuous_on {0..1} (g1 +++ g2 \<circ> (*) (inverse 2))"
+    by (intro continuous_intros) force
+qed (auto simp: joinpaths_def)
 
 lemma continuous_on_joinpaths_D2:
     "\<lbrakk>continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> continuous_on {0..1} g2"
-  apply (rule continuous_on_eq [of _ "(g1 +++ g2) \<circ> (\<lambda>x. inverse 2*x + 1/2)"])
-  apply (rule continuous_intros | simp)+
-  apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def)
-  done
+  using path_def path_join by blast
 
 lemma piecewise_differentiable_D1:
   assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}"
@@ -145,8 +138,8 @@
   proof (rule differentiable_transform_within)
     show "g1 +++ g2 \<circ> (*) (inverse 2) differentiable at x"
       using that g12D
-      apply (simp only: joinpaths_def)
-      by (rule differentiable_chain_at derivative_intros | force)+
+      unfolding joinpaths_def
+      by (intro differentiable_chain_at derivative_intros | force)+
     show "\<And>x'. \<lbrakk>dist x' x < dist (x/2) (1/2)\<rbrakk>
           \<Longrightarrow> (g1 +++ g2 \<circ> (*) (inverse 2)) x' = g1 x'"
       using that by (auto simp: dist_real_def joinpaths_def)
@@ -173,10 +166,9 @@
   qed
   have "continuous_on ({0..1} - insert 1 ((*) 2 ` S))
                       ((\<lambda>x. 1/2 * vector_derivative (g1 \<circ> (*)2) (at x)) \<circ> (*)(1/2))"
-    apply (rule continuous_intros)+
     using coDhalf
-    apply (simp add: scaleR_conv_of_real image_set_diff image_image)
-    done
+    apply (intro continuous_intros)
+    by (simp add: scaleR_conv_of_real image_set_diff image_image)
   then have con_g1: "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) (\<lambda>x. vector_derivative g1 (at x))"
     by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
   have "continuous_on {0..1} g1"
@@ -201,7 +193,7 @@
   proof (rule differentiable_transform_within)
     show "g1 +++ g2 \<circ> (\<lambda>x. (x + 1) / 2) differentiable at x"
       using g12D that
-      apply (simp only: joinpaths_def)
+      unfolding joinpaths_def
       apply (drule_tac x= "(x+1) / 2" in bspec, force simp: field_split_simps)
       apply (rule differentiable_chain_at derivative_intros | force)+
       done
@@ -238,10 +230,7 @@
   have "continuous_on {0..1} g2"
     using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast
   with \<open>finite S\<close> show ?thesis
-    apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-    apply (rule_tac x="insert 0 ((\<lambda>x. 2 * x - 1) ` S)" in exI)
-    apply (simp add: g2D con_g2)
-  done
+    by (meson C1_differentiable_on_eq con_g2 finite_imageI finite_insert g2D piecewise_C1_differentiable_on_def)
 qed
 
 
@@ -305,12 +294,8 @@
     next
       show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \<circ> g) (at t)"
           when "t \<in> {0..1} - S" for t
-        proof (rule vector_derivative_chain_at_general[symmetric])
-          show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that)
-        next
-          have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis
-          then show "f field_differentiable at (g t)" using der by auto
-        qed
+        by (metis C1_differentiable_on_eq DiffD1 der g_diff imageI path_image_def that 
+                  vector_derivative_chain_at_general)
     qed
   ultimately have "f \<circ> g C1_differentiable_on {0..1} - S"
     using C1_differentiable_on_eq by blast
@@ -379,10 +364,8 @@
     by (auto intro!: continuous_intros finite_vimageI [where h = "(\<lambda>x. 2*x - 1)"] inj_onI
              simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths)
   ultimately show ?thesis
-    apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def)
-    apply (rule piecewise_C1_differentiable_cases)
-    apply (auto simp: o_def)
-    done
+    unfolding valid_path_def continuous_on_joinpaths joinpaths_def
+    by (intro piecewise_C1_differentiable_cases) (auto simp: o_def)
 qed
 
 lemma valid_path_join_D1:
@@ -406,31 +389,24 @@
   assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
     shows "valid_path(shiftpath a g)"
   using assms
-  apply (auto simp: valid_path_def shiftpath_alt_def)
-  apply (rule piecewise_C1_differentiable_cases)
-  apply (auto simp: algebra_simps)
-  apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one])
-  apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
-  apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps])
-  apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
+  unfolding valid_path_def shiftpath_alt_def
+  apply (intro piecewise_C1_differentiable_cases)
+      apply (simp_all add: add.commute)
+    apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one])
+    apply (force simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
+   apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps])
+   apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
   done
 
 lemma vector_derivative_linepath_within:
     "x \<in> {0..1} \<Longrightarrow> vector_derivative (linepath a b) (at x within {0..1}) = b - a"
-  apply (rule vector_derivative_within_cbox [of 0 "1::real", simplified])
-  apply (auto simp: has_vector_derivative_linepath_within)
-  done
+  by (simp add: has_vector_derivative_linepath_within vector_derivative_at_within_ivl)
 
 lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a"
   by (simp add: has_vector_derivative_linepath_within vector_derivative_at)
 
 lemma valid_path_linepath [iff]: "valid_path (linepath a b)"
-  apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath)
-  apply (rule_tac x="{}" in exI)
-  apply (simp add: differentiable_on_def differentiable_def)
-  using has_vector_derivative_def has_vector_derivative_linepath_within
-  apply (fastforce simp add: continuous_on_eq_continuous_within)
-  done
+  using C1_differentiable_on_eq piecewise_C1_differentiable_on_def valid_path_def by fastforce
 
 lemma valid_path_subpath:
   fixes g :: "real \<Rightarrow> 'a :: real_normed_vector"
@@ -443,15 +419,19 @@
     by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise)
 next
   case False
-  have "(g \<circ> (\<lambda>x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}"
-    apply (rule piecewise_C1_differentiable_compose)
-    apply (simp add: C1_differentiable_imp_piecewise)
-     apply (simp add: image_affinity_atLeastAtMost)
-    using assms False
-    apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset)
-    apply (subst Int_commute)
-    apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI)
-    done
+  let ?f = "\<lambda>x. ((v-u) * x + u)"
+  have "(g \<circ> ?f) piecewise_C1_differentiable_on {0..1}"
+  proof (rule piecewise_C1_differentiable_compose)
+    show "?f piecewise_C1_differentiable_on {0..1}"
+      by (simp add: C1_differentiable_imp_piecewise)
+    have "g piecewise_C1_differentiable_on (if u \<le> v then {u..v} else {v..u})"
+      using assms piecewise_C1_differentiable_on_subset valid_path_def by force
+    then show "g piecewise_C1_differentiable_on ?f ` {0..1}"
+      by (simp add: image_affinity_atLeastAtMost split: if_split_asm)
+    show "\<And>x. finite ({0..1} \<inter> ?f -` {x})"
+      using False
+      by (simp add: Int_commute [of "{0..1}"] inj_on_def crossproduct_eq finite_vimage_IntI)
+  qed
   then show ?thesis
     by (auto simp: o_def valid_path_def subpath_def)
 qed
--- a/src/HOL/Analysis/Tagged_Division.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Analysis/Tagged_Division.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -14,99 +14,56 @@
     and "\<And>i. i \<in> S \<Longrightarrow> finite (T i)"
   shows "(\<Sum>i\<in>S. sum (x i) (T i)) = (\<Sum>(i, j)\<in>Sigma S T. x i j)"
   using assms
-proof induct
-  case empty
-  then show ?case
-    by simp
-next
-  case (insert a S)
-  show ?case
-    by (simp add: insert.hyps(1) insert.prems sum.Sigma)
-qed
+  by induction (auto simp: sum.Sigma)
 
 lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
   scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
   scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one
 
 
-subsection\<^marker>\<open>tag unimportant\<close> \<open>Sundries\<close>
-
-
-text\<open>A transitive relation is well-founded if all initial segments are finite.\<close>
-lemma wf_finite_segments:
-  assumes "irrefl r" and "trans r" and "\<And>x. finite {y. (y, x) \<in> r}"
-  shows "wf (r)"
-  apply (simp add: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms)
-  using acyclic_def assms irrefl_def trans_Restr by fastforce
-
-text\<open>For creating values between \<^term>\<open>u\<close> and \<^term>\<open>v\<close>.\<close>
-lemma scaling_mono:
-  fixes u::"'a::linordered_field"
-  assumes "u \<le> v" "0 \<le> r" "r \<le> s"
-    shows "u + r * (v - u) / s \<le> v"
-proof -
-  have "r/s \<le> 1" using assms
-    using divide_le_eq_1 by fastforce
-  then have "(r/s) * (v - u) \<le> 1 * (v - u)"
-    by (meson diff_ge_0_iff_ge mult_right_mono \<open>u \<le> v\<close>)
-  then show ?thesis
-    by (simp add: field_simps)
-qed
-
 subsection \<open>Some useful lemmas about intervals\<close>
 
 lemma interior_subset_union_intervals:
-  assumes "i = cbox a b"
-    and "j = cbox c d"
-    and "interior j \<noteq> {}"
+  fixes a b c d
+  defines "i \<equiv> cbox a b"
+  defines "j \<equiv> cbox c d"
+  assumes "interior j \<noteq> {}"
     and "i \<subseteq> j \<union> S"
     and "interior i \<inter> interior j = {}"
   shows "interior i \<subseteq> interior S"
-proof -
-  have "box a b \<inter> cbox c d = {}"
-     using Int_interval_mixed_eq_empty[of c d a b] assms
-     unfolding interior_cbox by auto
-  moreover
-  have "box a b \<subseteq> cbox c d \<union> S"
-    apply (rule order_trans,rule box_subset_cbox)
-    using assms by auto
-  ultimately
-  show ?thesis
-    unfolding assms interior_cbox
-      by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI)
-qed
+  by (smt (verit, del_insts) IntI Int_interval_mixed_eq_empty UnE assms empty_iff interior_cbox interior_maximal interior_subset open_interior subset_eq)
 
 lemma interior_Union_subset_cbox:
-  assumes "finite f"
-  assumes f: "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a b" "\<And>s. s \<in> f \<Longrightarrow> interior s \<subseteq> t"
-    and t: "closed t"
-  shows "interior (\<Union>f) \<subseteq> t"
+  assumes "finite \<F>"
+  assumes \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a b. S = cbox a b" "\<And>S. S \<in> \<F> \<Longrightarrow> interior S \<subseteq> T"
+    and T: "closed T"
+  shows "interior (\<Union>\<F>) \<subseteq> T"
 proof -
-  have [simp]: "s \<in> f \<Longrightarrow> closed s" for s
-    using f by auto
-  define E where "E = {s\<in>f. interior s = {}}"
-  then have "finite E" "E \<subseteq> {s\<in>f. interior s = {}}"
-    using \<open>finite f\<close> by auto
-  then have "interior (\<Union>f) = interior (\<Union>(f - E))"
+  have clo[simp]: "S \<in> \<F> \<Longrightarrow> closed S" for S
+    using \<F> by auto
+  define E where "E = {s\<in>\<F>. interior s = {}}"
+  then have "finite E" "E \<subseteq> {s\<in>\<F>. interior s = {}}"
+    using \<open>finite \<F>\<close> by auto
+  then have "interior (\<Union>\<F>) = interior (\<Union>(\<F> - E))"
   proof (induction E rule: finite_subset_induct')
     case (insert s f')
-    have "interior (\<Union>(f - insert s f') \<union> s) = interior (\<Union>(f - insert s f'))"
-      using insert.hyps \<open>finite f\<close> by (intro interior_closed_Un_empty_interior) auto
-    also have "\<Union>(f - insert s f') \<union> s = \<Union>(f - f')"
+    have "interior (\<Union>(\<F> - insert s f') \<union> s) = interior (\<Union>(\<F> - insert s f'))"
+      using insert.hyps \<open>finite \<F>\<close> by (intro interior_closed_Un_empty_interior) auto
+    also have "\<Union>(\<F> - insert s f') \<union> s = \<Union>(\<F> - f')"
       using insert.hyps by auto
     finally show ?case
       by (simp add: insert.IH)
   qed simp
-  also have "\<dots> \<subseteq> \<Union>(f - E)"
+  also have "\<dots> \<subseteq> \<Union>(\<F> - E)"
     by (rule interior_subset)
-  also have "\<dots> \<subseteq> t"
+  also have "\<dots> \<subseteq> T"
   proof (rule Union_least)
-    fix s assume "s \<in> f - E"
-    with f[of s] obtain a b where s: "s \<in> f" "s = cbox a b" "box a b \<noteq> {}"
+    fix s assume "s \<in> \<F> - E"
+    with \<F>[of s] obtain a b where s: "s \<in> \<F>" "s = cbox a b" "box a b \<noteq> {}"
       by (fastforce simp: E_def)
-    have "closure (interior s) \<subseteq> closure t"
-      by (intro closure_mono f \<open>s \<in> f\<close>)
-    with s \<open>closed t\<close> show "s \<subseteq> t"
+    have "closure (interior s) \<subseteq> closure T"
+      by (intro closure_mono \<F> \<open>s \<in> \<F>\<close>)
+    with s \<open>closed T\<close> show "s \<subseteq> T"
       by simp
   qed
   finally show ?thesis .
@@ -195,14 +152,12 @@
 definition\<^marker>\<open>tag important\<close> "gauge \<gamma> \<longleftrightarrow> (\<forall>x. x \<in> \<gamma> x \<and> open (\<gamma> x))"
 
 lemma gaugeI:
-  assumes "\<And>x. x \<in> \<gamma> x"
-    and "\<And>x. open (\<gamma> x)"
+  assumes "\<And>x. x \<in> \<gamma> x" and "\<And>x. open (\<gamma> x)"
   shows "gauge \<gamma>"
   using assms unfolding gauge_def by auto
 
 lemma gaugeD[dest]:
-  assumes "gauge \<gamma>"
-  shows "x \<in> \<gamma> x"
+  assumes "gauge \<gamma>" shows "x \<in> \<gamma> x"
     and "open (\<gamma> x)"
   using assms unfolding gauge_def by auto
 
@@ -213,7 +168,7 @@
   unfolding gauge_def by auto
 
 lemma gauge_trivial[intro!]: "gauge (\<lambda>x. ball x 1)"
-  by (rule gauge_ball) auto
+  by auto
 
 lemma gauge_Int[intro]: "gauge \<gamma>1 \<Longrightarrow> gauge \<gamma>2 \<Longrightarrow> gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)"
   unfolding gauge_def by auto
@@ -221,8 +176,7 @@
 lemma gauge_reflect:
   fixes \<gamma> :: "'a::euclidean_space \<Rightarrow> 'a set"
   shows "gauge \<gamma> \<Longrightarrow> gauge (\<lambda>x. uminus ` \<gamma> (- x))"
-  using equation_minus_iff
-  by (auto simp: gauge_def surj_def intro!: open_surjective_linear_image linear_uminus)
+  by (metis (mono_tags, lifting) gauge_def imageI open_negations minus_minus)
 
 lemma gauge_Inter:
   assumes "finite S"
@@ -233,7 +187,7 @@
     by auto
   show ?thesis
     unfolding gauge_def unfolding *
-    using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto
+    by (simp add: assms gaugeD open_INT)
 qed
 
 lemma gauge_existence_lemma:
@@ -290,18 +244,6 @@
   "s division_of cbox a (a::'a::euclidean_space) \<longleftrightarrow> s = {cbox a a}"
   (is "?l = ?r")
 proof
-  assume ?r
-  moreover
-  { fix K
-    assume "s = {{a}}" "K\<in>s"
-    then have "\<exists>x y. K = cbox x y"
-      apply (rule_tac x=a in exI)+
-      apply force
-      done
-  }
-  ultimately show ?l
-    unfolding division_of_def cbox_idem by auto
-next
   assume ?l
   have "x = {a}" if  "x \<in> s" for x
     by (metis \<open>s division_of cbox a a\<close> cbox_idem division_ofD(2) division_ofD(3) subset_singletonD that)
@@ -309,10 +251,10 @@
     using \<open>s division_of cbox a a\<close> by auto
   ultimately show ?r
     unfolding cbox_idem by auto
-qed
+qed (use cbox_idem in blast)
 
 lemma elementary_empty: obtains p where "p division_of {}"
-  unfolding division_of_trivial by auto
+  by simp
 
 lemma elementary_interval: obtains p where "p division_of (cbox a b)"
   by (metis division_of_trivial division_of_self)
@@ -334,31 +276,24 @@
     and "q \<subseteq> p"
   shows "q division_of (\<Union>q)"
 proof (rule division_ofI)
-  note * = division_ofD[OF assms(1)]
   show "finite q"
-    using "*"(1) assms(2) infinite_super by auto
-  {
-    fix k
-    assume "k \<in> q"
-    then have kp: "k \<in> p"
-      using assms(2) by auto
-    show "k \<subseteq> \<Union>q"
-      using \<open>k \<in> q\<close> by auto
-    show "\<exists>a b. k = cbox a b"
-      using *(4)[OF kp] by auto
-    show "k \<noteq> {}"
-      using *(3)[OF kp] by auto
-  }
+    using assms finite_subset by blast
+next
+  fix k
+  assume "k \<in> q"
+  show "k \<subseteq> \<Union>q"
+    using \<open>k \<in> q\<close> by auto
+  show "\<exists>a b. k = cbox a b" "k \<noteq> {}"
+    using assms \<open>k \<in> q\<close> by blast+
+next
   fix k1 k2
   assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2"
-  then have **: "k1 \<in> p" "k2 \<in> p" "k1 \<noteq> k2"
-    using assms(2) by auto
-  show "interior k1 \<inter> interior k2 = {}"
-    using *(5)[OF **] by auto
+  then show "interior k1 \<inter> interior k2 = {}"
+    using assms by blast
 qed auto
 
 lemma division_of_union_self[intro]: "p division_of s \<Longrightarrow> p division_of (\<Union>p)"
-  unfolding division_of_def by auto
+  by blast
 
 lemma division_inter:
   fixes s1 s2 :: "'a::euclidean_space set"
@@ -379,25 +314,20 @@
     ultimately show "finite ?A" by auto
     have *: "\<And>s. \<Union>{x\<in>s. x \<noteq> {}} = \<Union>s"
       by auto
-    show "\<Union>?A = s1 \<inter> s2"
+    show UA: "\<Union>?A = s1 \<inter> s2"
       unfolding * 
       using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto
     {
       fix k
-      assume "k \<in> ?A"
+      assume kA: "k \<in> ?A"
       then obtain k1 k2 where k: "k = k1 \<inter> k2" "k1 \<in> p1" "k2 \<in> p2" "k \<noteq> {}"
         by auto
       then show "k \<noteq> {}"
         by auto
       show "k \<subseteq> s1 \<inter> s2"
-        using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)]
-        unfolding k by auto
-      obtain a1 b1 where k1: "k1 = cbox a1 b1"
-        using division_ofD(4)[OF assms(1) k(2)] by blast
-      obtain a2 b2 where k2: "k2 = cbox a2 b2"
-        using division_ofD(4)[OF assms(2) k(3)] by blast
+        using UA kA by blast
       show "\<exists>a b. k = cbox a b"
-        unfolding k k1 k2 unfolding Int_interval by auto
+        using k by (metis (no_types, lifting) Int_interval assms division_ofD(4))
     }
     fix k1 k2
     assume "k1 \<in> ?A"
@@ -407,17 +337,9 @@
     then obtain x2 y2 where k2: "k2 = x2 \<inter> y2" "x2 \<in> p1" "y2 \<in> p2" "k2 \<noteq> {}"
       by auto
     assume "k1 \<noteq> k2"
-    then have th: "x1 \<noteq> x2 \<or> y1 \<noteq> y2"
-      unfolding k1 k2 by auto
-    have *: "interior x1 \<inter> interior x2 = {} \<or> interior y1 \<inter> interior y2 = {} \<Longrightarrow>
-      interior (x1 \<inter> y1) \<subseteq> interior x1 \<Longrightarrow> interior (x1 \<inter> y1) \<subseteq> interior y1 \<Longrightarrow>
-      interior (x2 \<inter> y2) \<subseteq> interior x2 \<Longrightarrow> interior (x2 \<inter> y2) \<subseteq> interior y2 \<Longrightarrow>
-      interior (x1 \<inter> y1) \<inter> interior (x2 \<inter> y2) = {}" by auto
-    show "interior k1 \<inter> interior k2 = {}"
+    then show "interior k1 \<inter> interior k2 = {}"
       unfolding k1 k2
-      apply (rule *)
-      using assms division_ofD(5) k1 k2(2) k2(3) th apply auto
-      done
+      using assms division_ofD(5) k1 k2 by auto
   qed
 qed
 
@@ -439,33 +361,20 @@
 
 lemma elementary_Int:
   fixes s t :: "'a::euclidean_space set"
-  assumes "p1 division_of s"
-    and "p2 division_of t"
+  assumes "p1 division_of s" and "p2 division_of t"
   shows "\<exists>p. p division_of (s \<inter> t)"
 using assms division_inter by blast
 
 lemma elementary_Inter:
-  assumes "finite f"
-    and "f \<noteq> {}"
-    and "\<forall>s\<in>f. \<exists>p. p division_of (s::('a::euclidean_space) set)"
-  shows "\<exists>p. p division_of (\<Inter>f)"
+  assumes "finite \<F>"
+    and "\<F> \<noteq> {}"
+    and "\<forall>s\<in>\<F>. \<exists>p. p division_of (s::('a::euclidean_space) set)"
+  shows "\<exists>p. p division_of (\<Inter>\<F>)"
   using assms
-proof (induct f rule: finite_induct)
-  case (insert x f)
-  show ?case
-  proof (cases "f = {}")
-    case True
-    then show ?thesis
-      unfolding True using insert by auto
-  next
-    case False
-    obtain p where "p division_of \<Inter>f"
-      using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] ..
-    moreover obtain px where "px division_of x"
-      using insert(5)[rule_format,OF insertI1] ..
-    ultimately show ?thesis
-      by (simp add: elementary_Int Inter_insert)
-  qed
+proof (induct \<F> rule: finite_induct)
+  case (insert x \<F>)
+  then show ?case
+    by (metis cInf_singleton complete_lattice_class.Inf_insert elementary_Int insert_iff)
 qed auto
 
 lemma division_disjoint_union:
@@ -476,40 +385,11 @@
 proof (rule division_ofI)
   note d1 = division_ofD[OF assms(1)]
   note d2 = division_ofD[OF assms(2)]
-  show "finite (p1 \<union> p2)"
-    using d1(1) d2(1) by auto
-  show "\<Union>(p1 \<union> p2) = s1 \<union> s2"
-    using d1(6) d2(6) by auto
-  {
-    fix k1 k2
-    assume as: "k1 \<in> p1 \<union> p2" "k2 \<in> p1 \<union> p2" "k1 \<noteq> k2"
-    moreover
-    let ?g="interior k1 \<inter> interior k2 = {}"
-    {
-      assume as: "k1\<in>p1" "k2\<in>p2"
-      have ?g
-        using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]]
-        using assms(3) by blast
-    }
-    moreover
-    {
-      assume as: "k1\<in>p2" "k2\<in>p1"
-      have ?g
-        using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]]
-        using assms(3) by blast
-    }
-    ultimately show ?g
-      using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto
-  }
-  fix k
-  assume k: "k \<in> p1 \<union> p2"
-  show "k \<subseteq> s1 \<union> s2"
-    using k d1(2) d2(2) by auto
-  show "k \<noteq> {}"
-    using k d1(3) d2(3) by auto
-  show "\<exists>a b. k = cbox a b"
-    using k d1(4) d2(4) by auto
-qed
+  fix k1 k2
+  assume "k1 \<in> p1 \<union> p2" "k2 \<in> p1 \<union> p2" "k1 \<noteq> k2"
+  with assms show "interior k1 \<inter> interior k2 = {}"
+    by (smt (verit, best) IntE Un_iff disjoint_iff_not_equal division_ofD(2) division_ofD(5) inf.orderE interior_Int)
+qed (use division_ofD assms in auto)
 
 lemma partial_division_extend_1:
   fixes a b c d :: "'a::euclidean_space"
@@ -533,71 +413,59 @@
     show "finite p"
       unfolding p_def by (auto intro!: finite_PiE)
     {
-      fix k
-      assume "k \<in> p"
-      then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
+      fix K
+      assume "K \<in> p"
+      then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "K = ?B f"
         by (auto simp: p_def)
-      then show "\<exists>a b. k = cbox a b"
-        by auto
-      have "k \<subseteq> cbox a b \<and> k \<noteq> {}"
-      proof (simp add: k box_eq_empty subset_box not_less, safe)
-        fix i :: 'a
-        assume i: "i \<in> Basis"
-        with f have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
-          by (auto simp: PiE_iff)
-        with i ord[of i]
-        show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
-          by auto
-      qed
-      then show "k \<noteq> {}" "k \<subseteq> cbox a b"
+      then show "\<exists>a b. K = cbox a b"
         by auto
       {
         fix l
         assume "l \<in> p"
         then obtain g where g: "g \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
           by (auto simp: p_def)
-        assume "l \<noteq> k"
+        assume "l \<noteq> K"
         have "\<exists>i\<in>Basis. f i \<noteq> g i"
-        proof (rule ccontr)
-          assume "\<not> ?thesis"
-          with f g have "f = g"
-            by (auto simp: PiE_iff extensional_def fun_eq_iff)
-          with \<open>l \<noteq> k\<close> show False
-            by (simp add: l k)
-        qed
+          using \<open>l \<noteq> K\<close> l k f g by auto
         then obtain i where *: "i \<in> Basis" "f i \<noteq> g i" ..
         then have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
-                  "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)"
+          "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)"
           using f g by (auto simp: PiE_iff)
-        with * ord[of i] show "interior l \<inter> interior k = {}"
+        with * ord[of i] show "interior l \<inter> interior K = {}"
           by (auto simp add: l k disjoint_interval intro!: bexI[of _ i])
       }
-      note \<open>k \<subseteq> cbox a b\<close>
+      have "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
+        if "i \<in> Basis" for i
+      proof -
+        have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
+          using that f by (auto simp: PiE_iff)
+        with that ord[of i]
+        show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
+          by auto
+      qed
+      then show "K \<noteq> {}" "K \<subseteq> cbox a b"
+        by (auto simp add: k box_eq_empty subset_box not_less)      
     }
     moreover
-    {
-      fix x assume x: "x \<in> cbox a b"
-      have "\<forall>i\<in>Basis. \<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
-      proof
-        fix i :: 'a
-        assume "i \<in> Basis"
-        with x ord[of i]
+    have "\<exists>k\<in>p. x \<in> k" if x: "x \<in> cbox a b" for x
+    proof -
+      have "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}" if "i \<in> Basis" for i
+      proof -
         have "(a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> c \<bullet> i) \<or> (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i) \<or>
             (d \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)"
+          using that x ord[of i]
           by (auto simp: cbox_def)
         then show "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
           by auto
       qed
       then obtain f where
         f: "\<forall>i\<in>Basis. x \<bullet> i \<in> {fst (f i) \<bullet> i..snd (f i) \<bullet> i} \<and> f i \<in> {(a, c), (c, d), (d, b)}"
-        unfolding bchoice_iff ..
-      moreover from f have "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}"
-        by auto
-      moreover from f have "x \<in> ?B (restrict f Basis)"
+        by metis 
+      moreover from f have "x \<in> ?B (restrict f Basis)" "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a,c), (c,d), (d,b)}"
         by (auto simp: mem_box)
-      ultimately have "\<exists>k\<in>p. x \<in> k"
+      ultimately show ?thesis
         unfolding p_def by blast
-    }
+    qed
     ultimately show "\<Union>p = cbox a b"
       by auto
   qed
@@ -608,14 +476,12 @@
   obtains q where "p \<subseteq> q" "q division_of cbox a (b::'a::euclidean_space)"
 proof (cases "p = {}")
   case True
-  obtain q where "q division_of (cbox a b)"
-    by (rule elementary_interval)
   then show ?thesis
-    using True that by blast
+    using elementary_interval that by auto
 next
   case False
   note p = division_ofD[OF assms(1)]
-  have div_cbox: "\<forall>k\<in>p. \<exists>q. q division_of cbox a b \<and> k \<in> q"
+  have "\<forall>k\<in>p. \<exists>q. q division_of cbox a b \<and> k \<in> q"
   proof
     fix k
     assume kp: "k \<in> p"
@@ -629,64 +495,55 @@
     then show "\<exists>q. q division_of cbox a b \<and> k \<in> q"
       unfolding k by auto
   qed
-  obtain q where q: "\<And>x. x \<in> p \<Longrightarrow> q x division_of cbox a b" "\<And>x. x \<in> p \<Longrightarrow> x \<in> q x"
-    using bchoice[OF div_cbox] by blast
+  then obtain q where q: "\<And>x. x \<in> p \<Longrightarrow> q x division_of cbox a b" "\<And>x. x \<in> p \<Longrightarrow> x \<in> q x"
+    by metis
   have "q x division_of \<Union>(q x)" if x: "x \<in> p" for x
-    apply (rule division_ofI)
-    using division_ofD[OF q(1)[OF x]] by auto
+    using q(1) x by blast
   then have di: "\<And>x. x \<in> p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})"
     by (meson Diff_subset division_of_subset)
-  have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
-    apply (rule elementary_Inter [OF finite_imageI[OF p(1)]])
-    apply (auto dest: di simp: False elementary_Inter [OF finite_imageI[OF p(1)]])
-    done
-  then obtain d where d: "d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)" ..
+  have "\<forall>s\<in>(\<lambda>i. \<Union> (q i - {i})) ` p. \<exists>d. d division_of s"
+    using di by blast
+  then obtain d where d: "d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
+    by (meson False elementary_Inter finite_imageI image_is_empty p(1))
   have "d \<union> p division_of cbox a b"
   proof -
     have te: "\<And>S f T. S \<noteq> {} \<Longrightarrow> \<forall>i\<in>S. f i \<union> i = T \<Longrightarrow> T = \<Inter>(f ` S) \<union> \<Union>S" by auto
     have cbox_eq: "cbox a b = \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p) \<union> \<Union>p"
     proof (rule te[OF False], clarify)
       fix i
-      assume i: "i \<in> p"
-      show "\<Union>(q i - {i}) \<union> i = cbox a b"
-        using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto
+      assume "i \<in> p"
+      then show "\<Union>(q i - {i}) \<union> i = cbox a b"
+        by (metis Un_commute complete_lattice_class.Sup_insert division_ofD(6) insert_Diff q)
     qed
-    { fix K
-      assume K: "K \<in> p"
-      note qk=division_ofD[OF q(1)[OF K]]
-      have *: "\<And>u T S. T \<inter> S = {} \<Longrightarrow> u \<subseteq> S \<Longrightarrow> u \<inter> T = {}"
+      have [simp]: "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<inter> interior K = {}" if K: "K \<in> p" for K
+      proof -
+      note qk = division_ofD[OF q(1)[OF K]]
+      have *: "\<And>U T S. T \<inter> S = {} \<Longrightarrow> U \<subseteq> S \<Longrightarrow> U \<inter> T = {}"
         by auto
-      have "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<inter> interior K = {}"
+      show ?thesis
       proof (rule *[OF Int_interior_Union_intervals])
         show "\<And>T. T\<in>q K - {K} \<Longrightarrow> interior K \<inter> interior T = {}"
-          using qk(5) using q(2)[OF K] by auto
+          using K q(2) qk(5) by auto
         show "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<subseteq> interior (\<Union>(q K - {K}))"
-          apply (rule interior_mono)+
-          using K by auto
-      qed (use qk in auto)} note [simp] = this
+          by (meson INF_lower K interior_mono)
+      qed (use qk in auto)
+    qed
     show "d \<union> p division_of (cbox a b)"
-      unfolding cbox_eq
-      apply (rule division_disjoint_union[OF d assms(1)])
-      apply (rule Int_interior_Union_intervals)
-      apply (rule p open_interior ballI)+
-      apply simp_all
-      done
+      by (simp add: Int_interior_Union_intervals assms(1) cbox_eq d division_disjoint_union p(1) p(4))
   qed
   then show ?thesis
     by (meson Un_upper2 that)
 qed
 
 lemma elementary_bounded[dest]:
-  fixes S :: "'a::euclidean_space set"
   shows "p division_of S \<Longrightarrow> bounded S"
   unfolding division_of_def by (metis bounded_Union bounded_cbox)
 
 lemma elementary_subset_cbox:
-  "p division_of S \<Longrightarrow> \<exists>a b. S \<subseteq> cbox a (b::'a::euclidean_space)"
+  "p division_of S \<Longrightarrow> \<exists>a b. S \<subseteq> cbox a b"
   by (meson bounded_subset_cbox_symmetric elementary_bounded)
 
 proposition division_union_intervals_exists:
-  fixes a b :: "'a::euclidean_space"
   assumes "cbox a b \<noteq> {}"
   obtains p where "(insert (cbox a b) p) division_of (cbox a b \<union> cbox c d)"
 proof (cases "cbox c d = {}")
@@ -707,26 +564,19 @@
     obtain p where pd: "p division_of cbox c d" and "cbox u v \<in> p"
       by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]])
     note p = this division_ofD[OF pd]
-    have "interior (cbox a b \<inter> \<Union>(p - {cbox u v})) = interior(cbox u v \<inter> \<Union>(p - {cbox u v}))"
-      apply (rule arg_cong[of _ _ interior])
-      using p(8) uv by auto
+    have "interior (cbox a b \<inter> \<Union>(p - {cbox u v})) = interior(cbox u v) \<inter> interior (\<Union>(p - {cbox u v}))"
+      by (metis Diff_subset Int_absorb1 Int_assoc Sup_subset_mono interior_Int p(8) uv)
     also have "\<dots> = {}"
-      unfolding interior_Int
-      apply (rule Int_interior_Union_intervals)
       using p(6) p(7)[OF p(2)] \<open>finite p\<close>
-      apply auto
-      done
-    finally have [simp]: "interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}" by simp
+      by (intro Int_interior_Union_intervals) auto
+    finally have disj: "interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}" 
+      by simp
     have cbe: "cbox a b \<union> cbox c d = cbox a b \<union> \<Union>(p - {cbox u v})"
       using p(8) unfolding uv[symmetric] by auto
     have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
-    proof -
-      have "{cbox a b} division_of cbox a b"
-        by (simp add: assms division_of_self)
-      then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
-        by (metis (no_types) Diff_subset \<open>interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}\<close> division_disjoint_union division_of_subset insert_is_Un p(1) p(8))
-    qed
-    with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe)
+      by (metis Diff_subset assms disj division_disjoint_union division_of_self division_of_subset insert_is_Un p(8) pd)
+    with that[of "p - {cbox u v}"] show ?thesis 
+      by (simp add: cbe)
   qed
 qed
 
@@ -735,7 +585,7 @@
     and "\<And>p. p \<in> f \<Longrightarrow> p division_of (\<Union>p)"
     and "\<And>k1 k2. k1 \<in> \<Union>f \<Longrightarrow> k2 \<in> \<Union>f \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
   shows "\<Union>f division_of \<Union>(\<Union>f)"
-  using assms  by (auto intro!: division_ofI)
+  using assms by (auto intro!: division_ofI)
 
 lemma elementary_union_interval:
   fixes a b :: "'a::euclidean_space"
@@ -756,21 +606,13 @@
   proof (cases "interior (cbox a b) = {}")
     case True
     show ?thesis
-      apply (rule that [of "insert (cbox a b) p", OF division_ofI])
-      using pdiv(1-4) True False apply auto
-       apply (metis IntI equals0D pdiv(5))
-      by (metis disjoint_iff_not_equal pdiv(5))
+      by (metis True assms division_disjoint_union elementary_interval inf_bot_left that)
   next
-    case False  
+    case nonempty: False  
     have "\<forall>K\<in>p. \<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> K)"
-    proof
-      fix K
-      assume kp: "K \<in> p"
-      from pdiv(4)[OF kp] obtain c d where "K = cbox c d" by blast
-      then show "\<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> K)"
-        by (meson \<open>cbox a b \<noteq> {}\<close> division_union_intervals_exists)
-    qed
-    from bchoice[OF this] obtain q where "\<forall>x\<in>p. insert (cbox a b) (q x) division_of (cbox a b) \<union> x" ..
+      by (metis \<open>cbox a b \<noteq> {}\<close> division_union_intervals_exists pdiv(4))
+    then obtain q where "\<forall>x\<in>p. insert (cbox a b) (q x) division_of (cbox a b) \<union> x" 
+      by metis
     note q = division_ofD[OF this[rule_format]]
     let ?D = "\<Union>{insert (cbox a b) (q K) | K. K \<in> p}"
     show thesis
@@ -814,11 +656,11 @@
           by (metis \<open>interior (cbox a b) \<noteq> {}\<close> K q(2) x interior_subset_union_intervals)
         moreover
         obtain c d where c_d: "K' = cbox c d"
-          using q(4)[OF x'(2,1)] by blast
+          using q(4) x'(1) x'(2) by presburger
         have "interior K' \<inter> interior (cbox a b) = {}"
           using as'(2) q(5) x' by blast
         then have "interior K' \<subseteq> interior x'"
-          by (metis \<open>interior (cbox a b) \<noteq> {}\<close> c_d interior_subset_union_intervals q(2) x'(1) x'(2))
+          by (metis nonempty c_d interior_subset_union_intervals q(2) x')
         moreover have "interior x \<inter> interior x' = {}"
           by (meson False assms division_ofD(5) x'(2) x(2))
         ultimately show ?thesis
@@ -831,39 +673,34 @@
 
 
 lemma elementary_unions_intervals:
-  assumes fin: "finite f"
-    and "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a (b::'a::euclidean_space)"
-  obtains p where "p division_of (\<Union>f)"
+  assumes fin: "finite \<F>"
+    and "\<And>s. s \<in> \<F> \<Longrightarrow> \<exists>a b. s = cbox a (b::'a::euclidean_space)"
+  obtains p where "p division_of (\<Union>\<F>)"
 proof -
-  have "\<exists>p. p division_of (\<Union>f)"
-  proof (induct_tac f rule:finite_subset_induct)
+  have "\<exists>p. p division_of (\<Union>\<F>)"
+  proof (induct_tac \<F> rule:finite_subset_induct)
     show "\<exists>p. p division_of \<Union>{}" using elementary_empty by auto
   next
     fix x F
-    assume as: "finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f"
-    from this(3) obtain p where p: "p division_of \<Union>F" ..
-    from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast
-    have *: "\<Union>F = \<Union>p"
-      using division_ofD[OF p] by auto
-    show "\<exists>p. p division_of \<Union>(insert x F)"
-      using elementary_union_interval[OF p[unfolded *], of a b]
-      unfolding Union_insert x * by metis
-  qed (insert assms, auto)
+    assume as: "finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>\<F>"
+    then obtain a b where x: "x = cbox a b"
+      by (meson assms(2)) 
+    then show "\<exists>p. p division_of \<Union>(insert x F)"
+      using elementary_union_interval by (metis Union_insert as(3) division_ofD(6))
+  qed (use assms in auto)
   then show ?thesis
     using that by auto
 qed
 
 lemma elementary_union:
-  fixes S T :: "'a::euclidean_space set"
   assumes "ps division_of S" "pt division_of T"
   obtains p where "p division_of (S \<union> T)"
 proof -
-  have *: "S \<union> T = \<Union>ps \<union> \<Union>pt"
+  have "S \<union> T = \<Union>ps \<union> \<Union>pt"
     using assms unfolding division_of_def by auto
+  with elementary_unions_intervals[of "ps \<union> pt"] assms 
   show ?thesis
-    apply (rule elementary_unions_intervals[of "ps \<union> pt"])
-    using assms apply auto
-    by (simp add: * that)
+    by (metis Un_iff Union_Un_distrib division_of_def finite_Un that)
 qed
 
 lemma partial_division_extend:
@@ -881,10 +718,7 @@
     by (metis ab dual_order.trans partial_division_extend_interval divp(6))
   note r1 = this division_ofD[OF this(2)]
   obtain p' where "p' division_of \<Union>(r1 - p)"
-    apply (rule elementary_unions_intervals[of "r1 - p"])
-    using r1(3,6)
-      apply auto
-    done
+    by (metis Diff_subset division_of_subset r1(2) r1(8))
   then obtain r2 where r2: "r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)"
     by (metis assms(2) divq(6) elementary_Int)
   {
@@ -894,11 +728,7 @@
       unfolding r1 using ab
       by (meson division_contains r1(2) subsetCE)
     moreover have "R \<notin> p"
-    proof
-      assume "R \<in> p"
-      then have "x \<in> S" using divp(2) r by auto
-      then show False using x by auto
-    qed
+      using divp(6) r(2) x(2) by blast
     ultimately have "x\<in>\<Union>(r1 - p)" by auto
   }
   then have Teq: "T = \<Union>p \<union> (\<Union>(r1 - p) \<inter> \<Union>q)"
@@ -923,14 +753,11 @@
   then have div: "p \<union> r2 division_of \<Union>p \<union> \<Union>(r1 - p) \<inter> \<Union>q"
     by (simp add: assms(1) division_disjoint_union divp(6) r2)
   show ?thesis
-    apply (rule that[of "p \<union> r2"])
-     apply (auto simp: div Teq)
-    done
+    by (metis Teq div sup_ge1 that)
 qed
 
 
 lemma division_split:
-  fixes a :: "'a::euclidean_space"
   assumes "p division_of (cbox a b)"
     and k: "k\<in>Basis"
   shows "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} division_of(cbox a b \<inter> {x. x\<bullet>k \<le> c})"
@@ -954,11 +781,10 @@
       using l p(2) uv by force
     show  "K \<noteq> {}"
       by (simp add: l)
-    show  "\<exists>a b. K = cbox a b"
-      apply (simp add: l uv p(2-3)[OF l(2)])
-      apply (subst interval_split[OF k])
-      apply (auto intro: order.trans)
-      done
+    have "\<exists>a b. cbox u v \<inter> {x. x \<bullet> k \<le> c} = cbox a b"
+      unfolding interval_split[OF k] by (auto intro: order.trans)
+    then show  "\<exists>a b. K = cbox a b"
+      by (simp add: l(1) uv)
     fix K'
     assume "K' \<in> ?p1"
     then obtain l' where l': "K' = l' \<inter> {x. x \<bullet> k \<le> c}" "l' \<in> p" "l' \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
@@ -978,11 +804,10 @@
       using l p(2) uv by force
     show  "K \<noteq> {}"
       by (simp add: l)
-    show  "\<exists>a b. K = cbox a b"
-      apply (simp add: l uv p(2-3)[OF l(2)])
-      apply (subst interval_split[OF k])
-      apply (auto intro: order.trans)
-      done
+    have "\<exists>a b. cbox u v \<inter> {x. c \<le> x \<bullet> k} = cbox a b"
+      unfolding interval_split[OF k] by (auto intro: order.trans)
+    then show  "\<exists>a b. K = cbox a b"
+      by (simp add: l(1) uv)
     fix K'
     assume "K' \<in> ?p2"
     then obtain l' where l': "K' = l' \<inter> {x. c \<le> x \<bullet> k}" "l' \<in> p" "l' \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
@@ -1037,9 +862,9 @@
     and "(\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
   shows "s tagged_division_of i"
   unfolding tagged_division_of
-  using assms  by fastforce
+  using assms by fastforce
 
-lemma tagged_division_ofD[dest]:  (*FIXME USE A LOCALE*)
+lemma tagged_division_ofD[dest]: 
   assumes "s tagged_division_of i"
   shows "finite s"
     and "\<And>x K. (x,K) \<in> s \<Longrightarrow> x \<in> K"
@@ -1057,18 +882,14 @@
   note assm = tagged_division_ofD[OF assms]
   show "\<Union>(snd ` s) = i" "finite (snd ` s)"
     using assm by auto
-  fix k
-  assume k: "k \<in> snd ` s"
-  then obtain xk where xk: "(xk, k) \<in> s"
-    by auto
-  then show "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = cbox a b"
+  fix K
+  assume k: "K \<in> snd ` s"
+  then show "K \<subseteq> i" "K \<noteq> {}" "\<exists>a b. K = cbox a b"
     using assm by fastforce+
-  fix k'
-  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
-  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
-    by auto
-  then show "interior k \<inter> interior k' = {}"
-    using assm(5) k'(2) xk by blast
+  fix K'
+  assume k': "K' \<in> snd ` s" "K \<noteq> K'"
+  then show "interior K \<inter> interior K' = {}"
+    by (metis (no_types, lifting) assms imageE k prod.collapse tagged_division_ofD(5))
 qed
 
 lemma partial_division_of_tagged_division:
@@ -1078,23 +899,18 @@
   note assm = tagged_partial_division_ofD[OF assms]
   show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)"
     using assm by auto
-  fix k
-  assume k: "k \<in> snd ` s"
-  then obtain xk where xk: "(xk, k) \<in> s"
-    by auto
-  then show "k \<noteq> {}" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>(snd ` s)"
-    using assm by auto
-  fix k'
-  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
-  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
-    by auto
-  then show "interior k \<inter> interior k' = {}"
-    using assm(5) k'(2) xk by auto
+  fix K
+  assume K: "K \<in> snd ` s"
+  then show "K \<noteq> {}" "\<exists>a b. K = cbox a b" "K \<subseteq> \<Union>(snd ` s)"
+    using assm by fastforce+
+  fix K'
+  assume "K' \<in> snd ` s" "K \<noteq> K'"
+  then show "interior K \<inter> interior K' = {}"
+    using assm(5) K by force
 qed
 
 lemma tagged_partial_division_subset:
-  assumes "s tagged_partial_division_of i"
-    and "t \<subseteq> s"
+  assumes "s tagged_partial_division_of i" and "t \<subseteq> s"
   shows "t tagged_partial_division_of i"
   using assms finite_subset[OF assms(2)]
   unfolding tagged_partial_division_of_def
@@ -1116,8 +932,7 @@
   by (rule tagged_division_ofI) auto
 
 lemma tagged_division_of_self_real: "x \<in> {a .. b::real} \<Longrightarrow> {(x,{a .. b})} tagged_division_of {a .. b}"
-  unfolding box_real[symmetric]
-  by (rule tagged_division_of_self)
+  by (metis box_real(2) tagged_division_of_self)
 
 lemma tagged_division_Un:
   assumes "p1 tagged_division_of s1"
@@ -1131,20 +946,16 @@
     using p1(1) p2(1) by auto
   show "\<Union>{k. \<exists>x. (x, k) \<in> p1 \<union> p2} = s1 \<union> s2"
     using p1(6) p2(6) by blast
-  fix x k
-  assume xk: "(x, k) \<in> p1 \<union> p2"
-  show "x \<in> k" "\<exists>a b. k = cbox a b"
-    using xk p1(2,4) p2(2,4) by auto
-  show "k \<subseteq> s1 \<union> s2"
-    using xk p1(3) p2(3) by blast
-  fix x' k'
-  assume xk': "(x', k') \<in> p1 \<union> p2" "(x, k) \<noteq> (x', k')"
+  fix x K
+  assume xK: "(x, K) \<in> p1 \<union> p2"
+  show "x \<in> K" "\<exists>a b. K = cbox a b" "K \<subseteq> s1 \<union> s2"
+    using xK p1 p2 by auto
+  fix x' K'
+  assume xk': "(x', K') \<in> p1 \<union> p2" "(x, K) \<noteq> (x', K')"
   have *: "\<And>a b. a \<subseteq> s1 \<Longrightarrow> b \<subseteq> s2 \<Longrightarrow> interior a \<inter> interior b = {}"
     using assms(3) interior_mono by blast
-  show "interior k \<inter> interior k' = {}"
-    apply (cases "(x, k) \<in> p1")
-    apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2))
-    by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2))
+  with assms show "interior K \<inter> interior K' = {}"
+    by (metis Un_iff inf_commute p1(3) p2(3) tagged_division_ofD(5) xK xk')
 qed
 
 lemma tagged_division_Union:
@@ -1186,21 +997,16 @@
 lemma tagged_partial_division_of_Union_self:
   assumes "p tagged_partial_division_of s"
   shows "p tagged_division_of (\<Union>(snd ` p))"
-  apply (rule tagged_division_ofI)
   using tagged_partial_division_ofD[OF assms]
-  apply auto
-  done
+  by (intro tagged_division_ofI) auto
 
 lemma tagged_division_of_union_self:
   assumes "p tagged_division_of s"
   shows "p tagged_division_of (\<Union>(snd ` p))"
-  apply (rule tagged_division_ofI)
   using tagged_division_ofD[OF assms]
-  apply auto
-  done
+  by (intro tagged_division_ofI) auto
 
 lemma tagged_division_Un_interval:
-  fixes a :: "'a::euclidean_space"
   assumes "p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> (c::real)})"
     and "p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
     and k: "k \<in> Basis"
@@ -1208,13 +1014,10 @@
 proof -
   have *: "cbox a b = (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<union> (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
     by auto
-  show ?thesis
-    apply (subst *)
-    apply (rule tagged_division_Un[OF assms(1-2)])
-    unfolding interval_split[OF k] interior_cbox
-    using k
-    apply (auto simp add: box_def elim!: ballE[where x=k])
-    done
+  have "interior (cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<inter> interior (cbox a b \<inter> {x. c \<le> x \<bullet> k}) = {}"
+    using k by (force simp: interval_split[OF k] box_def)
+  with assms show ?thesis
+    by (metis "*" tagged_division_Un)
 qed
 
 lemma tagged_division_Un_interval_real:
@@ -1223,35 +1026,23 @@
     and "p2 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<ge> c})"
     and k: "k \<in> Basis"
   shows "(p1 \<union> p2) tagged_division_of {a .. b}"
-  using assms
-  unfolding box_real[symmetric]
-  by (rule tagged_division_Un_interval)
+  using assms by (metis box_real(2) tagged_division_Un_interval)
 
 lemma tagged_division_split_left_inj:
   assumes d: "d tagged_division_of i"
-  and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
-  and "K1 \<noteq> K2"
-  and eq: "K1 \<inter> {x. x \<bullet> k \<le> c} = K2 \<inter> {x. x \<bullet> k \<le> c}"
-    shows "interior (K1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
-proof -
-  have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
-    using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
-  then show ?thesis
-    using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
-qed
+    and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
+    and "K1 \<noteq> K2"
+    and eq: "K1 \<inter> {x. x \<bullet> k \<le> c} = K2 \<inter> {x. x \<bullet> k \<le> c}"
+  shows "interior (K1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
+  by (smt (verit) Int_Un_eq(1) Un_Int_distrib interior_Int prod.inject sup_bot.right_neutral tagged_division_ofD(5) assms)
 
 lemma tagged_division_split_right_inj:
   assumes d: "d tagged_division_of i"
-  and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
-  and "K1 \<noteq> K2"
-  and eq: "K1 \<inter> {x. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}"
-    shows "interior (K1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
-proof -
-  have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
-    using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
-  then show ?thesis
-    using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
-qed
+    and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
+    and "K1 \<noteq> K2"
+    and eq: "K1 \<inter> {x. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}"
+  shows "interior (K1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
+  by (smt (verit) Int_Un_eq(1) Un_Int_distrib interior_Int prod.inject sup_bot.right_neutral tagged_division_ofD(5) assms)
 
 lemma (in comm_monoid_set) over_tagged_division_lemma:
   assumes "p tagged_division_of i"
@@ -1272,14 +1063,10 @@
       using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto
     have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y"
       using \<open>x \<in> p\<close> \<open>x \<noteq> y\<close> \<open>snd x = snd y\<close> [symmetric] by auto
-    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}"
-      by (intro assm(5)[of "fst x" _ "fst y"]) auto
-    then have "box a b = {}"
-      unfolding \<open>snd x = snd y\<close>[symmetric] ab by auto
-    then have "d (cbox a b) = \<^bold>1"
-      using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto
+    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "box a b = {}"
+      by (metis \<open>snd x = snd y\<close> ab assm(5) inf.idem interior_cbox prod.collapse)
     then show "d (snd x) = \<^bold>1"
-      unfolding ab by auto
+      by (simp add: ab assms(2))
   qed
 qed
 
@@ -1311,11 +1098,8 @@
     by standard (auto simp: lift_option_def ac_simps split: bind_split)
 qed
 
-lemma comm_monoid_and: "comm_monoid HOL.conj True"
-  by standard auto
-
 lemma comm_monoid_set_and: "comm_monoid_set HOL.conj True"
-  by (rule comm_monoid_set.intro) (fact comm_monoid_and)
+  by (simp add: comm_monoid_set.intro conj.comm_monoid_axioms)
 
 
 paragraph \<open>Misc\<close>
@@ -1323,9 +1107,7 @@
 lemma interval_real_split:
   "{a .. b::real} \<inter> {x. x \<le> c} = {a .. min b c}"
   "{a .. b} \<inter> {x. c \<le> x} = {max a c .. b}"
-  apply (metis Int_atLeastAtMostL1 atMost_def)
-  apply (metis Int_atLeastAtMostL2 atLeast_def)
-  done
+  by auto
 
 lemma bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)"
   by (meson zero_less_one)
@@ -1375,16 +1157,11 @@
   proof -
     obtain u v where l: "l = cbox u v"
       using \<open>l \<in> d\<close> assms(1) by blast
-    have *: "\<forall>i\<in>Basis. u \<bullet> i \<le> (\<Sum>i\<in>Basis. (if i = k then min (v \<bullet> k) c else v \<bullet> i) *\<^sub>R i) \<bullet> i"
-      using that(6) unfolding l interval_split[OF k] box_ne_empty that .
-    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
+    have "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
       using l using that(6) unfolding box_ne_empty[symmetric] by auto
-    show ?thesis
-      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
-      using that(1-3,5) \<open>x \<in> Basis\<close>
-      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that
-      apply (auto split: if_split_asm)
-      done
+    then show ?thesis
+      using that \<open>x \<in> Basis\<close> unfolding l interval_split[OF k] 
+      by (force split: if_split_asm)
   qed
   moreover have "\<And>x y. \<lbrakk>y < (if x = k then c else b \<bullet> x)\<rbrakk> \<Longrightarrow> y < b \<bullet> x"
     using \<open>c < b\<bullet>k\<close> by (auto split: if_split_asm)
@@ -1403,16 +1180,11 @@
   proof -
     obtain u v where l: "l = cbox u v"
       using \<open>l \<in> d\<close> assm(4) by blast
-    have *: "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (u \<bullet> k) c else u \<bullet> i) *\<^sub>R i) \<bullet> i \<le> v \<bullet> i"
-      using that(6) unfolding l interval_split[OF k] box_ne_empty that .
-    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
+    have "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
       using l using that(6) unfolding box_ne_empty[symmetric] by auto
-    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
-      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
-      using that(1-3,5) \<open>x \<in> Basis\<close>
-      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that
-      apply (auto split: if_split_asm)
-      done
+    then show "\<exists>i\<in>d. interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
+      using that \<open>x \<in> Basis\<close> unfolding l interval_split[OF k] 
+      by (force split: if_split_asm)
   qed
   ultimately show ?t2
     unfolding division_points_def interval_split[OF k, of a b]
@@ -1520,8 +1292,6 @@
   shows "(\<lambda>l. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e}) ` {l\<in>p. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e} \<noteq> {}}
          division_of  (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e})"
 proof -
-  have *: "\<And>x c. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
-    by auto
   have **: "\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'"
     by auto
   note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]]
@@ -1533,7 +1303,8 @@
       apply (rule equalityI)
       apply blast
       apply clarsimp
-      apply (rule_tac x="xa \<inter> {x. c + e \<ge> x \<bullet> k}" in exI)
+      apply (rename_tac S)
+      apply (rule_tac x="S \<inter> {x. c + e \<ge> x \<bullet> k}" in exI)
       apply auto
       done
     by (simp add: interval_split k interval_doublesplit)
@@ -1547,16 +1318,8 @@
     and Basis_imp: "\<And>a b c k. k \<in> Basis \<Longrightarrow> g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
 begin
 
-lemma empty [simp]:
-  "g {} = \<^bold>1"
-proof -
-  have *: "cbox One (-One) = ({}::'b set)"
-    by (auto simp: box_eq_empty inner_sum_left inner_Basis sum.If_cases ex_in_conv)
-  moreover have "box One (-One) = ({}::'b set)"
-    using box_subset_cbox[of One "-One"] by (auto simp: *)
-  ultimately show ?thesis
-    using box_empty_imp [of One "-One"] by simp
-qed
+lemma empty [simp]: "g {} = \<^bold>1"
+  by (metis box_empty_imp box_subset_cbox empty_as_interval subset_empty)
        
 lemma division:
   "F g d = g (cbox a b)" if "d division_of (cbox a b)"
@@ -1566,21 +1329,22 @@
   using that proof (induction C arbitrary: a b d rule: less_induct)
     case (less a b d)
     show ?case
-    proof cases
-      assume "box a b = {}"
+    proof (cases "box a b = {}")
+      case True
       { fix k assume "k\<in>d"
         then obtain a' b' where k: "k = cbox a' b'"
           using division_ofD(4)[OF less.prems] by blast
-        with \<open>k\<in>d\<close> division_ofD(2)[OF less.prems] have "cbox a' b' \<subseteq> cbox a b"
-          by auto
+        then have "cbox a' b' \<subseteq> cbox a b"
+          using \<open>k \<in> d\<close> less.prems by blast
         then have "box a' b' \<subseteq> box a b"
           unfolding subset_box by auto
         then have "g k = \<^bold>1"
-          using box_empty_imp [of a' b'] k by (simp add: \<open>box a b = {}\<close>) }
-      then show "box a b = {} \<Longrightarrow> F g d = g (cbox a b)"
+          using box_empty_imp [of a' b'] k by (simp add: True)
+      }
+      with True show "F g d = g (cbox a b)"
         by (auto intro!: neutral simp: box_empty_imp)
     next
-      assume "box a b \<noteq> {}"
+      case False
       then have ab: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" and ab': "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
         by (auto simp: box_ne_empty)
       show "F g d = g (cbox a b)"
@@ -1600,11 +1364,10 @@
           note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]]
           moreover
           have "a\<bullet>j \<le> u\<bullet>j" "v\<bullet>j \<le> b\<bullet>j"
-            using division_ofD(2,2,3)[OF \<open>d division_of cbox a b\<close> as]
-            apply (metis j subset_box(1) uv(1))
-            by (metis \<open>cbox u v \<subseteq> cbox a b\<close> j subset_box(1) uv(1))
+            by (meson as division_ofD(2) j less.prems subset_box(1) uv(1))+
           ultimately have "u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j"
-            unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force }
+            using uv(2) by force
+        }
         then have d': "\<forall>i\<in>d. \<exists>u v. i = cbox u v \<and>
           (\<forall>j\<in>Basis. u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j)"
           unfolding forall_in_division[OF less.prems] by blast
@@ -1624,9 +1387,9 @@
           proof safe
             fix j :: 'b
             assume j: "j \<in> Basis"
-            note i(2)[unfolded uv mem_box,rule_format,of j]
+            note i(2)[unfolded uv mem_box]
             then show "u \<bullet> j = a \<bullet> j" and "v \<bullet> j = b \<bullet> j"
-              using uv(2)[rule_format,of j] j by (auto simp: inner_simps)
+              by (smt (verit) False box_midpoint j mem_box(1) uv(2))+
           qed
           then have "i = cbox a b" using uv by auto
           then show ?thesis using i by auto
@@ -1650,17 +1413,11 @@
             unfolding euclidean_eq_iff[where 'a='b] by auto
           then have "u\<bullet>j = v\<bullet>j"
             using uv(2)[rule_format,OF j] by auto
-          then have "box u v = {}"
-            using j unfolding box_eq_empty by (auto intro!: bexI[of _ j])
           then show "g x = \<^bold>1"
-            unfolding uv(1) by (rule box_empty_imp)
+            by (smt (verit) box_empty_imp box_eq_empty(1) j uv(1))
         qed
         then show "F g d = g (cbox a b)"
-          using division_ofD[OF less.prems]
-          apply (subst deq)
-          apply (subst insert)
-          apply auto
-          done
+          by (metis deq division_of_finite insert_Diff_single insert_remove less.prems right_neutral)
       next
         case False
         then have "\<exists>x. x \<in> division_points (cbox a b) d"
@@ -1748,12 +1505,7 @@
 proposition tagged_division:
   assumes "d tagged_division_of (cbox a b)"
   shows "F (\<lambda>(_, l). g l) d = g (cbox a b)"
-proof -
-  have "F (\<lambda>(_, k). g k) d = F g (snd ` d)"
-    using assms box_empty_imp by (rule over_tagged_division_lemma)
-  then show ?thesis
-    unfolding assms [THEN division_of_tagged_division, THEN division] .
-  qed
+  by (metis assms box_empty_imp division division_of_tagged_division over_tagged_division_lemma)
 
 end
 
@@ -1778,14 +1530,14 @@
       from that have [simp]: "k = 1"
         by simp
       from neutral [of 0 1] neutral [of a a for a] coalesce_less
-  have [simp]: "g {} = \<^bold>1" "\<And>a. g {a} = \<^bold>1"
-    "\<And>a b c. a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    by auto
+      have [simp]: "g {} = \<^bold>1" "\<And>a. g {a} = \<^bold>1"
+        "\<And>a b c. a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
+        by auto
       have "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
-    by (auto simp: min_def max_def le_less)
+        by (auto simp: min_def max_def le_less)
       then show "g (cbox a b) = g (cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. c \<le> x \<bullet> k})"
         by (simp add: atMost_def [symmetric] atLeast_def [symmetric])
-qed
+    qed
   qed
   show "box = (greaterThanLessThan :: real \<Rightarrow> _)"
     and "cbox = (atLeastAtMost :: real \<Rightarrow> _)"
@@ -1795,17 +1547,7 @@
 
 lemma coalesce_less_eq:
   "g {a..c} \<^bold>* g {c..b} = g {a..b}" if "a \<le> c" "c \<le> b"
-  proof (cases "c = a \<or> c = b")
-    case False
-  with that have "a < c" "c < b"
-    by auto
-    then show ?thesis
-    by (rule coalesce_less)
-  next
-    case True
-  with that box_empty_imp [of a a] box_empty_imp [of b b] show ?thesis
-    by safe simp_all
-    qed
+  by (metis coalesce_less commute left_neutral less_eq_real_def neutral that)
 
 end
 
@@ -1819,22 +1561,22 @@
     show "g {a..b} = z" if "b \<le> a" for a b
       using that box_empty_imp by simp
     show "f (g {a..c}) (g {c..b}) = g {a..b}" if "a < c" "c < b" for a b c
-      using that
-    using Basis_imp [of 1 a b c]
+      using that Basis_imp [of 1 a b c]
       by (simp_all add: atMost_def [symmetric] atLeast_def [symmetric] max_def min_def)
-qed
+  qed
 qed
 
 
 subsection \<open>Special case of additivity we need for the FTC\<close>
 (*fix add explanation here *)
+
 lemma additive_tagged_division_1:
   fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
   assumes "a \<le> b"
     and "p tagged_division_of {a..b}"
-  shows "sum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
+  shows "sum (\<lambda>(x,K). f(Sup K) - f(Inf K)) p = f b - f a"
 proof -
-  let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
+  let ?f = "(\<lambda>K::real set. if K = {} then 0 else f(interval_upperbound K) - f(interval_lowerbound K))"
   interpret operative_real plus 0 ?f
     rewrites "comm_monoid_set.F (+) 0 = sum"
     by standard[1] (auto simp add: sum_def)
@@ -1843,12 +1585,7 @@
   have **: "cbox a b \<noteq> {}"
     using assms(1) by auto
   then have "f b - f a = (\<Sum>(x, l)\<in>p. if l = {} then 0 else f (interval_upperbound l) - f (interval_lowerbound l))"
-    proof -
-      have "(if cbox a b = {} then 0 else f (interval_upperbound (cbox a b)) - f (interval_lowerbound (cbox a b))) = f b - f a"
-        using assms by auto
-      then show ?thesis
-        using p_td assms by (simp add: tagged_division) 
-    qed 
+    using assms(2) tagged_division by force
   then show ?thesis
     using assms by (auto intro!: sum.cong)
 qed
@@ -1897,12 +1634,16 @@
   obtain pfn where pfn:
     "\<And>x. x \<in> I \<Longrightarrow> pfn x tagged_division_of x"
     "\<And>x. x \<in> I \<Longrightarrow> d fine pfn x"
-    using bchoice[OF assms(2)] by auto
+    using assms by metis
   show thesis
-    apply (rule_tac p="\<Union>(pfn ` I)" in that)
-    using assms(1) assms(3) assms(4) pfn(1) tagged_division_Union apply force
-    by (metis (mono_tags, lifting) fine_Union imageE pfn(2))
+  proof
+    show "\<Union> (pfn ` I) tagged_division_of i"
+      using assms pfn(1) tagged_division_Union by force
+    show "d fine \<Union> (pfn ` I)"
+      by (metis (mono_tags, lifting) fine_Union imageE pfn(2))
+  qed
 qed
+
 (* FIXME structure here, do not have one lemma in each subsection *)
 
 subsection\<^marker>\<open>tag unimportant\<close> \<open>The set we're concerned with must be closed\<close>
@@ -1910,9 +1651,11 @@
 lemma division_of_closed:
   fixes i :: "'n::euclidean_space set"
   shows "s division_of i \<Longrightarrow> closed i"
-  unfolding division_of_def by fastforce
+  by blast
 (* FIXME structure here, do not have one lemma in each subsection *)
+
 subsection \<open>General bisection principle for intervals; might be useful elsewhere\<close>
+
 (* FIXME  move? *)
 lemma interval_bisection_step:
   fixes type :: "'a::euclidean_space"
@@ -1952,11 +1695,10 @@
       assume "x \<in> ?A"
       then obtain c d
         where x:  "x = cbox c d"
-          "\<And>i. i \<in> Basis \<Longrightarrow>
-                        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = ?ab i \<or> c \<bullet> i = ?ab i \<and> d \<bullet> i = b \<bullet> i" 
+          "\<And>i. i \<in> Basis \<Longrightarrow> c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = ?ab i \<or> c \<bullet> i = ?ab i \<and> d \<bullet> i = b \<bullet> i" 
         by blast
       have "c = (\<Sum>i\<in>Basis. (if c \<bullet> i = a \<bullet> i then a \<bullet> i else ?ab i) *\<^sub>R i)"
-           "d = (\<Sum>i\<in>Basis. (if c \<bullet> i = a \<bullet> i then ?ab i else b \<bullet> i) *\<^sub>R i)"
+        "d = (\<Sum>i\<in>Basis. (if c \<bullet> i = a \<bullet> i then ?ab i else b \<bullet> i) *\<^sub>R i)"
         using x(2) ab by (fastforce simp add: euclidean_eq_iff[where 'a='a])+
       then show "x \<in> ?B"
         unfolding x by (rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in image_eqI) auto
@@ -1968,7 +1710,7 @@
     assume "S \<in> ?A"
     then obtain c d
       where s: "S = cbox c d"
-               "\<And>i. i \<in> Basis \<Longrightarrow> c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = ?ab i \<or> c \<bullet> i = ?ab i \<and> d \<bullet> i = b \<bullet> i"
+        "\<And>i. i \<in> Basis \<Longrightarrow> c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = ?ab i \<or> c \<bullet> i = ?ab i \<and> d \<bullet> i = b \<bullet> i"
       by blast
     show "P S"
       unfolding s using ab s(2) by (fastforce intro!: that)
@@ -1986,8 +1728,7 @@
     then obtain i where "c\<bullet>i \<noteq> e\<bullet>i \<or> d\<bullet>i \<noteq> f\<bullet>i" and i': "i \<in> Basis"
       unfolding euclidean_eq_iff[where 'a='a] by auto
     then have i: "c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i"
-      using s(2) t(2) apply fastforce
-      using t(2)[OF i'] \<open>c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i\<close> i' s(2) t(2) by fastforce
+      using s(2) t(2) by fastforce+
     have *: "\<And>s t. (\<And>a. a \<in> s \<Longrightarrow> a \<in> t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}"
       by auto
     show "interior S \<inter> interior T = {}"
@@ -1995,10 +1736,10 @@
     proof (rule *)
       fix x
       assume "x \<in> box c d" "x \<in> box e f"
-      then have x: "c\<bullet>i < d\<bullet>i" "e\<bullet>i < f\<bullet>i" "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i"
-        unfolding mem_box using i'  by force+
-      show False  using s(2)[OF i'] t(2)[OF i'] and i x  
-        by auto
+      then have "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i"
+        unfolding mem_box using i' by force+
+      with i i' show False
+        using s(2) t(2) by fastforce
     qed
   qed
   also have "\<Union>?A = cbox a b"
@@ -2024,7 +1765,7 @@
       by (force simp add: mem_box)
   qed
   finally show thesis
-      by (metis (no_types, lifting) assms(3) that)
+    by (metis (no_types, lifting) assms(3) that)
 qed
 
 lemma interval_bisection:
@@ -2066,9 +1807,10 @@
             snd (f x) \<bullet> i \<le> snd x \<bullet> i \<and>
             2 * (snd (f x) \<bullet> i - fst (f x) \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i)" by metis
   define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n
-  have [simp]: "A 0 = a" "B 0 = b" and ABRAW: "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and>
-    (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
-    2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
+  have [simp]: "A 0 = a" "B 0 = b" 
+    and ABRAW: "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and>
+                  (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
+                  2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
   proof -
     show "A 0 = a" "B 0 = b"
       unfolding ab_def by auto
@@ -2080,13 +1822,9 @@
         unfolding S using \<open>\<not> P (cbox a b)\<close> f by auto
     next
       case (Suc n)
-      show ?case
+      then show ?case
         unfolding S
-        apply (rule f[rule_format])
-        using Suc
-        unfolding S
-        apply auto
-        done
+        by (force intro!: f[rule_format])
     qed
   qed
   then have AB: "A(n)\<bullet>i \<le> A(Suc n)\<bullet>i" "A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i" 
@@ -2109,10 +1847,9 @@
       also have "\<dots> \<le> sum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
       proof (rule sum_mono)
         fix i :: 'a
-        assume i: "i \<in> Basis"
-        show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
-          using xy[unfolded mem_box,THEN bspec, OF i]
-          by (auto simp: inner_diff_left)
+        assume "i \<in> Basis"
+        with xy show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
+          by (smt (verit, best) inner_diff_left mem_box(2))
       qed
       also have "\<dots> \<le> sum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
         unfolding sum_divide_distrib
@@ -2136,23 +1873,21 @@
       finally show "dist x y < e" .
     qed
   qed
-  {
-    fix n m :: nat
-    assume "m \<le> n" then have "cbox (A n) (B n) \<subseteq> cbox (A m) (B m)"
-    proof (induction rule: inc_induct)
-      case (step i)
-      show ?case
-        using AB by (intro order_trans[OF step.IH] subset_box_imp) auto
-    qed simp
-  } note ABsubset = this
+  have ABsubset: "cbox (A n) (B n) \<subseteq> cbox (A m) (B m)" if "m \<le> n" for m n
+    using that
+  proof (induction rule: inc_induct)
+    case (step i) show ?case
+      by (smt (verit, ccfv_SIG) ABRAW in_mono mem_box(2) step.IH subsetI)
+  qed simp
   have "\<And>n. cbox (A n) (B n) \<noteq> {}"
     by (meson AB dual_order.trans interval_not_empty)
   then obtain x0 where x0: "\<And>n. x0 \<in> cbox (A n) (B n)"
     using decreasing_closed_nest [OF closed_cbox] ABsubset interv by blast
   show thesis
-  proof (rule that[rule_format, of x0])
-    show "x0\<in>cbox a b"
+  proof (intro that strip)
+    show "x0 \<in> cbox a b"
       using \<open>A 0 = a\<close> \<open>B 0 = b\<close> x0 by blast
+  next
     fix e :: real
     assume "e > 0"
     from interv[OF this] obtain n
@@ -2170,17 +1905,14 @@
     moreover have "cbox (A n) (B n) \<subseteq> cbox a b"
       using ABsubset \<open>A 0 = a\<close> \<open>B 0 = b\<close> by blast
     ultimately show "\<exists>c d. x0 \<in> cbox c d \<and> cbox c d \<subseteq> ball x0 e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
-      apply (rule_tac x="A n" in exI)
-      apply (rule_tac x="B n" in exI)
-      apply (auto simp: x0)
-      done
+      by (meson x0)
   qed
 qed
 
 
 subsection \<open>Cousin's lemma\<close>
 
-lemma fine_division_exists: (*FIXME rename(?) *)
+lemma fine_division_exists: 
   fixes a b :: "'a::euclidean_space"
   assumes "gauge g"
   obtains p where "p tagged_division_of (cbox a b)" "g fine p"
@@ -2199,23 +1931,21 @@
           cbox c d \<subseteq> ball x e \<and>
           cbox c d \<subseteq> (cbox a b) \<and>
           \<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
-    apply (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p", OF _ _ False])
-    apply (simp add: fine_def)
-    apply (metis tagged_division_Un fine_Un)
-    apply auto
-    done
+  proof (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> _ p", OF _ _ False])
+    show "\<exists>p. p tagged_division_of {} \<and> g fine p"
+      by (simp add: fine_def)
+  qed (meson tagged_division_Un fine_Un)+
   obtain e where e: "e > 0" "ball x e \<subseteq> g x"
-    using gaugeD[OF assms, of x] unfolding open_contains_ball by auto
-  from x(2)[OF e(1)]
-  obtain c d where c_d: "x \<in> cbox c d"
+    by (meson assms gauge_def openE)
+  then obtain c d where c_d: "x \<in> cbox c d"
                         "cbox c d \<subseteq> ball x e"
                         "cbox c d \<subseteq> cbox a b"
                         "\<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
-    by blast
+    by (meson x(2))
   have "g fine {(x, cbox c d)}"
     unfolding fine_def using e using c_d(2) by auto
   then show ?thesis
-    using tagged_division_of_self[OF c_d(1)] using c_d by auto
+    using tagged_division_of_self[OF c_d(1)] using c_d by simp
 qed
 
 lemma fine_division_exists_real:
@@ -2232,7 +1962,7 @@
     and "gauge d"
   obtains q where "q tagged_division_of (cbox a b)"
     and "d fine q"
-    and "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
+    and "\<forall>(x,K) \<in> p. K \<subseteq> d(x) \<longrightarrow> (x,K) \<in> q"
 proof -
   have p: "finite p" "p tagged_partial_division_of (cbox a b)"
     using ptag tagged_division_of_def by blast+
@@ -2250,48 +1980,48 @@
     obtain q1 where q1: "q1 tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> p}"
                 and "d fine q1"
                 and q1I: "\<And>x k. \<lbrakk>(x, k)\<in>p;  k \<subseteq> d x\<rbrakk> \<Longrightarrow> (x, k) \<in> q1"
-      using case_prodD tagged_partial_division_subset[OF insert(4) subset_insertI]
-      by (metis (mono_tags, lifting) insert.hyps(3) insert.prems(2))
+      using insert
+      by (metis (mono_tags, lifting) case_prodD subset_insertI tagged_partial_division_subset)
     have *: "\<Union>{l. \<exists>y. (y,l) \<in> insert xk p} = k \<union> \<Union>{l. \<exists>y. (y,l) \<in> p}"
       unfolding xk by auto
     note p = tagged_partial_division_ofD[OF insert(4)]
     obtain u v where uv: "k = cbox u v"
       using p(4) xk by blast
-    have "finite {k. \<exists>x. (x, k) \<in> p}"
-      apply (rule finite_subset[of _ "snd ` p"])
-      using image_iff apply fastforce
-      using insert.hyps(1) by blast
+    have "{K. \<exists>x. (x, K) \<in> p} \<subseteq> snd ` p"
+      by force
+    then have "finite {K. \<exists>x. (x, K) \<in> p}"
+      using finite_surj insert.hyps(1) by blast
     then have int: "interior (cbox u v) \<inter> interior (\<Union>{k. \<exists>x. (x, k) \<in> p}) = {}"
     proof (rule Int_interior_Union_intervals)
       show "open (interior (cbox u v))"
         by auto
-      show "\<And>T. T \<in> {k. \<exists>x. (x, k) \<in> p} \<Longrightarrow> \<exists>a b. T = cbox a b"
+      show "\<And>T. T \<in> {K. \<exists>x. (x, K) \<in> p} \<Longrightarrow> \<exists>a b. T = cbox a b"
         using p(4) by auto
-      show "\<And>T. T \<in> {k. \<exists>x. (x, k) \<in> p} \<Longrightarrow> interior (cbox u v) \<inter> interior T = {}"
-        by clarify (metis insert.hyps(2) insert_iff interior_cbox p(5) uv xk)
+      show "\<And>T. T \<in> {K. \<exists>x. (x, K) \<in> p} \<Longrightarrow> interior (cbox u v) \<inter> interior T = {}"
+        using insert.hyps(2) p(5) uv xk by blast
     qed
     show ?case
     proof (cases "cbox u v \<subseteq> d x")
       case True
       have "{(x, cbox u v)} tagged_division_of cbox u v"
         by (simp add: p(2) uv xk tagged_division_of_self)
-      then have "{(x, cbox u v)} \<union> q1 tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> insert xk p}"
-        unfolding * uv by (metis (no_types, lifting) int q1 tagged_division_Un)
-      with True show ?thesis
-        apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI)
-        using \<open>d fine q1\<close> fine_def q1I uv xk apply fastforce
-        done
+      then have "{(x, cbox u v)} \<union> q1 tagged_division_of \<Union>{K. \<exists>x. (x, K) \<in> insert xk p}"
+        by (smt (verit, best) "*" int q1 tagged_division_Un uv)
+      moreover have "d fine ({(x,cbox u v)} \<union> q1)"
+        using True \<open>d fine q1\<close> fine_def by fastforce
+      ultimately show ?thesis
+        by (metis (no_types, lifting) case_prodI2 insert_iff insert_is_Un q1I uv xk)
     next
       case False
       obtain q2 where q2: "q2 tagged_division_of cbox u v" "d fine q2"
         using fine_division_exists[OF assms(2)] by blast
       show ?thesis
-        apply (rule_tac x="q2 \<union> q1" in exI)
-        apply (intro conjI)
-        unfolding * uv
-        apply (rule tagged_division_Un q2 q1 int fine_Un)+
-          apply (auto intro: q1 q2 fine_Un \<open>d fine q1\<close> simp add: False q1I uv xk)
-        done
+      proof (intro exI conjI)
+        show "q2 \<union> q1 tagged_division_of \<Union> {k. \<exists>x. (x, k) \<in> insert xk p}"
+          by (smt (verit, best) "*" int q1 q2(1) tagged_division_Un uv)
+        show "d fine q2 \<union> q1"
+          by (simp add: \<open>d fine q1\<close> fine_Un q2(2))
+      qed (use False uv xk q1I in auto)
     qed
   qed
   with p obtain q where q: "q tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> p}" "d fine q" "\<forall>(x, k)\<in>p. k \<subseteq> d x \<longrightarrow> (x, k) \<in> q"
@@ -2335,8 +2065,7 @@
     show "\<Union>?D0 \<subseteq> cbox a b"
       apply (simp add: UN_subset_iff)
       apply (intro conjI allI ballI subset_box_imp)
-       apply (simp add: field_simps)
-      apply (auto intro: mult_right_mono aibi)
+       apply (simp add: field_simps aibi mult_right_mono)
       apply (force simp: aibi scaling_mono nat_less_real_le dest: PiE_mem intro: mult_right_mono)
       done
   next
@@ -2352,22 +2081,19 @@
       fix v w m and n::nat
       assume "m \<le> n" \<comment> \<open>WLOG we can assume \<^term>\<open>m \<le> n\<close>, when the first disjunct becomes impossible\<close>
       have "?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
+        apply (rule ccontr)
         apply (simp add: subset_box disjoint_interval)
-        apply (rule ccontr)
         apply (clarsimp simp add: aibi mult_le_cancel_right divide_le_cancel not_less not_le)
         apply (drule_tac x=i in bspec, assumption)
         using \<open>m\<le>n\<close> realff [of _ _ "1+_"] realff [of "1+_"_ "1+_"]
         apply (auto simp: divide_simps add.commute not_le nat_le_iff_add realff)
-         apply (simp_all add: power_add)
-        apply (metis (no_types, opaque_lifting) mult_Suc mult_less_cancel2 not_less_eq mult.assoc)
-        apply (metis (no_types, opaque_lifting) mult_Suc mult_less_cancel2 not_less_eq mult.assoc)
+        apply (metis (no_types, opaque_lifting) power_add mult_Suc mult_less_cancel2 not_less_eq mult.assoc)+
         done
       then show "?K0(m,v) \<subseteq> ?K0(n,w) \<or> ?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
         by meson
     qed auto
     show "\<And>A B. \<lbrakk>A \<in> ?D0; B \<in> ?D0\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
-      apply (erule imageE SigmaE)+
-      using * by simp
+      using * by fastforce
   next
     show "\<exists>K \<in> ?D0. x \<in> K \<and> K \<subseteq> g x" if "x \<in> S" for x
     proof (simp only: bex_simps split_paired_Bex_Sigma)
@@ -2389,9 +2115,8 @@
               by (meson sum_bounded_above that)
             also have "... = \<epsilon> / 2"
               by (simp add: field_split_simps)
-            also have "... < \<epsilon>"
-              by (simp add: \<open>0 < \<epsilon>\<close>)
-            finally show ?thesis .
+            finally show ?thesis
+              using \<open>0 < \<epsilon>\<close> by linarith 
           qed
           then show ?thesis
             by (rule_tac e = "\<epsilon> / 2 / DIM('a)" in that) (simp_all add:  \<open>0 < \<epsilon>\<close> dist_norm subsetD [OF \<epsilon>])
@@ -2404,13 +2129,7 @@
         then have "norm (b - a) < e * 2^n"
           by (auto simp: field_split_simps)
         then have bai: "b \<bullet> i - a \<bullet> i < e * 2 ^ n" if "i \<in> Basis" for i
-        proof -
-          have "b \<bullet> i - a \<bullet> i \<le> norm (b - a)"
-            by (metis abs_of_nonneg dual_order.trans inner_diff_left linear norm_ge_zero Basis_le_norm that)
-          also have "... < e * 2 ^ n"
-            using \<open>norm (b - a) < e * 2 ^ n\<close> by blast
-          finally show ?thesis .
-        qed
+          by (smt (verit, del_insts) Basis_le_norm diff_add_cancel inner_simps(1) that)
         have D: "(a + n \<le> x \<and> x \<le> a + m) \<Longrightarrow> (a + n \<le> y \<and> y \<le> a + m) \<Longrightarrow> abs(x - y) \<le> m - n"
                  for a m n x and y::real
           by auto
@@ -2438,9 +2157,9 @@
             next
               have "a \<bullet> i + real ?k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le>
                     a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
+                using aibi [OF \<open>i \<in> Basis\<close>] xab 2
                 apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
-                using aibi [OF \<open>i \<in> Basis\<close>] xab 2
-                  apply (simp_all add: \<open>i \<in> Basis\<close> mem_box field_split_simps)
+                  apply (auto simp: \<open>i \<in> Basis\<close> mem_box field_split_simps)
                 done
               also have "... = x \<bullet> i"
                 using abi_less by (simp add: field_split_simps)
@@ -2449,8 +2168,8 @@
               have "x \<bullet> i \<le> a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
                 using abi_less by (simp add: field_split_simps)
               also have "... \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
+                using aibi [OF \<open>i \<in> Basis\<close>] xab
                 apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
-                using aibi [OF \<open>i \<in> Basis\<close>] xab
                   apply (auto simp: \<open>i \<in> Basis\<close> mem_box divide_simps)
                 done
               finally show "x \<bullet> i \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n" .
@@ -2540,7 +2259,7 @@
   qed
   let ?\<D> = "{K \<in> \<D>. \<forall>K'. K' \<in> \<D> \<and> K \<noteq> K' \<longrightarrow> \<not>(K \<subseteq> K')}"
   show ?thesis
-  proof (rule that)
+  proof 
     show "countable ?\<D>"
       by (blast intro: countable_subset [OF _ count])
     show "\<Union>?\<D> \<subseteq> cbox a b"
@@ -2597,6 +2316,6 @@
 
 lemma eventually_division_filter_tagged_division:
   "eventually (\<lambda>p. p tagged_division_of s) (division_filter s)"
-  unfolding eventually_division_filter by (intro exI[of _ "\<lambda>x. ball x 1"]) auto
+  using eventually_division_filter by auto
 
 end
--- a/src/HOL/Complex_Analysis/Conformal_Mappings.thy	Tue Aug 01 11:27:55 2023 +0200
+++ b/src/HOL/Complex_Analysis/Conformal_Mappings.thy	Thu Aug 03 19:10:43 2023 +0200
@@ -1751,7 +1751,7 @@
         have "(*) ((cmod z)\<^sup>2) integrable_on {0..1}"
           by (metis ident_integrable_on integrable_0 integrable_eq integrable_on_cmult_iff lambda_zero)
         then show "(\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}"
-          using integrable_on_cmult_right[where 'b=real, simplified] integrable_on_cdivide [where 'b=real, simplified]
+          using integrable_on_cmult_right[where 'b=real, simplified] integrable_on_divide [where 'b=real, simplified]
           by blast
       qed (simp add: norm_mult power2_eq_square 4)
       then have int_le: "norm (f z - deriv f 0 * z) \<le> (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))"