merged
authordesharna
Tue, 03 Jan 2023 18:23:52 +0100
changeset 76878 b3c5bc06f5be
parent 76876 c9ffd9cf58db (diff)
parent 76877 c9e091867206 (current diff)
child 76892 7fd3e461d3b6
child 76896 c0459ee8220c
merged
--- a/Admin/components/bundled	Mon Dec 26 14:34:32 2022 +0100
+++ b/Admin/components/bundled	Tue Jan 03 18:23:52 2023 +0100
@@ -1,2 +1,2 @@
 #additional components to be bundled for release
-naproche-20221024
+#naproche-20221024
--- a/etc/build.props	Mon Dec 26 14:34:32 2022 +0100
+++ b/etc/build.props	Tue Jan 03 18:23:52 2023 +0100
@@ -315,7 +315,7 @@
   isabelle.Scala$Handler \
   isabelle.Scala_Functions \
   isabelle.Server_Commands \
-  isabelle.Sessions$File_Format \
+  isabelle.Sessions$ROOTS_File_Format \
   isabelle.Simplifier_Trace$Handler \
   isabelle.Tools \
   isabelle.jedit.JEdit_Plugin0 \
--- a/src/HOL/Analysis/Affine.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Affine.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -51,7 +51,7 @@
     using dim_unique[of B B "card B"] assms span_superset[of B] by auto
   have "dim B \<le> card (Basis :: 'a set)"
     using dim_subset_UNIV[of B] by simp
-  from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B"
+  from obtain_subset_with_card_n[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B"
     by auto
   let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
   have "\<exists>f. linear f \<and> f ` B = d \<and> f ` span B = ?t \<and> inj_on f (span B)"
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -83,11 +83,8 @@
 
 lemma retraction_openin_vimage_iff:
   "openin (top_of_set S) (S \<inter> r -` U) \<longleftrightarrow> openin (top_of_set T) U"
-  if retraction: "retraction S T r" and "U \<subseteq> T"
-  using retraction apply (rule retractionE)
-  apply (rule continuous_right_inverse_imp_quotient_map [where g=r])
-  using \<open>U \<subseteq> T\<close> apply (auto elim: continuous_on_subset)
-  done
+  if "retraction S T r" and "U \<subseteq> T"
+  by (simp add: retraction_openin_vimage_iff that)
 
 lemma retract_of_locally_compact:
     fixes S :: "'a :: {heine_borel,real_normed_vector} set"
@@ -251,7 +248,7 @@
   then obtain u l where "l \<in> s" "\<forall>b\<in>s. l \<le> b" "u \<in> s" "\<forall>b\<in>s. b \<le> u"
     using insert by auto
   with * show "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. a \<le> x" "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. x \<le> a"
-    using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+
+    by (metis insert_iff order.trans)+
 qed auto
 
 lemma kuhn_labelling_lemma:
@@ -1081,7 +1078,7 @@
       qed }
     with ks_f' eq \<open>a \<noteq> b.enum 0\<close> \<open>n \<noteq> 0\<close> show ?thesis
       apply (intro ex1I[of _ "b.enum ` {.. n}"])
-      apply auto []
+      apply fastforce
       apply metis
       done
   next
@@ -1111,16 +1108,11 @@
     have "{i} \<subseteq> {..n}"
       using i by auto
     { fix j assume "j \<le> n"
-      moreover have "j < i \<or> i = j \<or> i < j" by arith
-      moreover note i
-      ultimately have "enum j = b.enum j \<longleftrightarrow> j \<noteq> i"
-        apply (simp only: fun_eq_iff enum_def b.enum_def flip: image_comp)
+      with i Suc_i' have "enum j = b.enum j \<longleftrightarrow> j \<noteq> i"
+        unfolding fun_eq_iff enum_def b.enum_def image_comp [symmetric]
         apply (cases \<open>i = j\<close>)
-         apply simp
-         apply (metis Suc_i' \<open>i \<le> n\<close> imageI in_upd_image lessI lessThan_iff lessThan_subset_iff less_irrefl_nat transpose_apply_second transpose_commute)
-        apply (subst transpose_image_eq)
-         apply (auto simp add: i'_def)
-        done
+         apply (metis imageI in_upd_image lessI lessThan_iff lessThan_subset_iff order_less_le transpose_apply_first)
+        by (metis lessThan_iff linorder_not_less not_less_eq_eq order_less_le transpose_image_eq)
       }
     note enum_eq_benum = this
     then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})"
@@ -1295,7 +1287,7 @@
 
 lemma reduced_labelling_unique:
   "r \<le> n \<Longrightarrow> \<forall>i<r. x i = 0 \<Longrightarrow> r = n \<or> x r \<noteq> 0 \<Longrightarrow> reduced n x = r"
- unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) (metis le_less not_le)+
+  by (metis linorder_less_linear linorder_not_le reduced_labelling)
 
 lemma reduced_labelling_zero: "j < n \<Longrightarrow> x j = 0 \<Longrightarrow> reduced n x \<noteq> j"
   using reduced_labelling[of n x] by auto
@@ -1369,7 +1361,7 @@
       by (auto simp: out_eq_p)
     moreover
     { fix x assume "x \<in> s"
-      with lab_1[rule_format, of n x] all_eq_p s_le_p[of x]
+      with lab_1[rule_format, of n x] all_eq_p s_le_p[of x] 
       have "?rl x \<le> n"
         by (auto intro!: reduced_labelling_nonzero)
       then have "?rl x = reduced n (lab x)"
@@ -1532,34 +1524,9 @@
              (\<forall>x i. P x \<and> Q i \<and> x i = 1 \<longrightarrow> l x i = 1) \<and>
              (\<forall>x i. P x \<and> Q i \<and> l x i = 0 \<longrightarrow> x i \<le> f x i) \<and>
              (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)"
-proof -
-  have and_forall_thm: "\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)"
-    by auto
-  have *: "\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x"
-    by auto
-  show ?thesis
-    unfolding and_forall_thm
-    apply (subst choice_iff[symmetric])+
-    apply rule
-    apply rule
-  proof -
-    fix x x'
-    let ?R = "\<lambda>y::nat.
-      (P x \<and> Q x' \<and> x x' = 0 \<longrightarrow> y = 0) \<and>
-      (P x \<and> Q x' \<and> x x' = 1 \<longrightarrow> y = 1) \<and>
-      (P x \<and> Q x' \<and> y = 0 \<longrightarrow> x x' \<le> (f x) x') \<and>
-      (P x \<and> Q x' \<and> y = 1 \<longrightarrow> (f x) x' \<le> x x')"
-    have "0 \<le> f x x' \<and> f x x' \<le> 1" if "P x" "Q x'"
-      using assms(2)[rule_format,of "f x" x'] that
-      apply (drule_tac assms(1)[rule_format])
-      apply auto
-      done
-    then have "?R 0 \<or> ?R 1"
-      by auto
-    then show "\<exists>y\<le>1. ?R y"
-      by auto
-  qed
-qed
+  unfolding all_conj_distrib [symmetric] 
+  apply (subst choice_iff[symmetric])+
+  by (metis assms bot_nat_0.extremum nle_le zero_neq_one)
 
 subsection \<open>Brouwer's fixed point theorem\<close>
 
@@ -1785,11 +1752,7 @@
     using b' unfolding bij_betw_def by auto
   define r' ::'a where "r' = (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)"
   have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> p"
-    apply (rule order_trans)
-    apply (rule rs(1)[OF b'_im,THEN conjunct2])
-    using q(1)[rule_format,OF b'_im]
-    apply (auto simp: Suc_le_eq)
-    done
+    using b'_im q(1) rs(1) by fastforce
   then have "r' \<in> cbox 0 One"
     unfolding r'_def cbox_def
     using b'_Basis
@@ -1899,9 +1862,8 @@
   have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point S) x = x"
   proof (rule_tac brouwer_ball[OF e(1)])
     show "continuous_on (cball 0 e) (f \<circ> closest_point S)"
-      apply (rule continuous_on_compose)
-      using S compact_eq_bounded_closed continuous_on_closest_point apply blast
-      by (meson S contf closest_point_in_set compact_imp_closed continuous_on_subset image_subsetI)
+      by (meson assms closest_point_in_set compact_eq_bounded_closed contf continuous_on_closest_point
+          continuous_on_compose continuous_on_subset image_subsetI)
     show "(f \<circ> closest_point S) ` cball 0 e \<subseteq> cball 0 e"
       by clarsimp (metis S fim closest_point_exists(1) compact_eq_bounded_closed e(2) image_subset_iff mem_cball_0 subsetCE)
   qed (use assms in auto)
@@ -2033,9 +1995,7 @@
     using \<open>bounded S\<close> bounded_subset_ballD by blast
   have notga: "g x \<noteq> a" for x
     unfolding g_def using fros fim \<open>a \<notin> T\<close>
-    apply (auto simp: frontier_def)
-    using fid interior_subset apply fastforce
-    by (simp add: \<open>a \<in> S\<close> closure_def)
+    by (metis Un_iff \<open>a \<in> S\<close> closure_Un_frontier fid imageI subset_eq)
   define h where "h \<equiv> (\<lambda>y. a + (B / norm(y - a)) *\<^sub>R (y - a)) \<circ> g"
   have "\<not> (frontier (cball a B) retract_of (cball a B))"
     by (metis no_retraction_cball \<open>0 < B\<close>)
@@ -2119,8 +2079,7 @@
           using k \<open>x \<noteq> 0\<close> \<open>0 < dd x\<close>
           apply (simp add: in_segment)
           apply (rule_tac x = "dd x / k" in exI)
-          apply (simp add: field_simps that)
-          apply (simp add: vector_add_divide_simps algebra_simps)
+          apply (simp add: that vector_add_divide_simps algebra_simps)
           done
         ultimately show ?thesis
           using segsub by (auto simp: rel_frontier_def)
@@ -2163,8 +2122,7 @@
         then have "\<not> dd (x - a) \<le> 0 \<and> a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S"
           using \<open>x \<noteq> a\<close> dd1 by fastforce
         with \<open>x \<noteq> a\<close> show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<noteq> a"
-          apply (auto simp: iff)
-          using less_eq_real_def mult_le_0_iff not_less u by fastforce
+          using less_eq_real_def mult_le_0_iff not_less u by (fastforce simp: iff)
       qed
     qed
     show "retraction (T - {a}) (rel_frontier S) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
@@ -2216,9 +2174,8 @@
   fixes S :: "'a::euclidean_space set"
   assumes "bounded S" "convex S" "a \<in> rel_interior S"
     shows "rel_frontier S retract_of (affine hull S - {a})"
-apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S "affine hull S" a])
-apply (auto simp: affine_imp_convex rel_frontier_affine_hull retract_of_def assms)
-done
+  by (meson assms convex_affine_hull dual_order.refl rel_frontier_affine_hull 
+      rel_frontier_deformation_retract_of_punctured_convex retract_of_def)
 
 corollary rel_boundary_retract_of_punctured_affine_hull:
   fixes S :: "'a::euclidean_space set"
@@ -2231,29 +2188,24 @@
   fixes S :: "'a::euclidean_space set"
   assumes "convex S" "bounded S" "a \<in> rel_interior S" "convex T" "rel_frontier S \<subseteq> T" "T \<subseteq> affine hull S"
   shows "(rel_frontier S) homotopy_eqv (T - {a})"
-  apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S T])
-  using assms
-  apply auto
-  using deformation_retract_imp_homotopy_eqv homotopy_equivalent_space_sym by blast
+  by (meson assms deformation_retract_imp_homotopy_eqv homotopy_equivalent_space_sym 
+      rel_frontier_deformation_retract_of_punctured_convex[of S T])
 
 lemma homotopy_eqv_rel_frontier_punctured_affine_hull:
   fixes S :: "'a::euclidean_space set"
   assumes "convex S" "bounded S" "a \<in> rel_interior S"
     shows "(rel_frontier S) homotopy_eqv (affine hull S - {a})"
-apply (rule homotopy_eqv_rel_frontier_punctured_convex)
-  using assms rel_frontier_affine_hull  by force+
+  by (simp add: assms homotopy_eqv_rel_frontier_punctured_convex rel_frontier_affine_hull)
 
 lemma path_connected_sphere_gen:
   assumes "convex S" "bounded S" "aff_dim S \<noteq> 1"
   shows "path_connected(rel_frontier S)"
-proof (cases "rel_interior S = {}")
-  case True
+proof -
+  have "convex (closure S)" 
+    using assms by auto
   then show ?thesis
-    by (simp add: \<open>convex S\<close> convex_imp_path_connected rel_frontier_def)
-next
-  case False
-  then show ?thesis
-    by (metis aff_dim_affine_hull affine_affine_hull affine_imp_convex all_not_in_conv assms path_connected_punctured_convex rel_frontier_retract_of_punctured_affine_hull retract_of_path_connected)
+    by (metis Diff_empty aff_dim_affine_hull assms convex_affine_hull convex_imp_path_connected equals0I 
+       path_connected_punctured_convex rel_frontier_def rel_frontier_retract_of_punctured_affine_hull retract_of_path_connected)
 qed
 
 lemma connected_sphere_gen:
@@ -2282,8 +2234,7 @@
   then show ?thesis
     apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def)
     apply (rule_tac x = "\<lambda>z. inverse(norm(snd z - (g \<circ> fst)z)) *\<^sub>R (snd z - (g \<circ> fst)z)" in exI)
-    apply (intro conjI continuous_intros)
-    apply (rule continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+
+    apply (rule conjI continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+
     done
 qed
 
@@ -2315,8 +2266,8 @@
       using ball_subset_cball [of a r] r by auto
     have cont1: "continuous_on (s \<union> connected_component_set (- s) a)
                      (\<lambda>x. a + r *\<^sub>R g x)"
-      apply (rule continuous_intros)+
-      using \<open>continuous_on (s \<union> c) g\<close> ceq by blast
+      using \<open>continuous_on (s \<union> c) g\<close> ceq
+      by (intro continuous_intros) blast
     have cont2: "continuous_on (cball a r - connected_component_set (- s) a)
             (\<lambda>x. a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))"
       by (rule continuous_intros | force simp: \<open>a \<notin> s\<close>)+
@@ -2326,10 +2277,8 @@
                   else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))"
       apply (subst cb_eq)
       apply (rule continuous_on_cases [OF _ _ cont1 cont2])
-        using ceq cin
-      apply (auto intro: closed_Un_complement_component
-                  simp: \<open>closed s\<close> open_Compl open_connected_component)
-      done
+      using \<open>closed s\<close> ceq cin 
+      by (force simp: closed_Diff open_Compl closed_Un_complement_component open_connected_component)+
     have 2: "(\<lambda>x. a + r *\<^sub>R g x) ` (cball a r \<inter> connected_component_set (- s) a)
              \<subseteq> sphere a r "
       using \<open>0 < r\<close> by (force simp: dist_norm ceq)
@@ -2384,12 +2333,7 @@
     and "x \<in> S"
     and "\<And>x y. \<lbrakk>x\<in>S; y\<in>cball a e\<rbrakk> \<Longrightarrow> x + (y - f y) \<in> cball a e"
   shows "\<exists>y\<in>cball a e. f y = x"
-  apply (rule brouwer_surjective)
-  apply (rule compact_cball convex_cball)+
-  unfolding cball_eq_empty
-  using assms
-  apply auto
-  done
+  by (smt (verit, best) assms brouwer_surjective cball_eq_empty compact_cball convex_cball)
 
 subsubsection \<open>Inverse function theorem\<close>
 
@@ -2510,8 +2454,7 @@
 
 lemma has_derivative_inverse_strong:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
-  assumes "open S"
-    and "x \<in> S"
+  assumes S: "open S" "x \<in> S"
     and contf: "continuous_on S f"
     and gf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
     and derf: "(f has_derivative f') (at x)"
@@ -2524,23 +2467,16 @@
     unfolding linear_conv_bounded_linear[symmetric]
     using id right_inverse_linear by blast
   moreover have "g' \<circ> f' = id"
-    using id linf ling
-    unfolding linear_conv_bounded_linear[symmetric]
-    using linear_inverse_left
-    by auto
+    using id linear_inverse_left linear_linear linf ling by blast
   moreover have *: "\<And>T. \<lbrakk>T \<subseteq> S; x \<in> interior T\<rbrakk> \<Longrightarrow> f x \<in> interior (f ` T)"
-    apply (rule sussmann_open_mapping)
-    apply (rule assms ling)+
-    apply auto
-    done
+    using S derf contf id ling sussmann_open_mapping by blast
   have "continuous (at (f x)) g"
     unfolding continuous_at Lim_at
-  proof (rule, rule)
+  proof (intro strip)
     fix e :: real
     assume "e > 0"
     then have "f x \<in> interior (f ` (ball x e \<inter> S))"
-      using *[rule_format,of "ball x e \<inter> S"] \<open>x \<in> S\<close>
-      by (auto simp add: interior_open[OF open_ball] interior_open[OF assms(1)])
+      by (simp add: "*" S interior_open)
     then obtain d where d: "0 < d" "ball (f x) d \<subseteq> f ` (ball x e \<inter> S)"
       unfolding mem_interior by blast
     show "\<exists>d>0. \<forall>y. 0 < dist y (f x) \<and> dist y (f x) < d \<longrightarrow> dist (g y) (g (f x)) < e"
@@ -2550,16 +2486,11 @@
       then have "g y \<in> g ` f ` (ball x e \<inter> S)"
         by (metis d(2) dist_commute mem_ball rev_image_eqI subset_iff)
       then show "dist (g y) (g (f x)) < e"
-        using gf[OF \<open>x \<in> S\<close>]
-        by (simp add: assms(4) dist_commute image_iff)
+        using \<open>x \<in> S\<close> by (simp add: gf dist_commute image_iff)
     qed (use d in auto)
   qed
   moreover have "f x \<in> interior (f ` S)"
-    apply (rule sussmann_open_mapping)
-    apply (rule assms ling)+
-    using interior_open[OF assms(1)] and \<open>x \<in> S\<close>
-    apply auto
-    done
+    using "*" S interior_eq by blast
   moreover have "f (g y) = y" if "y \<in> interior (f ` S)" for y
     by (metis gf imageE interiorE subsetD that)
   ultimately show ?thesis using assms
@@ -2576,26 +2507,20 @@
     and "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
     and "(f has_derivative f') (at (g y))"
     and "f' \<circ> g' = id"
-    and "f (g y) = y"
+    and f: "f (g y) = y"
   shows "(g has_derivative g') (at y)"
-  using has_derivative_inverse_strong[OF assms(1-6)]
-  unfolding assms(7)
-  by simp
+  using has_derivative_inverse_strong[OF assms(1-6)] by (simp add: f)
 
 text \<open>On a region.\<close>
 
 theorem has_derivative_inverse_on:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
   assumes "open S"
-    and derf: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f'(x)) (at x)"
+    and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f'(x)) (at x)"
     and "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
     and "f' x \<circ> g' x = id"
     and "x \<in> S"
   shows "(g has_derivative g'(x)) (at (f x))"
-proof (rule has_derivative_inverse_strong[where g'="g' x" and f=f])
-  show "continuous_on S f"
-  unfolding continuous_on_eq_continuous_at[OF \<open>open S\<close>]
-  using derf has_derivative_continuous by blast
-qed (use assms in auto)
+  by (meson assms continuous_on_eq_continuous_at has_derivative_continuous has_derivative_inverse_strong)
 
 end
--- a/src/HOL/Analysis/Complex_Transcendental.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Complex_Transcendental.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -660,15 +660,10 @@
 
 lemma norm_sin_squared:
   "norm(sin z) ^ 2 = (exp(2 * Im z) + inverse(exp(2 * Im z)) - 2 * cos(2 * Re z)) / 4"
-proof (cases z)
-  case (Complex x1 x2)
-  then show ?thesis
-    apply (simp only: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double Complex_eq)
-    apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
-    apply (simp only: left_diff_distrib [symmetric] power_mult_distrib cos_squared_eq)
-    apply (simp add: power2_eq_square field_split_simps)
-    done
-qed
+  using cos_double_sin [of "Re z"]
+  apply (simp add: sin_cos_eq norm_cos_squared exp_minus mult.commute [of _ 2] exp_double)
+  apply (simp add: algebra_simps power2_eq_square)
+  done
 
 lemma exp_uminus_Im: "exp (- Im z) \<le> exp (cmod z)"
   using abs_Im_le_cmod linear order_trans by fastforce
@@ -1002,17 +997,14 @@
 
 lemma Arg2pi_times_of_real [simp]:
   assumes "0 < r" shows "Arg2pi (of_real r * z) = Arg2pi z"
-proof (cases "z=0")
-  case False
-  show ?thesis
-    by (rule Arg2pi_unique [of "r * norm z"]) (use Arg2pi False assms is_Arg_def in auto)
-qed auto
+  by (metis (no_types, lifting) Arg2pi Arg2pi_eq Arg2pi_unique assms mult.assoc 
+      mult_eq_0_iff mult_pos_pos of_real_mult zero_less_norm_iff)
 
 lemma Arg2pi_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg2pi (z * of_real r) = Arg2pi z"
   by (metis Arg2pi_times_of_real mult.commute)
 
 lemma Arg2pi_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg2pi (z / of_real r) = Arg2pi z"
-  by (metis Arg2pi_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff)
+  by (metis Arg2pi_times_of_real2 less_irrefl nonzero_eq_divide_eq of_real_eq_0_iff)
 
 lemma Arg2pi_le_pi: "Arg2pi z \<le> pi \<longleftrightarrow> 0 \<le> Im z"
 proof (cases "z=0")
@@ -1095,12 +1087,14 @@
 
 lemma Arg2pi_eq_iff:
   assumes "w \<noteq> 0" "z \<noteq> 0"
-     shows "Arg2pi w = Arg2pi z \<longleftrightarrow> (\<exists>x. 0 < x & w = of_real x * z)"
-  using assms Arg2pi_eq [of z] Arg2pi_eq [of w]
-  apply auto
-  apply (rule_tac x="norm w / norm z" in exI)
-  apply (simp add: field_split_simps)
-  by (metis mult.commute mult.left_commute)
+  shows "Arg2pi w = Arg2pi z \<longleftrightarrow> (\<exists>x. 0 < x & w = of_real x * z)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then have "(cmod w) * (z / cmod z) = w"
+    by (metis Arg2pi_eq assms(2) mult_eq_0_iff nonzero_mult_div_cancel_left)
+  then show ?rhs
+    by (metis assms divide_pos_pos of_real_divide times_divide_eq_left times_divide_eq_right zero_less_norm_iff)
+qed auto
 
 lemma Arg2pi_inverse_eq_0: "Arg2pi(inverse z) = 0 \<longleftrightarrow> Arg2pi z = 0"
   by (metis Arg2pi_eq_0 Arg2pi_inverse inverse_inverse_eq)
@@ -1142,19 +1136,20 @@
   using Arg2pi_add [OF assms]
   by auto
 
-lemma Arg2pi_cnj_eq_inverse: "z\<noteq>0 \<Longrightarrow> Arg2pi (cnj z) = Arg2pi (inverse z)"
-  apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric])
-  by (metis of_real_power zero_less_norm_iff zero_less_power)
+lemma Arg2pi_cnj_eq_inverse:
+  assumes "z \<noteq> 0" shows "Arg2pi (cnj z) = Arg2pi (inverse z)"
+proof -
+  have "\<exists>r>0. of_real r / z = cnj z"
+    by (metis assms complex_norm_square nonzero_mult_div_cancel_left zero_less_norm_iff zero_less_power)
+  then show ?thesis
+    by (metis Arg2pi_times_of_real2 divide_inverse_commute)
+qed
 
 lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \<in> \<real> then Arg2pi z else 2*pi - Arg2pi z)"
-proof (cases "z=0")
-  case False
-  then show ?thesis
-    by (simp add: Arg2pi_cnj_eq_inverse Arg2pi_inverse)
-qed auto
+  by (metis Arg2pi_cnj_eq_inverse Arg2pi_inverse Reals_cnj_iff complex_cnj_zero)
 
 lemma Arg2pi_exp: "0 \<le> Im z \<Longrightarrow> Im z < 2*pi \<Longrightarrow> Arg2pi(exp z) = Im z"
-  by (rule Arg2pi_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar)
+  by (simp add: Arg2pi_unique exp_eq_polar)
 
 lemma complex_split_polar:
   obtains r a::real where "z = complex_of_real r * (cos a + \<i> * sin a)" "0 \<le> r" "0 \<le> a" "a < 2*pi"
@@ -2215,21 +2210,13 @@
 
 lemma Arg_times_of_real [simp]:
   assumes "0 < r" shows "Arg (of_real r * z) = Arg z"
-proof (cases "z=0")
-  case True
-  then show ?thesis
-    by (simp add: Arg_def)
-next
-  case False
-  with Arg_eq assms show ?thesis
-  by (auto simp: mpi_less_Arg Arg_le_pi intro!:  Arg_unique [of "r * norm z"])
-qed
+  using Arg_def Ln_times_of_real assms by auto
 
 lemma Arg_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg (z * of_real r) = Arg z"
   by (metis Arg_times_of_real mult.commute)
 
 lemma Arg_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg (z / of_real r) = Arg z"
-  by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff)
+  by (metis Arg_times_of_real2 less_irrefl nonzero_eq_divide_eq of_real_eq_0_iff)
 
 lemma Arg_less_0: "0 \<le> Arg z \<longleftrightarrow> 0 \<le> Im z"
   using Im_Ln_le_pi Im_Ln_pos_le
@@ -2243,18 +2230,13 @@
   by (auto simp: order.order_iff_strict Arg_def)
 
 lemma Arg_eq_0: "Arg z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
-  using complex_is_Real_iff
-  by (simp add: Arg_def Im_Ln_eq_0) (metis less_eq_real_def of_real_Re of_real_def scale_zero_left)
+  using Arg_def Im_Ln_eq_0 complex_eq_iff complex_is_Real_iff by auto
 
 corollary\<^marker>\<open>tag unimportant\<close> Arg_ne_0: assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0" shows "Arg z \<noteq> 0"
   using assms by (auto simp: nonneg_Reals_def Arg_eq_0)
 
 lemma Arg_eq_pi_iff: "Arg z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
-proof (cases "z=0")
-  case False
-  then show ?thesis
-    using Arg_eq_0 [of "-z"] Arg_eq_pi complex_is_Real_iff by blast
-qed (simp add: Arg_def)
+  using Arg_eq_pi complex_is_Real_iff by blast
 
 lemma Arg_eq_0_pi: "Arg z = 0 \<or> Arg z = pi \<longleftrightarrow> z \<in> \<real>"
   using Arg_eq_pi_iff Arg_eq_0 by force
@@ -2266,15 +2248,11 @@
 proof (cases "z \<in> \<real>")
   case True
   then show ?thesis
-    by simp (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse)
+    by (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse)
 next
   case False
-  then have z: "Arg z < pi" "z \<noteq> 0"
-    using Arg_eq_0_pi Arg_le_pi by (auto simp: less_eq_real_def)
-  show ?thesis
-    apply (rule Arg_unique [of "inverse (norm z)"])
-    using False z mpi_less_Arg [of z] Arg_eq [of z]
-    by (auto simp: exp_minus field_simps)
+  then show ?thesis
+    by (simp add: Arg_def Ln_inverse complex_is_Real_iff complex_nonpos_Reals_iff)
 qed
 
 lemma Arg_eq_iff:
@@ -2282,7 +2260,7 @@
   shows "Arg w = Arg z \<longleftrightarrow> (\<exists>x. 0 < x \<and> w = of_real x * z)" (is "?lhs = ?rhs")
 proof
   assume ?lhs
-  then have "w = complex_of_real (cmod w / cmod z) * z"
+  then have "w = (cmod w / cmod z) * z"
     by (metis Arg_eq assms divide_divide_eq_right eq_divide_eq exp_not_eq_zero of_real_divide)
   then show ?rhs
     using assms divide_pos_pos zero_less_norm_iff by blast
@@ -2307,14 +2285,11 @@
   assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
     shows "continuous (at z) Arg"
 proof -
-  have [simp]: "(\<lambda>z. Im (Ln z)) \<midarrow>z\<rightarrow> Arg z"
+  have "(\<lambda>z. Im (Ln z)) \<midarrow>z\<rightarrow> Arg z"
     using Arg_def assms continuous_at by fastforce
-  show ?thesis
+  then show ?thesis
     unfolding continuous_at
-  proof (rule Lim_transform_within_open)
-    show "\<And>w. \<lbrakk>w \<in> - \<real>\<^sub>\<le>\<^sub>0; w \<noteq> z\<rbrakk> \<Longrightarrow> Im (Ln w) = Arg w"
-      by (metis Arg_def Compl_iff nonpos_Reals_zero_I)
-  qed (use assms in auto)
+    by (smt (verit, del_insts) Arg_eq_Im_Ln Lim_transform_away_at assms nonpos_Reals_zero_I)
 qed
 
 lemma continuous_within_Arg: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z within S) Arg"
@@ -2388,44 +2363,17 @@
   qed
 next
   show "arctan (Im z / Re z) > -pi"
-    by (rule le_less_trans[OF _ arctan_lbound]) auto
+    by (smt (verit, ccfv_SIG) arctan field_sum_of_halves)
 next
-  have "arctan (Im z / Re z) < pi / 2"
-    by (rule arctan_ubound)
-  also have "\<dots> \<le> pi" by simp
-  finally show "arctan (Im z / Re z) \<le> pi"
-    by simp
+ show "arctan (Im z / Re z) \<le> pi"
+   by (smt (verit, best) arctan field_sum_of_halves)
 qed
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Relation between Ln and Arg2pi, and hence continuity of Arg2pi\<close>
 
-lemma Arg2pi_Ln:
-  assumes "0 < Arg2pi z" shows "Arg2pi z = Im(Ln(-z)) + pi"
-proof (cases "z = 0")
-  case True
-  with assms show ?thesis
-    by simp
-next
-  case False
-  then have "z / of_real(norm z) = exp(\<i> * of_real(Arg2pi z))"
-    using Arg2pi [of z]
-    by (metis is_Arg_def abs_norm_cancel nonzero_mult_div_cancel_left norm_of_real zero_less_norm_iff)
-  then have "- z / of_real(norm z) = exp (\<i> * (of_real (Arg2pi z) - pi))"
-    using cis_conv_exp cis_pi
-    by (auto simp: exp_diff algebra_simps)
-  then have "ln (- z / of_real(norm z)) = ln (exp (\<i> * (of_real (Arg2pi z) - pi)))"
-    by simp
-  also have "\<dots> = \<i> * (of_real(Arg2pi z) - pi)"
-    using Arg2pi [of z] assms pi_not_less_zero
-    by auto
-  finally have "Arg2pi z =  Im (Ln (- z / of_real (cmod z))) + pi"
-    by simp
-  also have "\<dots> = Im (Ln (-z) - ln (cmod z)) + pi"
-    by (metis diff_0_right minus_diff_eq zero_less_norm_iff Ln_divide_of_real False)
-  also have "\<dots> = Im (Ln (-z)) + pi"
-    by simp
-  finally show ?thesis .
-qed
+lemma Arg2pi_Ln: "0 < Arg2pi z \<Longrightarrow> Arg2pi z = Im(Ln(-z)) + pi"
+  by (smt (verit, best) Arg2pi_0 Arg2pi_exp Arg2pi_minus Arg_exp Arg_minus Im_Ln_le_pi 
+      exp_Ln mpi_less_Im_Ln neg_equal_0_iff_equal)
 
 lemma continuous_at_Arg2pi:
   assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
@@ -2433,18 +2381,14 @@
 proof -
   have *: "isCont (\<lambda>z. Im (Ln (- z)) + pi) z"
     by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+
-  have [simp]: "Im x \<noteq> 0 \<Longrightarrow> Im (Ln (- x)) + pi = Arg2pi x" for x
-    using Arg2pi_Ln  by (simp add: Arg2pi_gt_0 complex_nonneg_Reals_iff)
   consider "Re z < 0" | "Im z \<noteq> 0" using assms
     using complex_nonneg_Reals_iff not_le by blast
-  then have [simp]: "(\<lambda>z. Im (Ln (- z)) + pi) \<midarrow>z\<rightarrow> Arg2pi z"
+  then have "(\<lambda>z. Im (Ln (- z)) + pi) \<midarrow>z\<rightarrow> Arg2pi z"
     using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within)
-  show ?thesis
+  then show ?thesis
     unfolding continuous_at
-  proof (rule Lim_transform_within_open)
-    show "\<And>x. \<lbrakk>x \<in> - \<real>\<^sub>\<ge>\<^sub>0; x \<noteq> z\<rbrakk> \<Longrightarrow> Im (Ln (- x)) + pi = Arg2pi x"
-      by (auto simp add: Arg2pi_Ln [OF Arg2pi_gt_0] complex_nonneg_Reals_iff)
-  qed (use assms in auto)
+    by (metis (mono_tags, lifting) Arg2pi_Ln Arg2pi_gt_0 Compl_iff Lim_transform_within_open assms 
+        closed_nonneg_Reals_complex open_Compl)
 qed
 
 
@@ -2467,18 +2411,7 @@
 lemma Arg2pi_eq_Im_Ln:
   assumes "0 \<le> Im z" "0 < Re z"
     shows "Arg2pi z = Im (Ln z)"
-proof (cases "Im z = 0")
-  case True then show ?thesis
-    using Arg2pi_eq_0 Ln_in_Reals assms(2) complex_is_Real_iff by auto
-next
-  case False
-  then have *: "Arg2pi z > 0"
-    using Arg2pi_gt_0 complex_is_Real_iff by blast
-  then have "z \<noteq> 0"
-    by auto
-  with * assms False show ?thesis
-    by (subst Arg2pi_Ln) (auto simp: Ln_minus)
-qed
+  by (smt (verit, ccfv_SIG) Arg2pi_exp Im_Ln_pos_le assms exp_Ln pi_neq_zero zero_complex.simps(1))
 
 lemma continuous_within_upperhalf_Arg2pi:
   assumes "z \<noteq> 0"
@@ -2681,9 +2614,7 @@
     also have "\<dots> \<longleftrightarrow> ((\<lambda>z. g z ^ nat n) has_field_derivative dd') (at z within S)"
     proof (rule has_field_derivative_cong_eventually)
       show "\<forall>\<^sub>F x in at z within S. g x powr of_int n = g x ^ nat n"
-        unfolding eventually_at
-        apply (rule exI[where x=e])
-        using powr_of_int that \<open>e>0\<close> e_dist by (simp add: dist_commute)
+        unfolding eventually_at by (metis e_dist \<open>e>0\<close> dist_commute powr_of_int that)
     qed (use powr_of_int \<open>g z\<noteq>0\<close> that in simp)
     also have "\<dots>" unfolding dd'_def using gderiv that
       by (auto intro!: derivative_eq_intros)
@@ -2707,8 +2638,7 @@
     proof (rule has_field_derivative_cong_eventually)
       show "\<forall>\<^sub>F x in at z within S. g x powr of_int n = inverse (g x ^ nat (- n))"
          unfolding eventually_at
-        apply (rule exI[where x=e])
-        using powr_of_int that \<open>e>0\<close> e_dist by (simp add: dist_commute)
+         by (metis \<open>e>0\<close> e_dist dist_commute linorder_not_le powr_of_int that)
     qed (use powr_of_int \<open>g z\<noteq>0\<close> that in simp)
     also have "\<dots>"
     proof -
@@ -2759,8 +2689,8 @@
 using field_differentiable_def has_field_derivative_powr_right by blast
 
 lemma holomorphic_on_powr_right [holomorphic_intros]:
-  assumes "f holomorphic_on s"
-  shows "(\<lambda>z. w powr (f z)) holomorphic_on s"
+  assumes "f holomorphic_on S"
+  shows "(\<lambda>z. w powr (f z)) holomorphic_on S"
 proof (cases "w = 0")
   case False
   with assms show ?thesis
@@ -2769,21 +2699,9 @@
 qed simp
 
 lemma holomorphic_on_divide_gen [holomorphic_intros]:
-  assumes f: "f holomorphic_on s" and g: "g holomorphic_on s" and 0: "\<And>z z'. \<lbrakk>z \<in> s; z' \<in> s\<rbrakk> \<Longrightarrow> g z = 0 \<longleftrightarrow> g z' = 0"
-shows "(\<lambda>z. f z / g z) holomorphic_on s"
-proof (cases "\<exists>z\<in>s. g z = 0")
-  case True
-  with 0 have "g z = 0" if "z \<in> s" for z
-    using that by blast
-  then show ?thesis
-    using g holomorphic_transform by auto
-next
-  case False
-  with 0 have "g z \<noteq> 0" if "z \<in> s" for z
-    using that by blast
-  with holomorphic_on_divide show ?thesis
-    using f g by blast
-qed
+  assumes "f holomorphic_on S" "g holomorphic_on S" and "\<And>z z'. \<lbrakk>z \<in> S; z' \<in> S\<rbrakk> \<Longrightarrow> g z = 0 \<longleftrightarrow> g z' = 0"
+  shows "(\<lambda>z. f z / g z) holomorphic_on S"
+  by (metis (no_types, lifting) assms division_ring_divide_zero holomorphic_on_divide holomorphic_transform)
 
 lemma norm_powr_real_powr:
   "w \<in> \<real> \<Longrightarrow> 0 \<le> Re w \<Longrightarrow> cmod (w powr z) = Re w powr Re z"
@@ -2879,6 +2797,7 @@
 lemma tendsto_neg_powr_complex_of_nat:
   assumes "filterlim f at_top F" and "Re s < 0"
   shows   "((\<lambda>x. of_nat (f x) powr s) \<longlongrightarrow> 0) F"
+  using tendsto_neg_powr_complex_of_real [of "real o f" F s]
 proof -
   have "((\<lambda>x. of_real (real (f x)) powr s) \<longlongrightarrow> 0) F" using assms(2)
     by (intro filterlim_compose[OF _ tendsto_neg_powr_complex_of_real]
@@ -3093,29 +3012,18 @@
   also have "\<dots> = exp (Ln z / 2)"
     apply (rule csqrt_square)
     using cos_gt_zero_pi [of "(Im (Ln z) / 2)"] Im_Ln_le_pi mpi_less_Im_Ln assms
-    by (fastforce simp: Re_exp Im_exp )
+    by (fastforce simp: Re_exp Im_exp)
   finally show ?thesis using assms csqrt_square
     by simp
 qed
 
 lemma csqrt_inverse:
-  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
-    shows "csqrt (inverse z) = inverse (csqrt z)"
-proof (cases "z=0")
-  case False
-  then show ?thesis
-    using assms csqrt_exp_Ln Ln_inverse exp_minus
-    by (simp add: csqrt_exp_Ln Ln_inverse exp_minus)
-qed auto
-
-lemma cnj_csqrt:
-  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
-    shows "cnj(csqrt z) = csqrt(cnj z)"
-proof (cases "z=0")
-  case False
-  then show ?thesis
-     by (simp add: assms cnj_Ln csqrt_exp_Ln exp_cnj)
-qed auto
+  "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> csqrt (inverse z) = inverse (csqrt z)"
+  by (metis Ln_inverse csqrt_eq_0 csqrt_exp_Ln divide_minus_left exp_minus 
+      inverse_nonzero_iff_nonzero)
+
+lemma cnj_csqrt: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> cnj(csqrt z) = csqrt(cnj z)"
+  by (metis cnj_Ln complex_cnj_divide complex_cnj_numeral complex_cnj_zero_iff csqrt_eq_0 csqrt_exp_Ln exp_cnj)
 
 lemma has_field_derivative_csqrt:
   assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
@@ -3175,8 +3083,8 @@
     "continuous (at z within (\<real> \<inter> {w. 0 \<le> Re(w)})) csqrt"
 proof (cases "z \<in> \<real>\<^sub>\<le>\<^sub>0")
   case True
-  have [simp]: "Im z = 0" and 0: "Re z < 0 \<or> z = 0"
-    using True cnj.code complex_cnj_zero_iff  by (auto simp: Complex_eq complex_nonpos_Reals_iff) fastforce
+  then have [simp]: "Im z = 0" and 0: "Re z < 0 \<or> z = 0"
+    using complex_nonpos_Reals_iff complex_eq_iff by force+
   show ?thesis
     using 0
   proof
@@ -3296,11 +3204,7 @@
   have nzi: "((1 - \<i>*z) * inverse (1 + \<i>*z)) \<noteq> 0"
     using nz1 nz2 by auto
   have "Im ((1 - \<i>*z) / (1 + \<i>*z)) = 0 \<Longrightarrow> 0 < Re ((1 - \<i>*z) / (1 + \<i>*z))"
-    apply (simp add: divide_complex_def)
-    apply (simp add: divide_simps split: if_split_asm)
-    using assms
-    apply (auto simp: algebra_simps abs_square_less_1 [unfolded power2_eq_square])
-    done
+    by (simp add: Im_complex_div_lemma Re_complex_div_lemma assms cmod_eq_Im)
   then have *: "((1 - \<i>*z) / (1 + \<i>*z)) \<notin> \<real>\<^sub>\<le>\<^sub>0"
     by (auto simp add: complex_nonpos_Reals_iff)
   show "\<bar>Re(Arctan z)\<bar> < pi/2"
@@ -3483,7 +3387,7 @@
   also have "\<dots> = x"
   proof -
     have "(complex_of_real x)\<^sup>2 \<noteq> - 1"
-      by (metis diff_0_right minus_diff_eq mult_zero_left not_le of_real_1 of_real_eq_iff of_real_minus of_real_power power2_eq_square real_minus_mult_self_le zero_less_one)
+      by (smt (verit, best) Im_complex_of_real imaginary_unit.sel(2) of_real_minus power2_eq_iff power2_i)
     then show ?thesis
       by simp
   qed
@@ -3526,22 +3430,8 @@
 qed
 
 lemma arctan_inverse:
-  assumes "0 < x"
-    shows "arctan(inverse x) = pi/2 - arctan x"
-proof -
-  have "arctan(inverse x) = arctan(inverse(tan(arctan x)))"
-    by (simp add: arctan)
-  also have "\<dots> = arctan (tan (pi / 2 - arctan x))"
-    by (simp add: tan_cot)
-  also have "\<dots> = pi/2 - arctan x"
-  proof -
-    have "0 < pi - arctan x"
-    using arctan_ubound [of x] pi_gt_zero by linarith
-    with assms show ?thesis
-      by (simp add: Transcendental.arctan_tan)
-  qed
-  finally show ?thesis .
-qed
+  "0 < x \<Longrightarrow>arctan(inverse x) = pi/2 - arctan x"
+  by (smt (verit, del_insts) arctan arctan_unique tan_cot zero_less_arctan_iff)
 
 lemma arctan_add_small:
   assumes "\<bar>x * y\<bar> < 1"
@@ -3590,7 +3480,7 @@
     show "(\<lambda>k. 1 / real (k * 2 + 1) * x ^ (k * 2 + 1)) \<longlonglongrightarrow> 0 * 0"
       using assms
       by (intro tendsto_mult real_tendsto_divide_at_top)
-        (auto simp: filterlim_real_sequentially filterlim_sequentially_iff_filterlim_real
+        (auto simp: filterlim_sequentially_iff_filterlim_real
           intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially
           tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top)
   qed simp
@@ -3705,7 +3595,7 @@
   ultimately show ?thesis
     apply (simp add: sin_exp_eq Arcsin_def Arcsin_body_lemma exp_minus divide_simps)
     apply (simp add: algebra_simps)
-    apply (simp add: power2_eq_square [symmetric] algebra_simps)
+    apply (simp add: right_diff_distrib flip: power2_eq_square)
     done
 qed
 
@@ -3748,13 +3638,13 @@
   by (metis Arcsin_sin)
 
 lemma Arcsin_0 [simp]: "Arcsin 0 = 0"
-  by (metis Arcsin_sin norm_zero pi_half_gt_zero real_norm_def sin_zero zero_complex.simps(1))
+  by (simp add: Arcsin_unique)
 
 lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2"
-  by (metis Arcsin_sin Im_complex_of_real Re_complex_of_real numeral_One of_real_numeral pi_half_ge_zero real_sqrt_abs real_sqrt_pow2 real_sqrt_power sin_of_real sin_pi_half)
+  using Arcsin_unique sin_of_real_pi_half by fastforce
 
 lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)"
-  by (metis Arcsin_1 Arcsin_sin Im_complex_of_real Re_complex_of_real abs_of_nonneg of_real_minus pi_half_ge_zero power2_minus real_sqrt_abs sin_Arcsin sin_minus)
+  by (simp add: Arcsin_unique)
 
 lemma has_field_derivative_Arcsin:
   assumes "Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1"
@@ -3800,8 +3690,7 @@
   using Arcsin_range_lemma [of "-z"]  by simp
 
 lemma Arccos_body_lemma: "z + \<i> * csqrt(1 - z\<^sup>2) \<noteq> 0"
-  using Arcsin_body_lemma [of z]
-  by (metis Arcsin_body_lemma complex_i_mult_minus diff_minus_eq_add power2_minus right_minus_eq)
+  by (metis Arcsin_body_lemma complex_i_mult_minus diff_0 diff_eq_eq power2_minus)
 
 lemma Re_Arccos: "Re(Arccos z) = Im (Ln (z + \<i> * csqrt(1 - z\<^sup>2)))"
   by (simp add: Arccos_def)
@@ -3811,7 +3700,7 @@
 
 text\<open>A very tricky argument to find!\<close>
 lemma isCont_Arccos_lemma:
-  assumes eq0: "Im (z + \<i> * csqrt (1 - z\<^sup>2)) = 0" and "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"
+  assumes eq0: "Im (z + \<i> * csqrt (1 - z\<^sup>2)) = 0" and "Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1"
     shows False
 proof (cases "Im z = 0")
   case True
@@ -3987,8 +3876,7 @@
   have "Arccos w \<bullet> Arccos w \<le> pi\<^sup>2 + (cmod w)\<^sup>2"
     using Re by (simp add: dot_square_norm cmod_power2 [of "Arccos w"])
   then have "cmod (Arccos w) \<le> pi + cmod (cos (Arccos w))"
-    apply (simp add: norm_le_square)
-    by (metis dot_square_norm norm_ge_zero norm_le_square pi_ge_zero triangle_lemma)
+    by (smt (verit) Im_Arccos_bound Re_Arccos_bound cmod_le cos_Arccos)
   then show "cmod (Arccos w) \<le> pi + cmod w"
     by auto
 qed
@@ -3996,68 +3884,11 @@
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Interrelations between Arcsin and Arccos\<close>
 
-lemma cos_Arcsin_nonzero:
-  assumes "z\<^sup>2 \<noteq> 1" shows "cos(Arcsin z) \<noteq> 0"
-proof -
-  have eq: "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = z\<^sup>2 * (z\<^sup>2 - 1)"
-    by (simp add: algebra_simps)
-  have "\<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> z\<^sup>2 - 1"
-  proof
-    assume "\<i> * z * (csqrt (1 - z\<^sup>2)) = z\<^sup>2 - 1"
-    then have "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (z\<^sup>2 - 1)\<^sup>2"
-      by simp
-    then have "z\<^sup>2 * (z\<^sup>2 - 1) = (z\<^sup>2 - 1)*(z\<^sup>2 - 1)"
-      using eq power2_eq_square by auto
-    then show False
-      using assms by simp
-  qed
-  then have "1 + \<i> * z * (csqrt (1 - z * z)) \<noteq> z\<^sup>2"
-    by (metis add_minus_cancel power2_eq_square uminus_add_conv_diff)
-  then have "2*(1 + \<i> * z * (csqrt (1 - z * z))) \<noteq> 2*z\<^sup>2"  (*FIXME cancel_numeral_factor*)
-    by (metis mult_cancel_left zero_neq_numeral)
-  then have "(\<i> * z + csqrt (1 - z\<^sup>2))\<^sup>2 \<noteq> -1"
-    using assms
-    apply (simp add: power2_sum)
-    apply (simp add: power2_eq_square algebra_simps)
-    done
-  then show ?thesis
-    apply (simp add: cos_exp_eq Arcsin_def exp_minus)
-    apply (simp add: divide_simps Arcsin_body_lemma)
-    apply (metis add.commute minus_unique power2_eq_square)
-    done
-qed
-
-lemma sin_Arccos_nonzero:
-  assumes "z\<^sup>2 \<noteq> 1" shows "sin(Arccos z) \<noteq> 0"
-proof -
-  have eq: "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = -(z\<^sup>2) * (1 - z\<^sup>2)"
-    by (simp add: algebra_simps)
-  have "\<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> 1 - z\<^sup>2"
-  proof
-    assume "\<i> * z * (csqrt (1 - z\<^sup>2)) = 1 - z\<^sup>2"
-    then have "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (1 - z\<^sup>2)\<^sup>2"
-      by simp
-    then have "-(z\<^sup>2) * (1 - z\<^sup>2) = (1 - z\<^sup>2)*(1 - z\<^sup>2)"
-      using eq power2_eq_square by auto
-    then have "-(z\<^sup>2) = (1 - z\<^sup>2)"
-      using assms
-      by (metis add.commute add.right_neutral diff_add_cancel mult_right_cancel)
-    then show False
-      using assms by simp
-  qed
-  then have "z\<^sup>2 + \<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> 1"
-    by (simp add: algebra_simps)
-  then have "2*(z\<^sup>2 + \<i> * z * (csqrt (1 - z\<^sup>2))) \<noteq> 2*1"
-    by (metis mult_cancel_left2 zero_neq_numeral)  (*FIXME cancel_numeral_factor*)
-  then have "(z + \<i> * csqrt (1 - z\<^sup>2))\<^sup>2 \<noteq> 1"
-    using assms
-    by (metis Arccos_def add.commute add.left_neutral cancel_comm_monoid_add_class.diff_cancel cos_Arccos csqrt_0 mult_zero_right)
-  then show ?thesis
-    apply (simp add: sin_exp_eq Arccos_def exp_minus)
-    apply (simp add: divide_simps Arccos_body_lemma)
-    apply (simp add: power2_eq_square)
-    done
-qed
+lemma cos_Arcsin_nonzero: "z\<^sup>2 \<noteq> 1 \<Longrightarrow>cos(Arcsin z) \<noteq> 0"
+  by (metis diff_0_right power_zero_numeral sin_Arcsin sin_squared_eq)
+
+lemma sin_Arccos_nonzero: "z\<^sup>2 \<noteq> 1 \<Longrightarrow>sin(Arccos z) \<noteq> 0"
+  by (metis add.right_neutral cos_Arccos power2_eq_square power_zero_numeral sin_cos_squared_add3)
 
 lemma cos_sin_csqrt:
   assumes "0 < cos(Re z)  \<or>  cos(Re z) = 0 \<and> Im z * sin(Re z) \<le> 0"
@@ -4076,185 +3907,51 @@
 qed (use assms in \<open>auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff\<close>)
 
 lemma Arcsin_Arccos_csqrt_pos:
-    "(0 < Re z | Re z = 0 & 0 \<le> Im z) \<Longrightarrow> Arcsin z = Arccos(csqrt(1 - z\<^sup>2))"
+    "(0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z) \<Longrightarrow> Arcsin z = Arccos(csqrt(1 - z\<^sup>2))"
   by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute)
 
 lemma Arccos_Arcsin_csqrt_pos:
-    "(0 < Re z | Re z = 0 & 0 \<le> Im z) \<Longrightarrow> Arccos z = Arcsin(csqrt(1 - z\<^sup>2))"
+    "(0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z) \<Longrightarrow> Arccos z = Arcsin(csqrt(1 - z\<^sup>2))"
   by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute)
 
 lemma sin_Arccos:
-    "0 < Re z | Re z = 0 & 0 \<le> Im z \<Longrightarrow> sin(Arccos z) = csqrt(1 - z\<^sup>2)"
+    "0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z \<Longrightarrow> sin(Arccos z) = csqrt(1 - z\<^sup>2)"
   by (simp add: Arccos_Arcsin_csqrt_pos)
 
 lemma cos_Arcsin:
-    "0 < Re z | Re z = 0 & 0 \<le> Im z \<Longrightarrow> cos(Arcsin z) = csqrt(1 - z\<^sup>2)"
+    "0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z \<Longrightarrow> cos(Arcsin z) = csqrt(1 - z\<^sup>2)"
   by (simp add: Arcsin_Arccos_csqrt_pos)
 
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Relationship with Arcsin on the Real Numbers\<close>
 
-lemma Im_Arcsin_of_real:
-  assumes "\<bar>x\<bar> \<le> 1"
-    shows "Im (Arcsin (of_real x)) = 0"
-proof -
-  have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 1))"
-    by (simp add: of_real_sqrt del: csqrt_of_real_nonneg)
-  then have "cmod (\<i> * of_real x + csqrt (1 - (of_real x)\<^sup>2))^2 = 1"
-    using assms abs_square_le_1
-    by (force simp add: Complex.cmod_power2)
-  then have "cmod (\<i> * of_real x + csqrt (1 - (of_real x)\<^sup>2)) = 1"
-    by (simp add: norm_complex_def)
-  then show ?thesis
-    by (simp add: Im_Arcsin exp_minus)
-qed
+lemma of_real_arcsin: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arcsin x) = Arcsin(of_real x)"
+  by (smt (verit, best) Arcsin_sin Im_complex_of_real Re_complex_of_real arcsin sin_of_real)
+
+lemma Im_Arcsin_of_real: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> Im (Arcsin (of_real x)) = 0"
+  by (metis Im_complex_of_real of_real_arcsin)
 
 corollary\<^marker>\<open>tag unimportant\<close> Arcsin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arcsin z \<in> \<real>"
   by (metis Im_Arcsin_of_real Re_complex_of_real Reals_cases complex_is_Real_iff)
 
-lemma arcsin_eq_Re_Arcsin:
-  assumes "\<bar>x\<bar> \<le> 1"
-    shows "arcsin x = Re (Arcsin (of_real x))"
-unfolding arcsin_def
-proof (rule the_equality, safe)
-  show "- (pi / 2) \<le> Re (Arcsin (complex_of_real x))"
-    using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"]
-    by (auto simp: Complex.in_Reals_norm Re_Arcsin)
-next
-  show "Re (Arcsin (complex_of_real x)) \<le> pi / 2"
-    using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"]
-    by (auto simp: Complex.in_Reals_norm Re_Arcsin)
-next
-  show "sin (Re (Arcsin (complex_of_real x))) = x"
-    using Re_sin [of "Arcsin (of_real x)"] Arcsin_body_lemma [of "of_real x"]
-    by (simp add: Im_Arcsin_of_real assms)
-next
-  fix x'
-  assume "- (pi / 2) \<le> x'" "x' \<le> pi / 2" "x = sin x'"
-  then show "x' = Re (Arcsin (complex_of_real (sin x')))"
-    unfolding sin_of_real [symmetric]
-    by (subst Arcsin_sin) auto
-qed
-
-lemma of_real_arcsin: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arcsin x) = Arcsin(of_real x)"
-  by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0)
+lemma arcsin_eq_Re_Arcsin: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> arcsin x = Re (Arcsin (of_real x))"
+  by (metis Re_complex_of_real of_real_arcsin)
 
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Relationship with Arccos on the Real Numbers\<close>
 
-lemma Im_Arccos_of_real:
-  assumes "\<bar>x\<bar> \<le> 1"
-    shows "Im (Arccos (of_real x)) = 0"
-proof -
-  have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 1))"
-    by (simp add: of_real_sqrt del: csqrt_of_real_nonneg)
-  then have "cmod (of_real x + \<i> * csqrt (1 - (of_real x)\<^sup>2))^2 = 1"
-    using assms abs_square_le_1
-    by (force simp add: Complex.cmod_power2)
-  then have "cmod (of_real x + \<i> * csqrt (1 - (of_real x)\<^sup>2)) = 1"
-    by (simp add: norm_complex_def)
-  then show ?thesis
-    by (simp add: Im_Arccos exp_minus)
-qed
+lemma of_real_arccos: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arccos x) = Arccos(of_real x)"
+  by (smt (verit, del_insts) Arccos_unique Im_complex_of_real Re_complex_of_real arccos_lbound 
+      arccos_ubound cos_arccos_abs cos_of_real)
+
+lemma Im_Arccos_of_real: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> Im (Arccos (of_real x)) = 0"
+  by (metis Im_complex_of_real of_real_arccos)
 
 corollary\<^marker>\<open>tag unimportant\<close> Arccos_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arccos z \<in> \<real>"
-  by (metis Im_Arccos_of_real Re_complex_of_real Reals_cases complex_is_Real_iff)
-
-lemma arccos_eq_Re_Arccos:
-  assumes "\<bar>x\<bar> \<le> 1"
-    shows "arccos x = Re (Arccos (of_real x))"
-unfolding arccos_def
-proof (rule the_equality, safe)
-  show "0 \<le> Re (Arccos (complex_of_real x))"
-    using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"]
-    by (auto simp: Complex.in_Reals_norm Re_Arccos)
-next
-  show "Re (Arccos (complex_of_real x)) \<le> pi"
-    using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"]
-    by (auto simp: Complex.in_Reals_norm Re_Arccos)
-next
-  show "cos (Re (Arccos (complex_of_real x))) = x"
-    using Re_cos [of "Arccos (of_real x)"] Arccos_body_lemma [of "of_real x"]
-    by (simp add: Im_Arccos_of_real assms)
-next
-  fix x'
-  assume "0 \<le> x'" "x' \<le> pi" "x = cos x'"
-  then show "x' = Re (Arccos (complex_of_real (cos x')))"
-    unfolding cos_of_real [symmetric]
-    by (subst Arccos_cos) auto
-qed
-
-lemma of_real_arccos: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arccos x) = Arccos(of_real x)"
-  by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0)
-
-subsection\<^marker>\<open>tag unimportant\<close>\<open>Some interrelationships among the real inverse trig functions\<close>
-
-lemma arccos_arctan:
-  assumes "-1 < x" "x < 1"
-    shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))"
-proof -
-  have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0"
-  proof (rule sin_eq_0_pi)
-    show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)"
-      using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"]  arccos_bounded [of x] assms
-      by (simp add: algebra_simps)
-  next
-    show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x) < pi"
-      using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"]  arccos_bounded [of x] assms
-      by (simp add: algebra_simps)
-  next
-    show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)) = 0"
-      using assms
-      by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan
-                    power2_eq_square square_eq_1_iff)
-  qed
-  then show ?thesis
-    by simp
-qed
-
-lemma arcsin_plus_arccos:
-  assumes "-1 \<le> x" "x \<le> 1"
-    shows "arcsin x + arccos x = pi/2"
-proof -
-  have "arcsin x = pi/2 - arccos x"
-    apply (rule sin_inj_pi)
-    using assms arcsin [OF assms] arccos [OF assms]
-    by (auto simp: algebra_simps sin_diff)
-  then show ?thesis
-    by (simp add: algebra_simps)
-qed
-
-lemma arcsin_arccos_eq: "-1 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arcsin x = pi/2 - arccos x"
-  using arcsin_plus_arccos by force
-
-lemma arccos_arcsin_eq: "-1 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arccos x = pi/2 - arcsin x"
-  using arcsin_plus_arccos by force
-
-lemma arcsin_arctan: "-1 < x \<Longrightarrow> x < 1 \<Longrightarrow> arcsin x = arctan(x / sqrt(1 - x\<^sup>2))"
-  by (simp add: arccos_arctan arcsin_arccos_eq)
-
-lemma csqrt_1_diff_eq: "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 1))"
-  by ( simp add: of_real_sqrt del: csqrt_of_real_nonneg)
-
-lemma arcsin_arccos_sqrt_pos: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arcsin x = arccos(sqrt(1 - x\<^sup>2))"
-  apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos)
-  apply (subst Arcsin_Arccos_csqrt_pos)
-  apply (auto simp: power_le_one csqrt_1_diff_eq)
-  done
-
-lemma arcsin_arccos_sqrt_neg: "-1 \<le> x \<Longrightarrow> x \<le> 0 \<Longrightarrow> arcsin x = -arccos(sqrt(1 - x\<^sup>2))"
-  using arcsin_arccos_sqrt_pos [of "-x"]
-  by (simp add: arcsin_minus)
-
-lemma arccos_arcsin_sqrt_pos: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arccos x = arcsin(sqrt(1 - x\<^sup>2))"
-  apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos)
-  apply (subst Arccos_Arcsin_csqrt_pos)
-  apply (auto simp: power_le_one csqrt_1_diff_eq)
-  done
-
-lemma arccos_arcsin_sqrt_neg: "-1 \<le> x \<Longrightarrow> x \<le> 0 \<Longrightarrow> arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))"
-  using arccos_arcsin_sqrt_pos [of "-x"]
-  by (simp add: arccos_minus)
+  by (metis Im_Arccos_of_real complex_is_Real_iff of_real_Re)
+
+lemma arccos_eq_Re_Arccos: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> arccos x = Re (Arccos (of_real x))"
+  by (metis Re_complex_of_real of_real_arccos)
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Continuity results for arcsin and arccos\<close>
 
--- a/src/HOL/Analysis/Derivative.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Derivative.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -47,8 +47,7 @@
     (\<forall>i \<in> Basis. ((\<lambda>x. f x \<bullet> i) has_derivative (\<lambda>x. f' x \<bullet> i)) (at a within S))"
   apply (simp add: has_derivative_within)
   apply (subst tendsto_componentwise_iff)
-  apply (simp add: bounded_linear_componentwise_iff [symmetric] ball_conj_distrib)
-  apply (simp add: algebra_simps)
+  apply (simp add: ball_conj_distrib  inner_diff_left inner_left_distrib flip: bounded_linear_componentwise_iff)
   done
 
 lemma has_derivative_at_withinI:
@@ -1181,9 +1180,7 @@
         (\<And>z. z \<in> t \<Longrightarrow> f (g z) = z)
         \<Longrightarrow> DERIV g y :> inverse (f')"
   unfolding has_field_derivative_def
-  apply (rule has_derivative_inverse_basic)
-  apply (auto simp:  bounded_linear_mult_right)
-  done
+  by (rule has_derivative_inverse_basic) (auto simp: bounded_linear_mult_right)
 
 text \<open>Simply rewrite that based on the domain point x.\<close>
 
@@ -1205,20 +1202,13 @@
 lemma has_derivative_inverse_dieudonne:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "open S"
-    and "open (f ` S)"
-    and "continuous_on S f"
-    and "continuous_on (f ` S) g"
-    and "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
-    and "x \<in> S"
-    and "(f has_derivative f') (at x)"
-    and "bounded_linear g'"
-    and "g' \<circ> f' = id"
+    and fS: "open (f ` S)"
+    and A: "continuous_on S f" "continuous_on (f ` S) g" 
+           "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" "x \<in> S"
+    and B: "(f has_derivative f') (at x)" "bounded_linear g'" "g' \<circ> f' = id"
   shows "(g has_derivative g') (at (f x))"
-  apply (rule has_derivative_inverse_basic_x[OF assms(7-9) _ _ assms(2)])
-  using assms(3-6)
-  unfolding continuous_on_eq_continuous_at[OF assms(1)] continuous_on_eq_continuous_at[OF assms(2)]
-  apply auto
-  done
+  using A fS continuous_on_eq_continuous_at
+  by (intro has_derivative_inverse_basic_x[OF B _ _ fS]) force+
 
 text \<open>Here's the simplest way of not assuming much about g.\<close>
 
@@ -1229,19 +1219,14 @@
     and fx: "f x \<in> interior (f ` S)"
     and "continuous_on S f"
     and gf: "\<And>y. y \<in> S \<Longrightarrow> g (f y) = y"
-    and "(f has_derivative f') (at x)"
-    and "bounded_linear g'"
-    and "g' \<circ> f' = id"
+    and B: "(f has_derivative f') (at x)" "bounded_linear g'" "g' \<circ> f' = id"
   shows "(g has_derivative g') (at (f x))"
 proof -
   have *: "\<And>y. y \<in> interior (f ` S) \<Longrightarrow> f (g y) = y"
     by (metis gf image_iff interior_subset subsetCE)
   show ?thesis
-    apply (rule has_derivative_inverse_basic_x[OF assms(6-8), where T = "interior (f ` S)"])
-    apply (rule continuous_on_interior[OF _ fx])
-    apply (rule continuous_on_inv)
-    apply (simp_all add: assms *)
-    done
+    using assms * continuous_on_interior continuous_on_inv fx 
+    by (intro has_derivative_inverse_basic_x[OF B, where T = "interior (f`S)"]) blast+
 qed
 
 
@@ -1316,7 +1301,8 @@
         also have "\<dots> \<le> onorm g' * k"
           apply (rule mult_left_mono)
           using d1(2)[of u]
-          using onorm_neg[where f="\<lambda>x. f' u x - f' a x"] d u onorm_pos_le[OF bling] apply (auto simp: algebra_simps)
+          using onorm_neg[where f="\<lambda>x. f' u x - f' a x"] d u onorm_pos_le[OF bling] 
+           apply (auto simp: algebra_simps)
           done
         also have "\<dots> \<le> 1 / 2"
           unfolding k_def by auto
@@ -1498,17 +1484,16 @@
       fix x' y z :: 'a
       fix c :: real
       note lin = assms(2)[rule_format,OF \<open>x\<in>S\<close>,THEN has_derivative_bounded_linear]
-      show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'"
-        apply (rule tendsto_unique[OF trivial_limit_sequentially tog'])
+      have "(\<lambda>n. f' n x (c *\<^sub>R x')) \<longlonglongrightarrow> c *\<^sub>R g' x x'"
         unfolding lin[THEN bounded_linear.linear, THEN linear_cmul]
-        apply (intro tendsto_intros tog')
-        done
-      show "g' x (y + z) = g' x y + g' x z"
-        apply (rule tendsto_unique[OF trivial_limit_sequentially tog'])
+        by (intro tendsto_intros tog')
+      then show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'"
+        using LIMSEQ_unique tog' by blast
+      have "(\<lambda>n. f' n x (y + z)) \<longlonglongrightarrow> g' x y + g' x z"
         unfolding lin[THEN bounded_linear.linear, THEN linear_add]
-        apply (rule tendsto_add)
-        apply (rule tog')+
-        done
+        by (simp add: tendsto_add tog')
+      then show "g' x (y + z) = g' x y + g' x z"
+        using LIMSEQ_unique tog' by blast
       obtain N where N: "\<forall>h. norm (f' N x h - g' x h) \<le> 1 * norm h"
         using nle \<open>x \<in> S\<close> unfolding eventually_sequentially by (fast intro: zero_less_one)
       have "bounded_linear (f' N x)"
@@ -1621,9 +1606,8 @@
       fix n x h
       assume n: "N \<le> n" and x: "x \<in> S"
       have *: "inverse (real (Suc n)) \<le> e"
-        apply (rule order_trans[OF _ N[THEN less_imp_le]])
-        using n apply (auto simp add: field_simps)
-        done
+        using n N
+        by (smt (verit, best) le_imp_inverse_le of_nat_0_less_iff of_nat_Suc of_nat_le_iff zero_less_Suc)
       show "norm (f' n x h - g' x h) \<le> e * norm h"
         by (meson "*" mult_right_mono norm_ge_zero order.trans x f')
     qed
@@ -1822,12 +1806,19 @@
 lemma has_vector_derivative_cong_ev:
   assumes *: "eventually (\<lambda>x. x \<in> S \<longrightarrow> f x = g x) (nhds x)" "f x = g x"
   shows "(f has_vector_derivative f') (at x within S) = (g has_vector_derivative f') (at x within S)"
+proof (cases "at x within S = bot")
+  case True
+  then show ?thesis   
+    by (simp add: has_derivative_def has_vector_derivative_def)
+next
+  case False
+  then show ?thesis
   unfolding has_vector_derivative_def has_derivative_def
   using *
-  apply (cases "at x within S \<noteq> bot")
   apply (intro refl conj_cong filterlim_cong)
   apply (auto simp: Lim_ident_at eventually_at_filter elim: eventually_mono)
   done
+qed
 
 lemma vector_derivative_cong_eq:
   assumes "eventually (\<lambda>x. x \<in> A \<longrightarrow> f x = g x) (nhds x)" "x = y" "A = B" "x \<in> A"
@@ -1900,18 +1891,15 @@
 lemma vector_derivative_scaleR_at [simp]:
     "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk>
    \<Longrightarrow> vector_derivative (\<lambda>x. f x *\<^sub>R g x) (at a) = f a *\<^sub>R vector_derivative g (at a) + vector_derivative f (at a) *\<^sub>R g a"
-apply (rule vector_derivative_at)
-apply (rule has_vector_derivative_scaleR)
-apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs)
-done
+  apply (intro vector_derivative_at has_vector_derivative_scaleR)
+   apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs)
+  done
 
 lemma vector_derivative_within_cbox:
   assumes ab: "a < b" "x \<in> cbox a b"
   assumes f: "(f has_vector_derivative f') (at x within cbox a b)"
   shows "vector_derivative f (at x within cbox a b) = f'"
-  by (intro vector_derivative_unique_within_closed_interval[OF ab _ f]
-            vector_derivative_works[THEN iffD1] differentiableI_vector)
-     fact
+  by (metis assms box_real(2) f islimpt_Icc trivial_limit_within vector_derivative_within)
 
 lemma vector_derivative_within_closed_interval:
   fixes f::"real \<Rightarrow> 'a::euclidean_space"
@@ -2328,8 +2316,8 @@
 lemma vector_derivative_chain_at_general:
   assumes "f differentiable at x" "g field_differentiable at (f x)"
   shows "vector_derivative (g \<circ> f) (at x) = vector_derivative f (at x) * deriv g (f x)"
-  apply (rule vector_derivative_at [OF field_vector_diff_chain_at])
-  using assms vector_derivative_works by (auto simp: field_differentiable_derivI)
+  using assms field_differentiable_derivI field_vector_diff_chain_at 
+      vector_derivative_at vector_derivative_works by blast
 
 lemma deriv_chain:
   "f field_differentiable at x \<Longrightarrow> g field_differentiable at (f x)
@@ -2409,10 +2397,14 @@
   by (auto intro!: DERIV_imp_deriv derivative_intros)
 
 lemma deriv_compose_linear:
-  "f field_differentiable at (c * z) \<Longrightarrow> deriv (\<lambda>w. f (c * w)) z = c * deriv f (c * z)"
-apply (rule DERIV_imp_deriv)
-  unfolding DERIV_deriv_iff_field_differentiable [symmetric]
-  by (metis (full_types) DERIV_chain2 DERIV_cmult_Id mult.commute)
+  assumes "f field_differentiable at (c * z)"
+  shows "deriv (\<lambda>w. f (c * w)) z = c * deriv f (c * z)"
+proof -
+  have "deriv (\<lambda>a. f (c * a)) z = deriv f (c * z) * c"
+    using assms by (simp add: DERIV_chain2 DERIV_deriv_iff_field_differentiable DERIV_imp_deriv)
+  then show ?thesis
+    by simp
+qed
 
 
 lemma nonzero_deriv_nonconstant:
@@ -2624,8 +2616,7 @@
             norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R
               norm ((x', y') - (x, y)))
             < e'"
-    apply eventually_elim
-  proof safe
+  proof (eventually_elim, safe)
     fix x' y'
     have "norm (f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) \<le>
         norm (f x' y' - f x' y - fy x' y (y' - y)) +
@@ -2727,12 +2718,10 @@
   shows "((\<lambda>x. if x \<in> S then f x else g x) has_vector_derivative
     (if x \<in> S then f' x else g' x)) (at x within u)"
   unfolding has_vector_derivative_def assms
-  using x_in
-  apply (intro has_derivative_If_within_closures[where ?f' = "\<lambda>x a. a *\<^sub>R f' x" and ?g' = "\<lambda>x a. a *\<^sub>R g' x",
-        THEN has_derivative_eq_rhs])
-  subgoal by (rule f'[unfolded has_vector_derivative_def]; assumption)
-  subgoal by (rule g'[unfolded has_vector_derivative_def]; assumption)
-  by (auto simp: assms)
+  using x_in f' g'
+  by (intro has_derivative_If_within_closures[where ?f' = "\<lambda>x a. a *\<^sub>R f' x" and ?g' = "\<lambda>x a. a *\<^sub>R g' x",
+        THEN has_derivative_eq_rhs]; force simp: assms has_vector_derivative_def)
+
 
 subsection\<^marker>\<open>tag important\<close>\<open>The Inverse Function Theorem\<close>
 
@@ -3100,17 +3089,12 @@
 lemma piecewise_differentiable_on_subset:
     "f piecewise_differentiable_on S \<Longrightarrow> T \<le> S \<Longrightarrow> f piecewise_differentiable_on T"
   using continuous_on_subset
-  unfolding piecewise_differentiable_on_def
-  apply safe
-  apply (blast elim: continuous_on_subset)
-  by (meson Diff_iff differentiable_within_subset subsetCE)
+  by (smt (verit) Diff_iff differentiable_within_subset in_mono piecewise_differentiable_on_def)
 
 lemma differentiable_on_imp_piecewise_differentiable:
   fixes a:: "'a::{linorder_topology,real_normed_vector}"
   shows "f differentiable_on {a..b} \<Longrightarrow> f piecewise_differentiable_on {a..b}"
-  apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on)
-  apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def)
-  done
+  using differentiable_imp_continuous_on differentiable_onD piecewise_differentiable_on_def by fastforce
 
 lemma differentiable_imp_piecewise_differentiable:
     "(\<And>x. x \<in> S \<Longrightarrow> f differentiable (at x within S))
--- a/src/HOL/Analysis/Determinants.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Determinants.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -1025,8 +1025,8 @@
     then show ?thesis
       using \<open>A *v axis k 1 = a\<close> that by auto
   next
-    from ex_card[OF 2] obtain h i::'n where "h \<noteq> i"
-      by (auto simp add: eval_nat_numeral card_Suc_eq)
+    from obtain_subset_with_card_n[OF 2] obtain h i::'n where "h \<noteq> i"
+      by (fastforce simp add: eval_nat_numeral card_Suc_eq)
     then obtain j where "j \<noteq> k"
       by (metis (full_types))
     let ?TA = "transpose A"
--- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -33,13 +33,13 @@
   by (simp add: sphere_def)
 
 lemma ball_trivial [simp]: "ball x 0 = {}"
-  by (simp add: ball_def)
+  by auto
 
 lemma cball_trivial [simp]: "cball x 0 = {x}"
-  by (simp add: cball_def)
+  by auto
 
 lemma sphere_trivial [simp]: "sphere x 0 = {x}"
-  by (simp add: sphere_def)
+  by auto
 
 lemma disjoint_ballI: "dist x y \<ge> r+s \<Longrightarrow> ball x r \<inter> ball y s = {}"
   using dist_triangle_less_add not_le by fastforce
@@ -196,9 +196,7 @@
   by (simp add: interior_open)
 
 lemma cball_eq_empty [simp]: "cball x e = {} \<longleftrightarrow> e < 0"
-  apply (simp add: set_eq_iff not_le)
-  apply (metis zero_le_dist dist_self order_less_le_trans)
-  done
+  by (smt (verit, best) Diff_empty ball_eq_empty cball_diff_sphere centre_in_ball centre_in_cball sphere_empty)
 
 lemma cball_empty [simp]: "e < 0 \<Longrightarrow> cball x e = {}"
   by simp
@@ -215,8 +213,7 @@
   using ball_divide_subset one_le_numeral by blast
 
 lemma cball_divide_subset: "d \<ge> 1 \<Longrightarrow> cball x (e/d) \<subseteq> cball x e"
-  apply (cases "e < 0", simp add: field_split_simps)
-  by (metis div_by_1 frac_le less_numeral_extra(1) not_le order_refl subset_cball)
+  by (smt (verit, best) cball_empty div_by_1 frac_le subset_cball zero_le_divide_iff)
 
 lemma cball_divide_subset_numeral: "cball x (e / numeral w) \<subseteq> cball x e"
   using cball_divide_subset one_le_numeral by blast
@@ -292,13 +289,7 @@
 
 lemma limpt_of_limpts: "x islimpt {y. y islimpt S} \<Longrightarrow> x islimpt S"
   for x :: "'a::metric_space"
-  apply (clarsimp simp add: islimpt_approachable)
-  apply (drule_tac x="e/2" in spec)
-  apply (auto simp: simp del: less_divide_eq_numeral1)
-  apply (drule_tac x="dist x' x" in spec)
-  apply (auto simp del: less_divide_eq_numeral1)
-  apply metric
-  done
+  by (metis islimpt_def islimpt_eq_acc_point mem_Collect_eq)
 
 lemma closed_limpts:  "closed {x::'a::metric_space. x islimpt S}"
   using closed_limpt limpt_of_limpts by blast
@@ -308,14 +299,12 @@
   by (auto simp: closure_def islimpt_Un dest: limpt_of_limpts)
 
 lemma islimpt_eq_infinite_ball: "x islimpt S \<longleftrightarrow> (\<forall>e>0. infinite(S \<inter> ball x e))"
-  apply (simp add: islimpt_eq_acc_point, safe)
-   apply (metis Int_commute open_ball centre_in_ball)
-  by (metis open_contains_ball Int_mono finite_subset inf_commute subset_refl)
+  unfolding islimpt_eq_acc_point
+  by (metis open_ball Int_commute Int_mono finite_subset open_contains_ball_eq subset_eq)
 
 lemma islimpt_eq_infinite_cball: "x islimpt S \<longleftrightarrow> (\<forall>e>0. infinite(S \<inter> cball x e))"
-  apply (simp add: islimpt_eq_infinite_ball, safe)
-   apply (meson Int_mono ball_subset_cball finite_subset order_refl)
-  by (metis open_ball centre_in_ball finite_Int inf.absorb_iff2 inf_assoc open_contains_cball_eq)
+  unfolding islimpt_eq_infinite_ball
+  by (metis open_ball ball_subset_cball centre_in_ball finite_Int inf.absorb_iff2 inf_assoc open_contains_cball_eq)
 
 
 subsection \<open>Perfect Metric Spaces\<close>
@@ -327,17 +316,11 @@
 lemma cball_eq_sing:
   fixes x :: "'a::{metric_space,perfect_space}"
   shows "cball x e = {x} \<longleftrightarrow> e = 0"
-proof (rule linorder_cases)
-  assume e: "0 < e"
-  obtain a where "a \<noteq> x" "dist a x < e"
-    using perfect_choose_dist [OF e] by auto
-  then have "a \<noteq> x" "dist x a \<le> e"
-    by (auto simp: dist_commute)
-  with e show ?thesis by (auto simp: set_eq_iff)
-qed auto
-
-
-subsection \<open>?\<close>
+  by (smt (verit, best) open_ball ball_eq_empty ball_subset_cball cball_empty cball_trivial 
+      not_open_singleton subset_singleton_iff)
+
+
+subsection \<open>Finite and discrete\<close>
 
 lemma finite_ball_include:
   fixes a :: "'a::metric_space"
@@ -364,19 +347,7 @@
   then obtain d where "d > 0" and d: "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> d \<le> dist a x"
     by blast
   show ?case
-  proof (cases "x = a")
-    case True
-    with \<open>d > 0 \<close>d show ?thesis by auto
-  next
-    case False
-    let ?d = "min d (dist a x)"
-    from False \<open>d > 0\<close> have dp: "?d > 0"
-      by auto
-    from d have d': "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> ?d \<le> dist a x"
-      by auto
-    with dp False show ?thesis
-      by (metis insert_iff le_less min_less_iff_conj not_less)
-  qed
+    by (smt (verit, del_insts) dist_pos_lt insert.IH insert_iff)
 qed (auto intro: zero_less_one)
 
 lemma discrete_imp_closed:
@@ -388,7 +359,7 @@
   have False if C: "\<And>e. e>0 \<Longrightarrow> \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" for x
   proof -
     from e have e2: "e/2 > 0" by arith
-    from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y \<noteq> x" "dist y x < e/2"
+    from C[OF e2] obtain y where y: "y \<in> S" "y \<noteq> x" "dist y x < e/2"
       by blast
     from e2 y(2) have mp: "min (e/2) (dist x y) > 0"
       by simp
@@ -438,9 +409,7 @@
 corollary Lim_withinI [intro?]:
   assumes "\<And>e. e > 0 \<Longrightarrow> \<exists>d>0. \<forall>x \<in> S. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l \<le> e"
   shows "(f \<longlongrightarrow> l) (at a within S)"
-  apply (simp add: Lim_within, clarify)
-  apply (rule ex_forward [OF assms [OF half_gt_zero]], auto)
-  done
+  unfolding Lim_within by (smt (verit, ccfv_SIG) assms zero_less_dist_iff)
 
 proposition Lim_at: "(f \<longlongrightarrow> l) (at a) \<longleftrightarrow>
     (\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d  \<longrightarrow> dist (f x) l < e)"
@@ -450,11 +419,7 @@
   fixes a :: "'a::metric_space" and l :: "'b::metric_space"
   shows "\<lbrakk>(f \<longlongrightarrow> l) (at a within S); eventually (\<lambda>x. x \<in> S \<longleftrightarrow> x \<in> T) (at a)\<rbrakk>
          \<Longrightarrow> (f \<longlongrightarrow> l) (at a within T)"
-apply (clarsimp simp: eventually_at Lim_within)
-apply (drule_tac x=e in spec, clarify)
-apply (rename_tac k)
-apply (rule_tac x="min d k" in exI, simp)
-done
+  by (simp add: eventually_at Lim_within) (smt (verit, best))
 
 text \<open>Another limit point characterization.\<close>
 
@@ -471,7 +436,7 @@
     by metis
   define f where "f \<equiv> rec_nat (y 1) (\<lambda>n fn. y (min (inverse(2 ^ (Suc n))) (dist fn x)))"
   have [simp]: "f 0 = y 1"
-               "f(Suc n) = y (min (inverse(2 ^ (Suc n))) (dist (f n) x))" for n
+            and fSuc: "f(Suc n) = y (min (inverse(2 ^ (Suc n))) (dist (f n) x))" for n
     by (simp_all add: f_def)
   have f: "f n \<in> S \<and> (f n \<noteq> x) \<and> dist (f n) x < inverse(2 ^ n)" for n
   proof (induction n)
@@ -479,11 +444,10 @@
       by (simp add: y)
   next
     case (Suc n) then show ?case
-      apply (auto simp: y)
-      by (metis half_gt_zero_iff inverse_positive_iff_positive less_divide_eq_numeral1(1) min_less_iff_conj y zero_less_dist_iff zero_less_numeral zero_less_power)
+      by (smt (verit, best) fSuc dist_pos_lt inverse_positive_iff_positive y zero_less_power)
   qed
   show ?rhs
-  proof (rule_tac x=f in exI, intro conjI allI)
+  proof (intro exI conjI allI)
     show "\<And>n. f n \<in> S - {x}"
       using f by blast
     have "dist (f n) x < dist (f m) x" if "m < n" for m n
@@ -497,7 +461,7 @@
       proof cases
         assume "m < n"
         have "dist (f(Suc n)) x = dist (y (min (inverse(2 ^ (Suc n))) (dist (f n) x))) x"
-          by simp
+          by (simp add: fSuc)
         also have "\<dots> < dist (f n) x"
           by (metis dist_pos_lt f min.strict_order_iff min_less_iff_conj y)
         also have "\<dots> < dist (f m) x"
@@ -505,17 +469,16 @@
         finally show ?thesis .
       next
         assume "m = n" then show ?case
-          by simp (metis dist_pos_lt f half_gt_zero_iff inverse_positive_iff_positive min_less_iff_conj y zero_less_numeral zero_less_power)
+          by (smt (verit, best) dist_pos_lt f fSuc y)
       qed
     qed
     then show "inj f"
       by (metis less_irrefl linorder_injI)
-    show "f \<longlonglongrightarrow> x"
-      apply (rule tendstoI)
-      apply (rule_tac c="nat (ceiling(1/e))" in eventually_sequentiallyI)
+    have "\<And>e n. \<lbrakk>0 < e; nat \<lceil>1 / e\<rceil> \<le> n\<rbrakk> \<Longrightarrow> dist (f n) x < e"
       apply (rule less_trans [OF f [THEN conjunct2, THEN conjunct2]])
-      apply (simp add: field_simps)
-      by (meson le_less_trans mult_less_cancel_left not_le of_nat_less_two_power)
+      by (simp add: divide_simps order_le_less_trans)
+    then show "f \<longlonglongrightarrow> x"
+      using lim_sequentially by blast
   qed
 next
   assume ?rhs
@@ -548,70 +511,49 @@
   assumes nondecF: "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y"
   shows "continuous (at_right a) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f (a + d) - f a < e)"
   apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le)
-  apply (intro all_cong ex_cong, safe)
-  apply (erule_tac x="a + d" in allE, simp)
-  apply (simp add: nondecF field_simps)
-  apply (drule nondecF, simp)
-  done
+  apply (intro all_cong ex_cong)
+  by (smt (verit, best) nondecF)
 
 lemma continuous_at_left_real_increasing:
   assumes nondecF: "\<And> x y. x \<le> y \<Longrightarrow> f x \<le> ((f y) :: real)"
-  shows "(continuous (at_left (a :: real)) f) = (\<forall>e > 0. \<exists>delta > 0. f a - f (a - delta) < e)"
+  shows "(continuous (at_left (a :: real)) f) \<longleftrightarrow> (\<forall>e > 0. \<exists>delta > 0. f a - f (a - delta) < e)"
   apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le)
-  apply (intro all_cong ex_cong, safe)
-  apply (erule_tac x="a - d" in allE, simp)
-  apply (simp add: nondecF field_simps)
-  apply (cut_tac x="a - d" and y=x in nondecF, simp_all)
-  done
+  apply (intro all_cong ex_cong)
+  by (smt (verit) nondecF)
 
 text\<open>Versions in terms of open balls.\<close>
 
 lemma continuous_within_ball:
-  "continuous (at x within s) f \<longleftrightarrow>
-    (\<forall>e > 0. \<exists>d > 0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)"
+  "continuous (at x within S) f \<longleftrightarrow>
+    (\<forall>e > 0. \<exists>d > 0. f ` (ball x d \<inter> S) \<subseteq> ball (f x) e)"
   (is "?lhs = ?rhs")
 proof
   assume ?lhs
   {
     fix e :: real
     assume "e > 0"
-    then obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
+    then obtain d where "d>0" and d: "\<forall>y\<in>S. 0 < dist y x \<and> dist y x < d \<longrightarrow> dist (f y) (f x) < e"
       using \<open>?lhs\<close>[unfolded continuous_within Lim_within] by auto
-    {
-      fix y
-      assume "y \<in> f ` (ball x d \<inter> s)"
-      then have "y \<in> ball (f x) e"
-        using d(2)
-        using \<open>e > 0\<close>
-        by (auto simp: dist_commute)
+    { fix y
+      assume "y \<in> f ` (ball x d \<inter> S)" then have "y \<in> ball (f x) e"
+        using d \<open>e > 0\<close> by (auto simp: dist_commute)
     }
-    then have "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e"
-      using \<open>d > 0\<close>
-      unfolding subset_eq ball_def by (auto simp: dist_commute)
+    then have "\<exists>d>0. f ` (ball x d \<inter> S) \<subseteq> ball (f x) e"
+      using \<open>d > 0\<close> by blast
   }
   then show ?rhs by auto
 next
   assume ?rhs
   then show ?lhs
-    unfolding continuous_within Lim_within ball_def subset_eq
-    apply (auto simp: dist_commute)
-    apply (erule_tac x=e in allE, auto)
-    done
+    apply (simp add: continuous_within Lim_within ball_def subset_eq)
+    by (metis (mono_tags, lifting) Int_iff dist_commute mem_Collect_eq)
 qed
 
 lemma continuous_at_ball:
-  "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs
-    unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
-    by (metis dist_commute dist_pos_lt dist_self)
-next
-  assume ?rhs
-  then show ?lhs
-    unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
-    by (metis dist_commute)
-qed
+  "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)"
+  apply (simp add: continuous_at Lim_at subset_eq Ball_def Bex_def image_iff)
+  by (smt (verit, ccfv_threshold) dist_commute dist_self)
+
 
 text\<open>Define setwise continuity in terms of limits within the set.\<close>
 
@@ -622,16 +564,14 @@
   by (metis dist_pos_lt dist_self)
 
 lemma continuous_within_E:
-  assumes "continuous (at x within s) f" "e>0"
-  obtains d where "d>0"  "\<And>x'. \<lbrakk>x'\<in> s; dist x' x \<le> d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
-  using assms apply (simp add: continuous_within_eps_delta)
-  apply (drule spec [of _ e], clarify)
-  apply (rule_tac d="d/2" in that, auto)
-  done
+  assumes "continuous (at x within S) f" "e>0"
+  obtains d where "d>0"  "\<And>x'. \<lbrakk>x'\<in> S; dist x' x \<le> d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
+  using assms unfolding continuous_within_eps_delta
+  by (metis dense order_le_less_trans)
 
 lemma continuous_onI [intro?]:
-  assumes "\<And>x e. \<lbrakk>e > 0; x \<in> s\<rbrakk> \<Longrightarrow> \<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) \<le> e"
-  shows "continuous_on s f"
+  assumes "\<And>x e. \<lbrakk>e > 0; x \<in> S\<rbrakk> \<Longrightarrow> \<exists>d>0. \<forall>x'\<in>S. dist x' x < d \<longrightarrow> dist (f x') (f x) \<le> e"
+  shows "continuous_on S f"
 apply (simp add: continuous_on_iff, clarify)
 apply (rule ex_forward [OF assms [OF half_gt_zero]], auto)
 done
@@ -642,10 +582,7 @@
     assumes "continuous_on s f" "x\<in>s" "e>0"
     obtains d where "d>0"  "\<And>x'. \<lbrakk>x' \<in> s; dist x' x \<le> d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
   using assms
-  apply (simp add: continuous_on_iff)
-  apply (elim ballE allE)
-  apply (auto intro: that [where d="d/2" for d])
-  done
+  unfolding continuous_on_iff by (metis dense order_le_less_trans)
 
 text\<open>The usual transformation theorems.\<close>
 
@@ -657,8 +594,7 @@
     and "\<And>x'. \<lbrakk>x' \<in> s; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'"
   shows "continuous (at x within s) g"
   using assms
-  unfolding continuous_within
-  by (force intro: Lim_transform_within)
+  unfolding continuous_within by (force intro: Lim_transform_within)
 
 
 subsection \<open>Closure and Limit Characterization\<close>
@@ -666,9 +602,7 @@
 lemma closure_approachable:
   fixes S :: "'a::metric_space set"
   shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)"
-  apply (auto simp: closure_def islimpt_approachable)
-  apply (metis dist_self)
-  done
+  using dist_self by (force simp: closure_def islimpt_approachable)
 
 lemma closure_approachable_le:
   fixes S :: "'a::metric_space set"
@@ -693,14 +627,13 @@
 proof -
   have *: "\<forall>x\<in>S. Inf S \<le> x"
     using cInf_lower[of _ S] assms by metis
-  {
-    fix e :: real
+  { fix e :: real
     assume "e > 0"
     then have "Inf S < Inf S + e" by simp
     with assms obtain x where "x \<in> S" "x < Inf S + e"
-      by (subst (asm) cInf_less_iff) auto
+      using cInf_lessD by blast
     with * have "\<exists>x\<in>S. dist x (Inf S) < e"
-      by (intro bexI[of _ x]) (auto simp: dist_real_def)
+      using dist_real_def by force
   }
   then show ?thesis unfolding closure_approachable by auto
 qed
@@ -717,9 +650,9 @@
     assume "e > 0"
     then have "Sup S - e < Sup S" by simp
     with assms obtain x where "x \<in> S" "Sup S - e < x"
-      by (subst (asm) less_cSup_iff) auto
+      using less_cSupE by blast
     with * have "\<exists>x\<in>S. dist x (Sup S) < e"
-      by (intro bexI[of _ x]) (auto simp: dist_real_def)
+      using dist_real_def by force
   }
   then show ?thesis unfolding closure_approachable by auto
 qed
@@ -730,8 +663,7 @@
 proof
   show ?rhs if ?lhs
   proof -
-    {
-      fix e :: real
+    { fix e :: real
       assume "e > 0"
       then obtain y where "y \<in> S - {x}" and "dist y x < e"
         using \<open>?lhs\<close> not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
@@ -744,8 +676,7 @@
   qed
   show ?lhs if ?rhs
   proof -
-    {
-      fix e :: real
+    { fix e :: real
       assume "e > 0"
       then obtain y where "y \<in> S \<inter> ball x e - {x}"
         using \<open>?rhs\<close> by blast
@@ -804,8 +735,7 @@
 proof -
   from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a"
     unfolding bounded_def by auto
-  {
-    fix y
+  { fix y
     assume "y \<in> closure S"
     then obtain f where f: "\<forall>n. f n \<in> S"  "(f \<longlongrightarrow> y) sequentially"
       unfolding closure_sequential by auto
@@ -1299,10 +1229,7 @@
     and xy: "\<forall>u\<in>S. \<forall>v\<in>S. dist u v \<le> dist x y"
     using compact_sup_maxdistance[OF assms] by auto
   then have "diameter S \<le> dist x y"
-    unfolding diameter_def
-    apply clarsimp
-    apply (rule cSUP_least, fast+)
-    done
+    unfolding diameter_def by (force intro!: cSUP_least)
   then show ?thesis
     by (metis b diameter_bounded_bound order_antisym xys)
 qed
@@ -1332,12 +1259,12 @@
   assumes "bounded S"
   shows "diameter(closure S) = diameter S"
 proof (rule order_antisym)
-  have "False" if "diameter S < diameter (closure S)"
+  have "False" if d_less_d: "diameter S < diameter (closure S)"
   proof -
     define d where "d = diameter(closure S) - diameter(S)"
     have "d > 0"
       using that by (simp add: d_def)
-    then have "diameter(closure(S)) - d / 2 < diameter(closure(S))"
+    then have dle: "diameter(closure(S)) - d / 2 < diameter(closure(S))"
       by simp
     have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2"
       by (simp add: d_def field_split_simps)
@@ -1346,14 +1273,13 @@
     moreover have "0 \<le> diameter S"
       using assms diameter_ge_0 by blast
     ultimately obtain x y where "x \<in> closure S" "y \<in> closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y"
-      using diameter_bounded(2) [OF bocl, rule_format, of "diameter(closure(S)) - d / 2"] \<open>d > 0\<close> d_def by auto
+      by (smt (verit) dle d_less_d d_def dd diameter_lower_bounded)
     then obtain x' y' where x'y': "x' \<in> S" "dist x' x < d/4" "y' \<in> S" "dist y' y < d/4"
-      using closure_approachable
-      by (metis \<open>0 < d\<close> zero_less_divide_iff zero_less_numeral)
+      by (metis \<open>0 < d\<close> zero_less_divide_iff zero_less_numeral closure_approachable)
     then have "dist x' y' \<le> diameter S"
       using assms diameter_bounded_bound by blast
     with x'y' have "dist x y \<le> d / 4 + diameter S + d / 4"
-      by (meson add_mono_thms_linordered_semiring(1) dist_triangle dist_triangle3 less_eq_real_def order_trans)
+      by (meson add_mono dist_triangle dist_triangle3 less_eq_real_def order_trans)
     then show ?thesis
       using xy d_def by linarith
   qed
@@ -1421,8 +1347,7 @@
       then have "T \<subseteq> ball y \<delta>"
         using \<open>y \<in> T\<close> dia diameter_bounded_bound by fastforce
       then show ?thesis
-        apply (rule_tac x=C in bexI)
-        using \<open>ball y \<delta> \<subseteq> C\<close> \<open>C \<in> \<C>\<close> by auto
+        by (meson \<open>C \<in> \<C>\<close> \<open>ball y \<delta> \<subseteq> C\<close> subset_eq)
     qed
   qed
 qed
@@ -1490,8 +1415,7 @@
   obtain r :: "nat \<Rightarrow> nat" where r: "strict_mono r" "monoseq (f \<circ> r)"
     unfolding comp_def by (metis seq_monosub)
   then have "Bseq (f \<circ> r)"
-    unfolding Bseq_eq_bounded using f
-    by (metis BseqI' bounded_iff comp_apply rangeI)
+    unfolding Bseq_eq_bounded by (metis f BseqI' bounded_iff comp_apply rangeI)
   with r show "\<exists>l r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
     using Bseq_monoseq_convergent[of "f \<circ> r"] by (auto simp: convergent_def)
 qed
@@ -1525,32 +1449,31 @@
       using k
       by (rule bounded_proj)
     obtain l1::"'a" and r1 where r1: "strict_mono r1"
-      and lr1: "\<forall>e > 0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially"
-      using insert(3) using insert(4) by auto
+      and lr1: "\<forall>e > 0. \<forall>\<^sub>F n in sequentially. \<forall>i\<in>d. dist (f (r1 n) proj i) (l1 proj i) < e"
+      using insert by auto
     have f': "\<forall>n. f (r1 n) proj k \<in> (\<lambda>x. x proj k) ` range f"
       by simp
     have "bounded (range (\<lambda>i. f (r1 i) proj k))"
       by (metis (lifting) bounded_subset f' image_subsetI s')
-    then obtain l2 r2 where r2:"strict_mono r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) proj k) \<longlongrightarrow> l2) sequentially"
+    then obtain l2 r2 where r2: "strict_mono r2" and lr2: "(\<lambda>i. f (r1 (r2 i)) proj k) \<longlonglongrightarrow> l2"
       using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) proj k"]
       by (auto simp: o_def)
     define r where "r = r1 \<circ> r2"
-    have r:"strict_mono r"
+    have r: "strict_mono r"
       using r1 and r2 unfolding r_def o_def strict_mono_def by auto
     moreover
     define l where "l = unproj (\<lambda>i. if i = k then l2 else l1 proj i)"
-    {
-      fix e::real
+    { fix e::real
       assume "e > 0"
-      from lr1 \<open>e > 0\<close> have N1: "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially"
+      from lr1 \<open>e > 0\<close> have N1: "\<forall>\<^sub>F n in sequentially. \<forall>i\<in>d. dist (f (r1 n) proj i) (l1 proj i) < e"
         by blast
-      from lr2 \<open>e > 0\<close> have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) proj k) l2 < e) sequentially"
+      from lr2 \<open>e > 0\<close> have N2: "\<forall>\<^sub>F n in sequentially. dist (f (r1 (r2 n)) proj k) l2 < e"
         by (rule tendstoD)
-      from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) proj i) (l1 proj i) < e) sequentially"
+      from r2 N1 have N1': "\<forall>\<^sub>F n in sequentially. \<forall>i\<in>d. dist (f (r1 (r2 n)) proj i) (l1 proj i) < e"
         by (rule eventually_subseq)
-      have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) proj i) (l proj i) < e) sequentially"
+      have "\<forall>\<^sub>F n in sequentially. \<forall>i\<in>insert k d. dist (f (r n) proj i) (l proj i) < e"
         using N1' N2
-        by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def proj_unproj)
+        by eventually_elim (use insert.prems in \<open>auto simp: l_def r_def o_def proj_unproj\<close>)
     }
     ultimately show ?case by auto
   qed
@@ -1576,7 +1499,7 @@
     using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast
   from f have s2: "bounded (range (snd \<circ> f \<circ> r1))"
     by (auto simp: image_comp intro: bounded_snd bounded_subset)
-  obtain l2 r2 where r2: "strict_mono r2" and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) \<longlongrightarrow> l2) sequentially"
+  obtain l2 r2 where r2: "strict_mono r2" and l2: "(\<lambda>n. snd (f (r1 (r2 n)))) \<longlonglongrightarrow> l2"
     using bounded_imp_convergent_subsequence [OF s2]
     unfolding o_def by fast
   have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) \<longlongrightarrow> l1) sequentially"
@@ -1619,13 +1542,9 @@
       fix e :: real
       assume "e > 0"
       from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2"
-        unfolding cauchy_def
-        using \<open>e > 0\<close>
-        apply (erule_tac x="e/2" in allE, auto)
-        done
-      from lr(3)[unfolded lim_sequentially, THEN spec[where x="e/2"]]
-      obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2"
-        using \<open>e > 0\<close> by auto
+        unfolding cauchy_def using \<open>e > 0\<close> by (meson half_gt_zero)
+      then obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2"
+        by (metis dist_self lim_sequentially lr(3))
       {
         fix n :: nat
         assume n: "n \<ge> max N M"
@@ -1756,58 +1675,28 @@
 instance heine_borel < complete_space
 proof
   fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
-  then have "bounded (range f)"
-    by (rule cauchy_imp_bounded)
-  then have "compact (closure (range f))"
-    unfolding compact_eq_bounded_closed by auto
-  then have "complete (closure (range f))"
-    by (rule compact_imp_complete)
-  moreover have "\<forall>n. f n \<in> closure (range f)"
-    using closure_subset [of "range f"] by auto
-  ultimately have "\<exists>l\<in>closure (range f). (f \<longlongrightarrow> l) sequentially"
-    using \<open>Cauchy f\<close> unfolding complete_def by auto
   then show "convergent f"
-    unfolding convergent_def by auto
+    unfolding convergent_def
+    using Cauchy_converges_subseq cauchy_imp_bounded bounded_imp_convergent_subsequence by blast
 qed
 
 lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)"
 proof (rule completeI)
   fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
-  then have "convergent f" by (rule Cauchy_convergent)
-  then show "\<exists>l\<in>UNIV. f \<longlonglongrightarrow> l" unfolding convergent_def by simp
+  then show "\<exists>l\<in>UNIV. f \<longlonglongrightarrow> l" 
+    using Cauchy_convergent convergent_def by blast
 qed
 
 lemma complete_imp_closed:
   fixes S :: "'a::metric_space set"
-  assumes "complete S"
-  shows "closed S"
-proof (unfold closed_sequential_limits, clarify)
-  fix f x assume "\<forall>n. f n \<in> S" and "f \<longlonglongrightarrow> x"
-  from \<open>f \<longlonglongrightarrow> x\<close> have "Cauchy f"
-    by (rule LIMSEQ_imp_Cauchy)
-  with \<open>complete S\<close> and \<open>\<forall>n. f n \<in> S\<close> obtain l where "l \<in> S" and "f \<longlonglongrightarrow> l"
-    by (rule completeE)
-  from \<open>f \<longlonglongrightarrow> x\<close> and \<open>f \<longlonglongrightarrow> l\<close> have "x = l"
-    by (rule LIMSEQ_unique)
-  with \<open>l \<in> S\<close> show "x \<in> S"
-    by simp
-qed
+  shows "complete S \<Longrightarrow> closed S"
+  by (metis LIMSEQ_imp_Cauchy LIMSEQ_unique closed_sequential_limits completeE)
 
 lemma complete_Int_closed:
   fixes S :: "'a::metric_space set"
   assumes "complete S" and "closed t"
   shows "complete (S \<inter> t)"
-proof (rule completeI)
-  fix f assume "\<forall>n. f n \<in> S \<inter> t" and "Cauchy f"
-  then have "\<forall>n. f n \<in> S" and "\<forall>n. f n \<in> t"
-    by simp_all
-  from \<open>complete S\<close> obtain l where "l \<in> S" and "f \<longlonglongrightarrow> l"
-    using \<open>\<forall>n. f n \<in> S\<close> and \<open>Cauchy f\<close> by (rule completeE)
-  from \<open>closed t\<close> and \<open>\<forall>n. f n \<in> t\<close> and \<open>f \<longlonglongrightarrow> l\<close> have "l \<in> t"
-    by (rule closed_sequentially)
-  with \<open>l \<in> S\<close> and \<open>f \<longlonglongrightarrow> l\<close> show "\<exists>l\<in>S \<inter> t. f \<longlonglongrightarrow> l"
-    by fast
-qed
+  by (metis Int_iff assms closed_sequentially completeE completeI)
 
 lemma complete_closed_subset:
   fixes S :: "'a::metric_space set"
@@ -1845,8 +1734,7 @@
 lemma continuous_closed_imp_Cauchy_continuous:
   fixes S :: "('a::complete_space) set"
   shows "\<lbrakk>continuous_on S f; closed S; Cauchy \<sigma>; \<And>n. (\<sigma> n) \<in> S\<rbrakk> \<Longrightarrow> Cauchy(f \<circ> \<sigma>)"
-  apply (simp add: complete_eq_closed [symmetric] continuous_on_sequentially)
-  by (meson LIMSEQ_imp_Cauchy complete_def)
+  by (meson LIMSEQ_imp_Cauchy completeE complete_eq_closed continuous_on_sequentially)
 
 lemma banach_fix_type:
   fixes f::"'a::complete_space\<Rightarrow>'a"
@@ -1866,15 +1754,13 @@
   assumes "closed S"
       and T: "T \<in> \<F>" "bounded T"
       and clof: "\<And>T. T \<in> \<F> \<Longrightarrow> closed T"
-      and none: "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> S \<inter> \<Inter>\<F>' \<noteq> {}"
+      and "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> S \<inter> \<Inter>\<F>' \<noteq> {}"
     shows "S \<inter> \<Inter>\<F> \<noteq> {}"
 proof -
   have "compact (S \<inter> T)"
     using \<open>closed S\<close> clof compact_eq_bounded_closed T by blast
   then have "(S \<inter> T) \<inter> \<Inter>\<F> \<noteq> {}"
-    apply (rule compact_imp_fip)
-     apply (simp add: clof)
-    by (metis Int_assoc complete_lattice_class.Inf_insert finite_insert insert_subset none \<open>T \<in> \<F>\<close>)
+    by (smt (verit, best) Inf_insert Int_assoc assms compact_imp_fip finite_insert insert_subset)
   then show ?thesis by blast
 qed
 
@@ -1884,33 +1770,27 @@
    "\<lbrakk>closed S; \<And>T. T \<in> \<F> \<Longrightarrow> compact T;
      \<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> S \<inter> \<Inter>\<F>' \<noteq> {}\<rbrakk>
         \<Longrightarrow> S \<inter> \<Inter>\<F> \<noteq> {}"
-by (metis Inf_greatest closed_imp_fip compact_eq_bounded_closed empty_subsetI finite.emptyI inf.orderE)
+  by (metis closed_imp_fip compact_eq_bounded_closed equals0I finite.emptyI order.refl)
 
 lemma closed_fip_Heine_Borel:
   fixes \<F> :: "'a::heine_borel set set"
-  assumes "closed S" "T \<in> \<F>" "bounded T"
+  assumes "T \<in> \<F>" "bounded T"
       and "\<And>T. T \<in> \<F> \<Longrightarrow> closed T"
       and "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> \<Inter>\<F>' \<noteq> {}"
     shows "\<Inter>\<F> \<noteq> {}"
-proof -
-  have "UNIV \<inter> \<Inter>\<F> \<noteq> {}"
-    using assms closed_imp_fip [OF closed_UNIV] by auto
-  then show ?thesis by simp
-qed
+  using closed_imp_fip [OF closed_UNIV] assms by auto
 
 lemma compact_fip_Heine_Borel:
   fixes \<F> :: "'a::heine_borel set set"
   assumes clof: "\<And>T. T \<in> \<F> \<Longrightarrow> compact T"
       and none: "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> \<Inter>\<F>' \<noteq> {}"
     shows "\<Inter>\<F> \<noteq> {}"
-by (metis InterI all_not_in_conv clof closed_fip_Heine_Borel compact_eq_bounded_closed none)
+  by (metis InterI clof closed_fip_Heine_Borel compact_eq_bounded_closed equals0D none)
 
 lemma compact_sequence_with_limit:
   fixes f :: "nat \<Rightarrow> 'a::heine_borel"
-  shows "(f \<longlongrightarrow> l) sequentially \<Longrightarrow> compact (insert l (range f))"
-apply (simp add: compact_eq_bounded_closed, auto)
-apply (simp add: convergent_imp_bounded)
-by (simp add: closed_limpt islimpt_insert sequence_unique_limpt)
+  shows "f \<longlonglongrightarrow> l \<Longrightarrow> compact (insert l (range f))"
+  by (simp add: closed_limpt compact_eq_bounded_closed convergent_imp_bounded islimpt_insert sequence_unique_limpt)
 
 
 subsection \<open>Properties of Balls and Spheres\<close>
@@ -1918,21 +1798,18 @@
 lemma compact_cball[simp]:
   fixes x :: "'a::heine_borel"
   shows "compact (cball x e)"
-  using compact_eq_bounded_closed bounded_cball closed_cball
-  by blast
+  using compact_eq_bounded_closed bounded_cball closed_cball by blast
 
 lemma compact_frontier_bounded[intro]:
   fixes S :: "'a::heine_borel set"
   shows "bounded S \<Longrightarrow> compact (frontier S)"
   unfolding frontier_def
-  using compact_eq_bounded_closed
-  by blast
+  using compact_eq_bounded_closed by blast
 
 lemma compact_frontier[intro]:
   fixes S :: "'a::heine_borel set"
   shows "compact S \<Longrightarrow> compact (frontier S)"
-  using compact_eq_bounded_closed compact_frontier_bounded
-  by blast
+  using compact_eq_bounded_closed compact_frontier_bounded by blast
 
 
 subsection \<open>Distance from a Set\<close>
@@ -2054,9 +1931,7 @@
     with infdist_nonneg[of x A] have "infdist x A > 0"
       by auto
     then have "ball x (infdist x A) \<inter> closure A = {}"
-      apply auto
-      apply (metis \<open>x \<in> closure A\<close> closure_approachable dist_commute infdist_le not_less)
-      done
+      by (smt (verit, best) \<open>x \<in> closure A\<close> closure_approachableD infdist_le)
     then have "x \<notin> closure A"
       by (metis \<open>0 < infdist x A\<close> centre_in_ball disjoint_iff_not_equal)
     then show False using \<open>x \<in> closure A\<close> by simp
@@ -2065,18 +1940,15 @@
   assume x: "infdist x A = 0"
   then obtain a where "a \<in> A"
     by atomize_elim (metis all_not_in_conv assms)
-  show "x \<in> closure A"
-    unfolding closure_approachable
-    apply safe
-  proof (rule ccontr)
-    fix e :: real
-    assume "e > 0"
-    assume "\<not> (\<exists>y\<in>A. dist y x < e)"
-    then have "infdist x A \<ge> e" using \<open>a \<in> A\<close>
-      unfolding infdist_def
+  have False if "e > 0" "\<not> (\<exists>y\<in>A. dist y x < e)" for e
+  proof -
+    have "infdist x A \<ge> e" using \<open>a \<in> A\<close>
+      unfolding infdist_def using that
       by (force simp: dist_commute intro: cINF_greatest)
     with x \<open>e > 0\<close> show False by auto
   qed
+  then show "x \<in> closure A"
+    using closure_approachable by blast
 qed
 
 lemma in_closed_iff_infdist_zero:
@@ -2084,14 +1956,14 @@
   shows "x \<in> A \<longleftrightarrow> infdist x A = 0"
 proof -
   have "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
-    by (rule in_closure_iff_infdist_zero) fact
+    by (simp add: \<open>A \<noteq> {}\<close> in_closure_iff_infdist_zero)
   with assms show ?thesis by simp
 qed
 
 lemma infdist_pos_not_in_closed:
   assumes "closed S" "S \<noteq> {}" "x \<notin> S"
   shows "infdist x S > 0"
-using in_closed_iff_infdist_zero[OF assms(1) assms(2), of x] assms(3) infdist_nonneg le_less by fastforce
+  by (smt (verit, best) assms in_closed_iff_infdist_zero infdist_nonneg)
 
 lemma
   infdist_attains_inf:
@@ -2103,10 +1975,9 @@
   have "bdd_below (dist y ` X)"
     by auto
   from distance_attains_inf[OF assms, of y]
-  obtain x where INF: "x \<in> X" "\<And>z. z \<in> X \<Longrightarrow> dist y x \<le> dist y z" by auto
-  have "infdist y X = dist y x"
-    by (auto simp: infdist_def assms
-      intro!: antisym cINF_lower[OF _ \<open>x \<in> X\<close>] cINF_greatest[OF assms(2) INF(2)])
+  obtain x where "x \<in> X" "\<And>z. z \<in> X \<Longrightarrow> dist y x \<le> dist y z" by auto
+  then have "infdist y X = dist y x"
+    by (metis antisym assms(2) cINF_greatest infdist_def infdist_le)
   with \<open>x \<in> X\<close> show ?thesis ..
 qed
 
@@ -2119,13 +1990,9 @@
   consider "S = {}" | "T = {}" | "S \<noteq> {} \<and> T \<noteq> {}" by auto
   then show "\<exists>U V. open U \<and> open V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> U \<inter> V = {}"
   proof (cases)
-    case 1
-    show ?thesis
-      apply (rule exI[of _ "{}"], rule exI[of _ UNIV]) using 1 by auto
+    case 1 then show ?thesis by blast
   next
-    case 2
-    show ?thesis
-      apply (rule exI[of _ UNIV], rule exI[of _ "{}"]) using 2 by auto
+    case 2 then show ?thesis by blast
   next
     case 3
     define U where "U = (\<Union>x\<in>S. ball x ((infdist x T)/2))"
@@ -2154,7 +2021,7 @@
     then have E: "U \<inter> V = {}"
       unfolding U_def V_def by auto
     show ?thesis
-      apply (rule exI[of _ U], rule exI[of _ V]) using A B C D E by auto
+      using A B C D E by blast
   qed
 qed
 
@@ -2218,57 +2085,42 @@
 subsection \<open>Separation between Points and Sets\<close>
 
 proposition separate_point_closed:
-  fixes s :: "'a::heine_borel set"
-  assumes "closed s" and "a \<notin> s"
-  shows "\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x"
-proof (cases "s = {}")
-  case True
-  then show ?thesis by(auto intro!: exI[where x=1])
-next
-  case False
-  from assms obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y"
-    using \<open>s \<noteq> {}\<close> by (blast intro: distance_attains_inf [of s a])
-  with \<open>x\<in>s\<close> show ?thesis using dist_pos_lt[of a x] and\<open>a \<notin> s\<close>
-    by blast
-qed
+  fixes S :: "'a::heine_borel set"
+  assumes "closed S" and "a \<notin> S"
+  shows "\<exists>d>0. \<forall>x\<in>S. d \<le> dist a x"
+  by (metis assms distance_attains_inf empty_iff gt_ex zero_less_dist_iff)
 
 proposition separate_compact_closed:
-  fixes s t :: "'a::heine_borel set"
-  assumes "compact s"
-    and t: "closed t" "s \<inter> t = {}"
-  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
+  fixes S T :: "'a::heine_borel set"
+  assumes "compact S"
+    and T: "closed T" "S \<inter> T = {}"
+  shows "\<exists>d>0. \<forall>x\<in>S. \<forall>y\<in>T. d \<le> dist x y"
 proof cases
-  assume "s \<noteq> {} \<and> t \<noteq> {}"
-  then have "s \<noteq> {}" "t \<noteq> {}" by auto
-  let ?inf = "\<lambda>x. infdist x t"
-  have "continuous_on s ?inf"
+  assume "S \<noteq> {} \<and> T \<noteq> {}"
+  then have "S \<noteq> {}" "T \<noteq> {}" by auto
+  let ?inf = "\<lambda>x. infdist x T"
+  have "continuous_on S ?inf"
     by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident)
-  then obtain x where x: "x \<in> s" "\<forall>y\<in>s. ?inf x \<le> ?inf y"
-    using continuous_attains_inf[OF \<open>compact s\<close> \<open>s \<noteq> {}\<close>] by auto
+  then obtain x where x: "x \<in> S" "\<forall>y\<in>S. ?inf x \<le> ?inf y"
+    using continuous_attains_inf[OF \<open>compact S\<close> \<open>S \<noteq> {}\<close>] by auto
   then have "0 < ?inf x"
-    using t \<open>t \<noteq> {}\<close> in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg)
-  moreover have "\<forall>x'\<in>s. \<forall>y\<in>t. ?inf x \<le> dist x' y"
+    using T \<open>T \<noteq> {}\<close> in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg)
+  moreover have "\<forall>x'\<in>S. \<forall>y\<in>T. ?inf x \<le> dist x' y"
     using x by (auto intro: order_trans infdist_le)
   ultimately show ?thesis by auto
 qed (auto intro!: exI[of _ 1])
 
 proposition separate_closed_compact:
-  fixes s t :: "'a::heine_borel set"
-  assumes "closed s"
-    and "compact t"
-    and "s \<inter> t = {}"
-  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
-proof -
-  have *: "t \<inter> s = {}"
-    using assms(3) by auto
-  show ?thesis
-    using separate_compact_closed[OF assms(2,1) *] by (force simp: dist_commute)
-qed
+  fixes S T :: "'a::heine_borel set"
+  assumes S: "closed S"
+    and T: "compact T"
+    and dis: "S \<inter> T = {}"
+  shows "\<exists>d>0. \<forall>x\<in>S. \<forall>y\<in>T. d \<le> dist x y"
+  by (metis separate_compact_closed[OF T S] dis dist_commute inf_commute)
 
 proposition compact_in_open_separated:
   fixes A::"'a::heine_borel set"
-  assumes "A \<noteq> {}"
-  assumes "compact A"
+  assumes A: "A \<noteq> {}" "compact A"
   assumes "open B"
   assumes "A \<subseteq> B"
   obtains e where "e > 0" "{x. infdist x A \<le> e} \<subseteq> B"
@@ -2287,13 +2139,8 @@
     assume "\<nexists>e. 0 < e \<and> {x. infdist x A \<le> e} \<subseteq> B"
     with \<open>d > 0\<close> obtain x where x: "infdist x A \<le> d" "x \<notin> B"
       by auto
-    from assms have "closed A" "A \<noteq> {}" by (auto simp: compact_eq_bounded_closed)
-    from infdist_attains_inf[OF this]
-    obtain y where y: "y \<in> A" "infdist x A = dist x y"
-      by auto
-    have "dist x y \<le> d" using x y by simp
-    also have "\<dots> < dist x y" using y d x by auto
-    finally show False by simp
+    then show False
+      by (metis A compact_eq_bounded_closed infdist_attains_inf x d linorder_not_less)
   qed
 qed
 
@@ -2303,8 +2150,8 @@
 lemma uniformly_continuous_onE:
   assumes "uniformly_continuous_on s f" "0 < e"
   obtains d where "d>0" "\<And>x x'. \<lbrakk>x\<in>s; x'\<in>s; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
-using assms
-by (auto simp: uniformly_continuous_on_def)
+  using assms
+  by (auto simp: uniformly_continuous_on_def)
 
 lemma uniformly_continuous_on_sequentially:
   "uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and>
@@ -2320,20 +2167,11 @@
       fix e :: real
       assume "e > 0"
       then obtain d where "d > 0" and d: "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
-        using \<open>?lhs\<close>[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
+        by (metis \<open>?lhs\<close> uniformly_continuous_onE)
       obtain N where N: "\<forall>n\<ge>N. dist (x n) (y n) < d"
         using xy[unfolded lim_sequentially dist_norm] and \<open>d>0\<close> by auto
-      {
-        fix n
-        assume "n\<ge>N"
-        then have "dist (f (x n)) (f (y n)) < e"
-          using N[THEN spec[where x=n]]
-          using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]]
-          using x and y
-          by (simp add: dist_commute)
-      }
       then have "\<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e"
-        by auto
+        using d x y by blast
     }
     then have "((\<lambda>n. dist (f(x n)) (f(y n))) \<longlongrightarrow> 0) sequentially"
       unfolding lim_sequentially and dist_real_def by auto
@@ -2348,8 +2186,7 @@
     then obtain fa where fa:
       "\<forall>x. 0 < x \<longrightarrow> fst (fa x) \<in> s \<and> snd (fa x) \<in> s \<and> dist (fst (fa x)) (snd (fa x)) < x \<and> \<not> dist (f (fst (fa x))) (f (snd (fa x))) < e"
       using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"]
-      unfolding Bex_def
-      by (auto simp: dist_commute)
+      by (auto simp: Bex_def dist_commute)
     define x where "x n = fst (fa (inverse (real n + 1)))" for n
     define y where "y n = snd (fa (inverse (real n + 1)))" for n
     have xyn: "\<forall>n. x n \<in> s \<and> y n \<in> s"
@@ -2362,17 +2199,8 @@
       assume "e > 0"
       then obtain N :: nat where "N \<noteq> 0" and N: "0 < inverse (real N) \<and> inverse (real N) < e"
         unfolding real_arch_inverse[of e] by auto
-      {
-        fix n :: nat
-        assume "n \<ge> N"
-        then have "inverse (real n + 1) < inverse (real N)"
-          using of_nat_0_le_iff and \<open>N\<noteq>0\<close> by auto
-        also have "\<dots> < e" using N by auto
-        finally have "inverse (real n + 1) < e" by auto
-        then have "dist (x n) (y n) < e"
-          using xy0[THEN spec[where x=n]] by auto
-      }
-      then have "\<exists>N. \<forall>n\<ge>N. dist (x n) (y n) < e" by auto
+      then have "\<exists>N. \<forall>n\<ge>N. dist (x n) (y n) < e"
+        by (smt (verit, ccfv_SIG) inverse_le_imp_le nat_le_real_less of_nat_le_0_iff xy0) 
     }
     then have "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e"
       using \<open>?rhs\<close>[THEN spec[where x=x], THEN spec[where x=y]] and xyn
@@ -2415,8 +2243,7 @@
     then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))"
       by simp
     also have "... \<le> 1 / real (Suc (max N1 N2))"
-      apply (simp add: field_split_simps del: max.bounded_iff)
-      using \<open>strict_mono r\<close> seq_suble by blast
+      by (meson Suc_le_mono \<open>strict_mono r\<close> inverse_of_nat_le nat.discI seq_suble)
     also have "... \<le> 1 / real (Suc N2)"
       by (simp add: field_simps)
     also have "... < e/2"
@@ -2447,7 +2274,7 @@
     using cont by metis
   let ?\<G> = "((\<lambda>x. ball x (d x (e/2))) ` S)"
   have Ssub: "S \<subseteq> \<Union> ?\<G>"
-    by clarsimp (metis d_pos \<open>0 < e\<close> dist_self half_gt_zero_iff)
+    using \<open>0 < e\<close> d_pos by auto
   then obtain k where "0 < k" and k: "\<And>x. x \<in> S \<Longrightarrow> \<exists>G \<in> ?\<G>. ball x k \<subseteq> G"
     by (rule Heine_Borel_lemma [OF \<open>compact S\<close>]) auto
   moreover have "dist (f v) (f u) < e" if "f \<in> \<F>" "u \<in> S" "v \<in> S" "dist v u < k" for f u v
@@ -2637,10 +2464,7 @@
   from uc have cont_f: "continuous_on X f"
     by (simp add: uniformly_continuous_imp_continuous)
   obtain y where y: "(f \<longlongrightarrow> y x) (at x within X)" if "x \<in> closure X" for x
-    apply atomize_elim
-    apply (rule choice)
-    using uniformly_continuous_on_extension_at_closure[OF assms]
-    by metis
+    using uniformly_continuous_on_extension_at_closure[OF assms] by meson
   let ?g = "\<lambda>x. if x \<in> X then f x else y x"
 
   have "uniformly_continuous_on (closure X) ?g"
@@ -2733,14 +2557,11 @@
 proof
   assume ?lhs
   then show ?rhs
-    apply (simp add: openin_open)
-    apply (metis Int_commute Int_mono inf.cobounded2 open_contains_ball order_refl subsetCE)
-    done
+    by (metis IntD2 Int_commute Int_lower1 Int_mono inf.idem openE openin_open)
 next
   assume ?rhs
   then show ?lhs
-    apply (simp add: openin_euclidean_subtopology_iff)
-    by (metis (no_types) Int_iff dist_commute inf.absorb_iff2 mem_ball)
+    by (smt (verit) open_ball Int_commute Int_iff centre_in_ball in_mono openin_open_Int openin_subopen)
 qed
 
 lemma openin_contains_cball:
@@ -2804,11 +2625,8 @@
           "\<And>e. e>0 \<Longrightarrow> \<exists>n. \<forall>x\<in>S n. \<forall>y\<in>S n. dist x y < e"
   obtains a where "\<And>n. a \<in> S n"
 proof -
-  have "\<forall>n. \<exists>x. x \<in> S n"
-    using assms(2) by auto
-  then have "\<exists>t. \<forall>n. t n \<in> S n"
-    using choice[of "\<lambda>n x. x \<in> S n"] by auto
-  then obtain t where t: "\<forall>n. t n \<in> S n" by auto
+  obtain t where t: "\<forall>n. t n \<in> S n"
+    by (meson assms(2) equals0I)
   {
     fix e :: real
     assume "e > 0"
@@ -3012,15 +2830,7 @@
 lemma continuous_at_real_range:
   fixes f :: "'a::real_normed_vector \<Rightarrow> real"
   shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'. norm(x' - x) < d --> \<bar>f x' - f x\<bar> < e)"
-  unfolding continuous_at
-  unfolding Lim_at
-  unfolding dist_norm
-  apply auto
-  apply (erule_tac x=e in allE, auto)
-  apply (rule_tac x=d in exI, auto)
-  apply (erule_tac x=x' in allE, auto)
-  apply (erule_tac x=e in allE, auto)
-  done
+  by (metis (mono_tags, opaque_lifting) LIM_eq continuous_within norm_eq_zero real_norm_def right_minus_eq)
 
 lemma continuous_on_real_range:
   fixes f :: "'a::real_normed_vector \<Rightarrow> real"
@@ -3091,12 +2901,7 @@
 lemma le_setdist_iff:
         "d \<le> setdist S T \<longleftrightarrow>
         (\<forall>x \<in> S. \<forall>y \<in> T. d \<le> dist x y) \<and> (S = {} \<or> T = {} \<longrightarrow> d \<le> 0)"
-  apply (cases "S = {} \<or> T = {}")
-  apply (force simp add: setdist_def)
-  apply (intro iffI conjI)
-  using setdist_le_dist apply fastforce
-  apply (auto simp: intro: le_setdistI)
-  done
+  by (smt (verit) le_setdistI setdist_def setdist_le_dist)
 
 lemma setdist_ltE:
   assumes "setdist S T < b" "S \<noteq> {}" "T \<noteq> {}"
@@ -3105,11 +2910,7 @@
 by (auto simp: not_le [symmetric] le_setdist_iff)
 
 lemma setdist_refl: "setdist S S = 0"
-  apply (cases "S = {}")
-  apply (force simp add: setdist_def)
-  apply (rule antisym [OF _ setdist_pos_le])
-  apply (metis all_not_in_conv dist_self setdist_le_dist)
-  done
+  by (metis dist_eq_0_iff ex_in_conv order_antisym setdist_def setdist_le_dist setdist_pos_le)
 
 lemma setdist_sym: "setdist S T = setdist T S"
   by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf])
@@ -3121,10 +2922,8 @@
 next
   case False
   then have "\<And>x. x \<in> S \<Longrightarrow> setdist S T - dist x a \<le> setdist {a} T"
-    apply (intro le_setdistI)
-    apply (simp_all add: algebra_simps)
-    apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist])
-    done
+    using  dist_self dist_triangle3 empty_not_insert le_setdist_iff setdist_le_dist singleton_iff
+    by (smt (verit, best) dist_self dist_triangle3 empty_not_insert le_setdist_iff setdist_le_dist singleton_iff)
   then have "setdist S T - setdist {a} T \<le> setdist S {a}"
     using False by (fastforce intro: le_setdistI)
   then show ?thesis
@@ -3148,9 +2947,7 @@
   by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz])
 
 lemma setdist_subset_right: "\<lbrakk>T \<noteq> {}; T \<subseteq> u\<rbrakk> \<Longrightarrow> setdist S u \<le> setdist S T"
-  apply (cases "S = {} \<or> u = {}", force)
-  apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono)
-  done
+  by (smt (verit, best) in_mono le_setdist_iff)
 
 lemma setdist_subset_left: "\<lbrakk>S \<noteq> {}; S \<subseteq> T\<rbrakk> \<Longrightarrow> setdist T u \<le> setdist S u"
   by (metis setdist_subset_right setdist_sym)
@@ -3166,12 +2963,9 @@
       by (auto simp: continuous_intros dist_norm)
     then have *: "\<And>x. x \<in> closure S \<Longrightarrow> setdist S T \<le> dist x y"
       by (fast intro: setdist_le_dist \<open>y \<in> T\<close> continuous_ge_on_closure)
-  } note * = this
+  } then
   show ?thesis
-    apply (rule antisym)
-     using False closure_subset apply (blast intro: setdist_subset_left)
-    using False * apply (force intro!: le_setdistI)
-    done
+    by (metis False antisym closure_eq_empty closure_subset le_setdist_iff setdist_subset_left)
 qed
 
 lemma setdist_closure_2 [simp]: "setdist T (closure S) = setdist T S"
@@ -3199,14 +2993,13 @@
     have "Inf {dist x y |x y. x \<in> A \<and> y \<in> B} \<le> Inf (dist x ` B)"
       if  "b \<in> B" "a \<in> A" "x \<in> A" for x 
     proof -
-      have *: "\<And>b'. b' \<in> B \<Longrightarrow> Inf {dist x y |x y. x \<in> A \<and> y \<in> B} \<le> dist x b'"
+      have "\<And>b'. b' \<in> B \<Longrightarrow> Inf {dist x y |x y. x \<in> A \<and> y \<in> B} \<le> dist x b'"
         by (metis (mono_tags, lifting) ex_in_conv setdist_def setdist_le_dist that(3))
-      show ?thesis
-        using that by (subst conditionally_complete_lattice_class.le_cInf_iff) (auto simp: *)+
+      then show ?thesis
+        by (smt (verit) cINF_greatest ex_in_conv \<open>b \<in> B\<close>)
     qed
     then show "Inf {dist x y |x y. x \<in> A \<and> y \<in> B} \<le> (INF x\<in>A. Inf (dist x ` B))"
-      using that
-      by (subst conditionally_complete_lattice_class.le_cInf_iff) (auto simp: bdd_below_def)
+      by (metis (mono_tags, lifting) cINF_greatest emptyE that)
   next
     have *: "\<And>x y. \<lbrakk>b \<in> B; a \<in> A; x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> \<exists>a\<in>A. Inf (dist a ` B) \<le> dist x y"
       by (meson bdd_below_image_dist cINF_lower)
@@ -3225,8 +3018,7 @@
   shows "infdist x B \<le> infdist x A"
   by (simp add: assms infdist_eq_setdist setdist_subset_right)
 
-lemma infdist_singleton [simp]:
-  "infdist x {y} = dist x y"
+lemma infdist_singleton [simp]: "infdist x {y} = dist x y"
   by (simp add: infdist_eq_setdist)
 
 proposition setdist_attains_inf:
@@ -3247,12 +3039,7 @@
     also have "\<dots> = infdist y A"
     proof (rule order_antisym)
       show "(INF y\<in>B. infdist y A) \<le> infdist y A"
-      proof (rule cInf_lower)
-        show "infdist y A \<in> (\<lambda>y. infdist y A) ` B"
-          using \<open>y \<in> B\<close> by blast
-        show "bdd_below ((\<lambda>y. infdist y A) ` B)"
-          by (meson bdd_belowI2 infdist_nonneg)
-      qed
+        by (meson \<open>y \<in> B\<close> bdd_belowI2 cInf_lower image_eqI infdist_nonneg)
     next
       show "infdist y A \<le> (INF y\<in>B. infdist y A)"
         by (simp add: \<open>B \<noteq> {}\<close> cINF_greatest min)
--- a/src/HOL/Analysis/Extended_Real_Limits.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Extended_Real_Limits.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -175,45 +175,12 @@
 lemma ereal_open_closed:
   fixes S :: "ereal set"
   shows "open S \<and> closed S \<longleftrightarrow> S = {} \<or> S = UNIV"
-proof -
-  {
-    assume lhs: "open S \<and> closed S"
-    {
-      assume "-\<infinity> \<notin> S"
-      then have "S = {}"
-        using lhs ereal_open_closed_aux by auto
-    }
-    moreover
-    {
-      assume "-\<infinity> \<in> S"
-      then have "- S = {}"
-        using lhs ereal_open_closed_aux[of "-S"] by auto
-    }
-    ultimately have "S = {} \<or> S = UNIV"
-      by auto
-  }
-  then show ?thesis
-    by auto
-qed
+  using ereal_open_closed_aux open_closed by auto
 
 lemma ereal_open_atLeast:
   fixes x :: ereal
   shows "open {x..} \<longleftrightarrow> x = -\<infinity>"
-proof
-  assume "x = -\<infinity>"
-  then have "{x..} = UNIV"
-    by auto
-  then show "open {x..}"
-    by auto
-next
-  assume "open {x..}"
-  then have "open {x..} \<and> closed {x..}"
-    by auto
-  then have "{x..} = UNIV"
-    unfolding ereal_open_closed by auto
-  then show "x = -\<infinity>"
-    by (simp add: bot_ereal_def atLeast_eq_UNIV_iff)
-qed
+  by (metis atLeast_eq_UNIV_iff bot_ereal_def closed_atLeast ereal_open_closed not_Ici_eq_empty)
 
 lemma mono_closed_real:
   fixes S :: "real set"
@@ -227,10 +194,7 @@
       then have *: "\<forall>x\<in>S. Inf S \<le> x"
         using cInf_lower[of _ S] ex by (metis bdd_below_def)
       then have "Inf S \<in> S"
-        apply (subst closed_contains_Inf)
-        using ex \<open>S \<noteq> {}\<close> \<open>closed S\<close>
-        apply auto
-        done
+        by (meson \<open>S \<noteq> {}\<close> assms(2) bdd_belowI closed_contains_Inf)
       then have "\<forall>x. Inf S \<le> x \<longleftrightarrow> x \<in> S"
         using mono[rule_format, of "Inf S"] *
         by auto
@@ -267,33 +231,18 @@
     and "closed S"
   shows "\<exists>a. S = {x. a \<le> ereal x}"
 proof -
-  {
-    assume "S = {}"
-    then have ?thesis
-      apply (rule_tac x=PInfty in exI)
-      apply auto
-      done
-  }
-  moreover
-  {
-    assume "S = UNIV"
-    then have ?thesis
-      apply (rule_tac x="-\<infinity>" in exI)
-      apply auto
-      done
-  }
-  moreover
-  {
-    assume "\<exists>a. S = {a ..}"
-    then obtain a where "S = {a ..}"
-      by auto
-    then have ?thesis
-      apply (rule_tac x="ereal a" in exI)
-      apply auto
-      done
-  }
-  ultimately show ?thesis
-    using mono_closed_real[of S] assms by auto
+  consider "S = {} \<or> S = UNIV" | "(\<exists>a. S = {a..})"
+    using assms(2) mono mono_closed_real by blast
+  then show ?thesis
+  proof cases
+    case 1
+    then show ?thesis
+      by (meson PInfty_neq_ereal(1) UNIV_eq_I bot.extremum empty_Collect_eq ereal_infty_less_eq(1) mem_Collect_eq)
+  next
+    case 2
+    then show ?thesis
+      by (metis atLeast_iff ereal_less_eq(3) mem_Collect_eq subsetI subset_antisym)
+  qed
 qed
 
 lemma Liminf_within:
@@ -349,10 +298,7 @@
 lemma min_Liminf_at:
   fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_linorder"
   shows "min (f x) (Liminf (at x) f) = (SUP e\<in>{0<..}. INF y\<in>ball x e. f y)"
-  apply (simp add: inf_min [symmetric] Liminf_at)
-  apply (subst inf_commute)
-  apply (subst SUP_inf)
-  apply auto
+  apply (simp add: inf_min [symmetric] Liminf_at inf_commute [of "f x"] SUP_inf)
   apply (metis (no_types, lifting) INF_insert centre_in_ball greaterThan_iff image_cong inf_commute insert_Diff)
   done
 
@@ -362,9 +308,11 @@
 lemma sum_constant_ereal:
   fixes a::ereal
   shows "(\<Sum>i\<in>I. a) = a * card I"
-apply (cases "finite I", induct set: finite, simp_all)
-apply (cases a, auto, metis (no_types, opaque_lifting) add.commute mult.commute semiring_normalization_rules(3))
-done
+proof (induction I rule: infinite_finite_induct)
+  case (insert i I)
+  then show ?case
+    by (simp add: ereal_right_distrib flip: plus_ereal.simps)
+qed auto
 
 lemma real_lim_then_eventually_real:
   assumes "(u \<longlongrightarrow> ereal l) F"
@@ -381,13 +329,13 @@
   assumes "c>(0::real)"
   shows "Inf {ereal c * x |x. P x} = ereal c * Inf {x. P x}"
 proof -
-  have "(\<lambda>x::ereal. c * x) (Inf {x::ereal. P x}) = Inf ((\<lambda>x::ereal. c * x)`{x::ereal. P x})"
-    apply (rule mono_bij_Inf)
-    apply (simp add: assms ereal_mult_left_mono less_imp_le mono_def)
-    apply (rule bij_betw_byWitness[of _ "\<lambda>x. (x::ereal) / c"], auto simp add: assms ereal_mult_divide)
-    using assms ereal_divide_eq apply auto
-    done
-  then show ?thesis by (simp only: setcompr_eq_image[symmetric])
+  have "bij ((*) (ereal c))"
+    apply (rule bij_betw_byWitness[of _ "\<lambda>x. (x::ereal) / c"], auto simp: assms ereal_mult_divide)
+    using assms ereal_divide_eq by auto
+  then have "ereal c * Inf {x. P x} = Inf ((*) (ereal c) ` {x. P x})"
+    by (simp add: assms ereal_mult_left_mono less_imp_le mono_def mono_bij_Inf)
+  then show ?thesis
+    by (simp add: setcompr_eq_image)
 qed
 
 
@@ -425,7 +373,7 @@
     fix M::real
     have "eventually (\<lambda>x. f x > ereal(M - C)) F" using f by (simp add: tendsto_PInfty)
     then have "eventually (\<lambda>x. (f x > ereal (M-C)) \<and> (g x > ereal C)) F"
-      by (auto simp add: ge eventually_conj_iff)
+      by (auto simp: ge eventually_conj_iff)
     moreover have "\<And>x. ((f x > ereal (M-C)) \<and> (g x > ereal C)) \<Longrightarrow> (f x + g x > ereal M)"
       using ereal_add_strict_mono2 by fastforce
     ultimately have "eventually (\<lambda>x. f x + g x > ereal M) F" using eventually_mono by force
@@ -462,7 +410,7 @@
     fix M::real
     have "eventually (\<lambda>x. f x < ereal(M - C)) F" using f by (simp add: tendsto_MInfty)
     then have "eventually (\<lambda>x. (f x < ereal (M- C)) \<and> (g x < ereal C)) F"
-      by (auto simp add: ge eventually_conj_iff)
+      by (auto simp: ge eventually_conj_iff)
     moreover have "\<And>x. ((f x < ereal (M-C)) \<and> (g x < ereal C)) \<Longrightarrow> (f x + g x < ereal M)"
       using ereal_add_strict_mono2 by fastforce
     ultimately have "eventually (\<lambda>x. f x + g x < ereal M) F" using eventually_mono by force
@@ -495,7 +443,7 @@
   shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
 proof -
   have "((\<lambda>x. g x + f x) \<longlongrightarrow> x + y) F"
-    using tendsto_add_ereal_general1[OF x, OF g, OF f] add.commute[of "y", of "x"] by simp
+    by (metis (full_types) add.commute f g tendsto_add_ereal_general1 x)
   moreover have "\<And>x. g x + f x = f x + g x" using add.commute by auto
   ultimately show ?thesis by simp
 qed
@@ -511,7 +459,8 @@
 proof (cases x)
   case (real r)
   show ?thesis
-    apply (rule tendsto_add_ereal_general2) using real assms by auto
+    using real assms
+    by (intro tendsto_add_ereal_general2; auto)
 next
   case (PInf)
   then have "y \<noteq> -\<infinity>" using assms by simp
@@ -519,7 +468,8 @@
 next
   case (MInf)
   then have "y \<noteq> \<infinity>" using assms by simp
-  then show ?thesis using tendsto_add_ereal_MInf MInf f g by (metis ereal_MInfty_eq_plus)
+  then show ?thesis 
+    by (metis ereal_MInfty_eq_plus tendsto_add_ereal_MInf MInf f g)
 qed
 
 subsubsection\<^marker>\<open>tag important\<close> \<open>Continuity of multiplication\<close>
@@ -541,13 +491,16 @@
 
   {
     fix n assume "u n = ereal(real_of_ereal(u n))" "v n = ereal(real_of_ereal(v n))"
-    then have "ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n" by (metis times_ereal.simps(1))
+    then have "ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n" 
+      by (metis times_ereal.simps(1))
   }
   then have *: "eventually (\<lambda>n. ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n) F"
     using eventually_elim2[OF ureal vreal] by auto
 
-  have "((\<lambda>n. real_of_ereal(u n) * real_of_ereal(v n)) \<longlongrightarrow> l * m) F" using tendsto_mult[OF limu limv] by auto
-  then have "((\<lambda>n. ereal(real_of_ereal(u n)) * real_of_ereal(v n)) \<longlongrightarrow> ereal(l * m)) F" by auto
+  have "((\<lambda>n. real_of_ereal(u n) * real_of_ereal(v n)) \<longlongrightarrow> l * m) F" 
+    using tendsto_mult[OF limu limv] by auto
+  then have "((\<lambda>n. ereal(real_of_ereal(u n)) * real_of_ereal(v n)) \<longlongrightarrow> ereal(l * m)) F" 
+    by auto
   then show ?thesis using * filterlim_cong by fastforce
 qed
 
@@ -556,8 +509,10 @@
   assumes "(f \<longlongrightarrow> l) F" "l>0" "(g \<longlongrightarrow> \<infinity>) F"
   shows "((\<lambda>x. f x * g x) \<longlongrightarrow> \<infinity>) F"
 proof -
-  obtain a::real where "0 < ereal a" "a < l" using assms(2) using ereal_dense2 by blast
-  have *: "eventually (\<lambda>x. f x > a) F" using \<open>a < l\<close> assms(1) by (simp add: order_tendsto_iff)
+  obtain a::real where "0 < ereal a" "a < l" 
+    using assms(2) using ereal_dense2 by blast
+  have *: "eventually (\<lambda>x. f x > a) F" 
+    using \<open>a < l\<close> assms(1) by (simp add: order_tendsto_iff)
   {
     fix K::real
     define M where "M = max K 1"
@@ -573,10 +528,10 @@
 
     have "eventually (\<lambda>x. g x > M/a) F" using assms(3) by (simp add: tendsto_PInfty)
     then have "eventually (\<lambda>x. (f x > a) \<and> (g x > M/a)) F"
-      using * by (auto simp add: eventually_conj_iff)
+      using * by (auto simp: eventually_conj_iff)
     then have "eventually (\<lambda>x. f x * g x > K) F" using eventually_mono Imp by force
   }
-  then show ?thesis by (auto simp add: tendsto_PInfty)
+  then show ?thesis by (auto simp: tendsto_PInfty)
 qed
 
 lemma tendsto_mult_ereal_pos:
@@ -611,12 +566,12 @@
 lemma ereal_sgn_abs:
   fixes l::ereal
   shows "sgn(l) * l = abs(l)"
-apply (cases l) by (auto simp add: sgn_if ereal_less_uminus_reorder)
+    by (cases l, auto simp: sgn_if ereal_less_uminus_reorder)
 
 lemma sgn_squared_ereal:
   assumes "l \<noteq> (0::ereal)"
   shows "sgn(l) * sgn(l) = 1"
-apply (cases l) using assms by (auto simp add: one_ereal_def sgn_if)
+  using assms by (cases l, auto simp: one_ereal_def sgn_if)
 
 lemma tendsto_mult_ereal [tendsto_intros]:
   fixes f g::"_ \<Rightarrow> ereal"
@@ -638,13 +593,13 @@
     by (metis abs_ereal_ge0 abs_ereal_less0 abs_ereal_pos ereal_uminus_uminus ereal_uminus_zero less_le not_less)+
   then have "sgn(l) * l > 0" "sgn(m) * m > 0" using ereal_sgn_abs by auto
   moreover have "((\<lambda>x. sgn(l) * f x) \<longlongrightarrow> (sgn(l) * l)) F"
-    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(1))
+    by (rule tendsto_cmult_ereal, auto simp: sgn_finite assms(1))
   moreover have "((\<lambda>x. sgn(m) * g x) \<longlongrightarrow> (sgn(m) * m)) F"
-    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(2))
+    by (rule tendsto_cmult_ereal, auto simp: sgn_finite assms(2))
   ultimately have *: "((\<lambda>x. (sgn(l) * f x) * (sgn(m) * g x)) \<longlongrightarrow> (sgn(l) * l) * (sgn(m) * m)) F"
     using tendsto_mult_ereal_pos by force
   have "((\<lambda>x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x))) \<longlongrightarrow> (sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m))) F"
-    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite2 *)
+    by (rule tendsto_cmult_ereal, auto simp: sgn_finite2 *)
   moreover have "\<And>x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x)) = f x * g x"
     by (metis mult.left_neutral sgn_squared_ereal[OF \<open>l \<noteq> 0\<close>] sgn_squared_ereal[OF \<open>m \<noteq> 0\<close>] mult.assoc mult.commute)
   moreover have "(sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m)) = l * m"
@@ -656,7 +611,7 @@
   fixes f::"_ \<Rightarrow> ereal" and c::ereal
   assumes "(f \<longlongrightarrow> l) F" "\<not> (l=0 \<and> abs(c) = \<infinity>)"
   shows "((\<lambda>x. c * f x) \<longlongrightarrow> c * l) F"
-by (cases "c = 0", auto simp add: assms tendsto_mult_ereal)
+by (cases "c = 0", auto simp: assms tendsto_mult_ereal)
 
 
 subsubsection\<^marker>\<open>tag important\<close> \<open>Continuity of division\<close>
@@ -675,13 +630,15 @@
       fix z::ereal assume "z>1/e"
       then have "z>0" using \<open>e>0\<close> using less_le_trans not_le by fastforce
       then have "1/z \<ge> 0" by auto
-      moreover have "1/z < e" using \<open>e>0\<close> \<open>z>1/e\<close>
-        apply (cases z) apply auto
-        by (metis (mono_tags, opaque_lifting) less_ereal.simps(2) less_ereal.simps(4) divide_less_eq ereal_divide_less_pos ereal_less(4)
-            ereal_less_eq(4) less_le_trans mult_eq_0_iff not_le not_one_less_zero times_ereal.simps(1))
+      moreover have "1/z < e" 
+      proof (cases z)
+        case (real r)
+        then show ?thesis
+          using \<open>0 < e\<close> \<open>0 < z\<close> \<open>ereal (1 / e) < z\<close> divide_less_eq ereal_divide_less_pos by fastforce 
+      qed (use \<open>0 < e\<close> \<open>0 < z\<close> in auto)
       ultimately have "1/z \<ge> 0" "1/z < e" by auto
     }
-    ultimately have "eventually (\<lambda>n. 1/u n<e) F" "eventually (\<lambda>n. 1/u n\<ge>0) F" by (auto simp add: eventually_mono)
+    ultimately have "eventually (\<lambda>n. 1/u n<e) F" "eventually (\<lambda>n. 1/u n\<ge>0) F" by (auto simp: eventually_mono)
   } note * = this
   show ?thesis
   proof (subst order_tendsto_iff, auto)
@@ -755,7 +712,7 @@
   define h where "h = (\<lambda>x. 1/ g x)"
   have *: "(h \<longlongrightarrow> 1/m) F" unfolding h_def using assms(2) assms(3) tendsto_inverse_ereal by auto
   have "((\<lambda>x. f x * h x) \<longlongrightarrow> l * (1/m)) F"
-    apply (rule tendsto_mult_ereal[OF assms(1) *]) using assms(3) assms(4) by (auto simp add: divide_ereal_def)
+    apply (rule tendsto_mult_ereal[OF assms(1) *]) using assms(3) assms(4) by (auto simp: divide_ereal_def)
   moreover have "f x * h x = f x / g x" for x unfolding h_def by (simp add: divide_ereal_def)
   moreover have "l * (1/m) = l/m" by (simp add: divide_ereal_def)
   ultimately show ?thesis unfolding h_def using Lim_transform_eventually by auto
@@ -771,9 +728,12 @@
   assumes "(u \<longlongrightarrow> l) F" "(v \<longlongrightarrow> m) F" "\<not>((l = \<infinity> \<and> m = \<infinity>) \<or> (l = -\<infinity> \<and> m = -\<infinity>))"
   shows "((\<lambda>n. u n - v n) \<longlongrightarrow> l - m) F"
 proof -
-  have "((\<lambda>n. u n + (-v n)) \<longlongrightarrow> l + (-m)) F"
-    apply (intro tendsto_intros assms) using assms by (auto simp add: ereal_uminus_eq_reorder)
-  then show ?thesis by (simp add: minus_ereal_def)
+  have "\<infinity> = l \<longrightarrow> ((\<lambda>a. u a + - v a) \<longlongrightarrow> l + - m) F"
+      by (metis (no_types) assms ereal_uminus_uminus tendsto_add_ereal_general tendsto_uminus_ereal)
+  then have "((\<lambda>n. u n + (-v n)) \<longlongrightarrow> l + (-m)) F"
+    by (simp add: assms ereal_uminus_eq_reorder tendsto_add_ereal_general)
+  then show ?thesis 
+    by (simp add: minus_ereal_def)
 qed
 
 lemma id_nat_ereal_tendsto_PInf [tendsto_intros]:
@@ -852,11 +812,11 @@
   then show ?thesis by (simp add: tendsto_eventually)
 next
   case (PInf)
-  then have "min x n = n" for n::nat by (auto simp add: min_def)
+  then have "min x n = n" for n::nat by (auto simp: min_def)
   then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
 next
   case (MInf)
-  then have "min x n = x" for n::nat by (auto simp add: min_def)
+  then have "min x n = x" for n::nat by (auto simp: min_def)
   then show ?thesis by auto
 qed
 
@@ -874,7 +834,7 @@
   then show ?thesis using real by auto
 next
   case (PInf)
-  then have "real_of_ereal(min x n) = n" for n::nat by (auto simp add: min_def)
+  then have "real_of_ereal(min x n) = n" for n::nat by (auto simp: min_def)
   then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
 qed (simp add: assms)
 
@@ -889,13 +849,13 @@
   then show ?thesis by (simp add: tendsto_eventually)
 next
   case (MInf)
-  then have "max x (-real n) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def)
+  then have "max x (-real n) = (-1)* ereal(real n)" for n::nat by (auto simp: max_def)
   moreover have "(\<lambda>n. (-1)* ereal(real n)) \<longlonglongrightarrow> -\<infinity>"
     using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def)
   ultimately show ?thesis using MInf by auto
 next
   case (PInf)
-  then have "max x (-real n) = x" for n::nat by (auto simp add: max_def)
+  then have "max x (-real n) = x" for n::nat by (auto simp: max_def)
   then show ?thesis by auto
 qed
 
@@ -913,7 +873,7 @@
   then show ?thesis using real by auto
 next
   case (MInf)
-  then have "real_of_ereal(max x (-real n)) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def)
+  then have "real_of_ereal(max x (-real n)) = (-1)* ereal(real n)" for n::nat by (auto simp: max_def)
   moreover have "(\<lambda>n. (-1)* ereal(real n)) \<longlonglongrightarrow> -\<infinity>"
     using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def)
   ultimately show ?thesis using MInf by auto
@@ -937,7 +897,7 @@
   have "continuous_on ({..0} \<union> {(0::ereal)..}) abs"
     apply (rule continuous_on_closed_Un, auto)
     apply (rule iffD1[OF continuous_on_cong, of "{..0}" _ "\<lambda>x. -x"])
-    using less_eq_ereal_def apply (auto simp add: continuous_uminus_ereal)
+    using less_eq_ereal_def apply (auto simp: continuous_uminus_ereal)
     apply (rule iffD1[OF continuous_on_cong, of "{0..}" _ "\<lambda>x. x"])
       apply (auto)
     done
@@ -979,9 +939,9 @@
     apply (intro tendsto_intros) using assms apply auto
     using enn2ereal_inject zero_ennreal.rep_eq by fastforce+
   moreover have "e2ennreal(enn2ereal (u n) * enn2ereal (v n)) = u n * v n" for n
-    by (subst times_ennreal.abs_eq[symmetric], auto simp add: eq_onp_same_args)
+    by (subst times_ennreal.abs_eq[symmetric], auto simp: eq_onp_same_args)
   moreover have "e2ennreal(enn2ereal l * enn2ereal m)  = l * m"
-    by (subst times_ennreal.abs_eq[symmetric], auto simp add: eq_onp_same_args)
+    by (subst times_ennreal.abs_eq[symmetric], auto simp: eq_onp_same_args)
   ultimately show ?thesis
     by auto
 qed
@@ -1167,7 +1127,7 @@
   obtain C where C: "limsup u < C" "C < \<infinity>" using assms ereal_dense2 by blast
   then have "C = ereal(real_of_ereal C)" using ereal_real by force
   have "eventually (\<lambda>n. u n < C) sequentially" using C(1) unfolding Limsup_def
-    apply (auto simp add: INF_less_iff)
+    apply (auto simp: INF_less_iff)
     using SUP_lessD eventually_mono by fastforce
   then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> u n < C" using eventually_sequentially by auto
   define D where "D = max (real_of_ereal C) (Max {u n |n. n \<le> N})"
@@ -1176,7 +1136,7 @@
     fix n show "u n \<le> D"
     proof (cases)
       assume *: "n \<le> N"
-      have "u n \<le> Max {u n |n. n \<le> N}" by (rule Max_ge, auto simp add: *)
+      have "u n \<le> Max {u n |n. n \<le> N}" by (rule Max_ge, auto simp: *)
       then show "u n \<le> D" unfolding D_def by linarith
     next
       assume "\<not>(n \<le> N)"
@@ -1197,7 +1157,7 @@
   obtain C where C: "liminf u > C" "C > -\<infinity>" using assms using ereal_dense2 by blast
   then have "C = ereal(real_of_ereal C)" using ereal_real by force
   have "eventually (\<lambda>n. u n > C) sequentially" using C(1) unfolding Liminf_def
-    apply (auto simp add: less_SUP_iff)
+    apply (auto simp: less_SUP_iff)
     using eventually_elim2 less_INF_D by fastforce
   then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> u n > C" using eventually_sequentially by auto
   define D where "D = min (real_of_ereal C) (Min {u n |n. n \<le> N})"
@@ -1206,7 +1166,7 @@
     fix n show "u n \<ge> D"
     proof (cases)
       assume *: "n \<le> N"
-      have "u n \<ge> Min {u n |n. n \<le> N}" by (rule Min_le, auto simp add: *)
+      have "u n \<ge> Min {u n |n. n \<le> N}" by (rule Min_le, auto simp: *)
       then show "u n \<ge> D" unfolding D_def by linarith
     next
       assume "\<not>(n \<le> N)"
@@ -1615,7 +1575,7 @@
   have "eventually (\<lambda>n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast
   moreover have "eventually (\<lambda>n. (u o s) n < \<infinity>) sequentially" using assms(3) * order_tendsto_iff by blast
   moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \<infinity>" for n
-    unfolding w_def using that by (auto simp add: ereal_divide_eq)
+    unfolding w_def using that by (auto simp: ereal_divide_eq)
   ultimately have "eventually (\<lambda>n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force
   moreover have "(\<lambda>n. (w o s) n / (u o s) n) \<longlonglongrightarrow> (limsup w) / a"
     apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto
@@ -1645,7 +1605,7 @@
   have "eventually (\<lambda>n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast
   moreover have "eventually (\<lambda>n. (u o s) n < \<infinity>) sequentially" using assms(3) * order_tendsto_iff by blast
   moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \<infinity>" for n
-    unfolding w_def using that by (auto simp add: ereal_divide_eq)
+    unfolding w_def using that by (auto simp: ereal_divide_eq)
   ultimately have "eventually (\<lambda>n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force
   moreover have "(\<lambda>n. (w o s) n / (u o s) n) \<longlonglongrightarrow> (liminf w) / a"
     apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto
--- a/src/HOL/Analysis/Function_Topology.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Function_Topology.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -118,12 +118,7 @@
           have *: "f \<in> U"
             if "U \<in> \<U>" and "\<forall>V\<in>\<U>. \<forall>i. i \<in> I \<and> Y V i \<noteq> topspace (X i) \<longrightarrow> f i \<in> Y V i"
               and "\<forall>i\<in>I. f i \<in> topspace (X i)" and "f \<in> extensional I" for f U
-          proof -
-            have "Pi\<^sub>E I (Y U) = U"
-              using Y \<open>U \<in> \<U>\<close> by blast
-            then show "f \<in> U"
-              using that unfolding PiE_def Pi_def by blast
-          qed
+            by (smt (verit) PiE_iff Y that)
           show "?TOP \<inter> \<Inter>(\<Union>U\<in>\<U>. ?F U) \<subseteq> \<Inter>\<U>"
             by (auto simp: PiE_iff *)
           show "\<Inter>\<U> \<subseteq> ?TOP \<inter> \<Inter>(\<Union>U\<in>\<U>. ?F U)"
@@ -158,18 +153,16 @@
           using Y by force
         show "?G ` \<U> \<subseteq> {Pi\<^sub>E I Y |Y. (\<forall>i. openin (X i) (Y i)) \<and> finite {i. Y i \<noteq> topspace (X i)}}"
           apply clarsimp
-          apply (rule_tac x=" (\<lambda>i. if i = J U then Y U else topspace (X i))" in exI)
+          apply (rule_tac x= "(\<lambda>i. if i = J U then Y U else topspace (X i))" in exI)
           apply (auto simp: *)
           done
       next
         show "(\<Inter>U\<in>\<U>. ?G U) = ?TOP \<inter> \<Inter>\<U>"
         proof
           have "(\<Pi>\<^sub>E i\<in>I. if i = J U then Y U else topspace (X i)) \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
-            apply (clarsimp simp: PiE_iff Ball_def all_conj_distrib split: if_split_asm)
-            using Y \<open>U \<in> \<U>\<close> openin_subset by blast+
+            by (simp add: PiE_mono Y \<open>U \<in> \<U>\<close> openin_subset)
           then have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> ?TOP"
-            using \<open>U \<in> \<U>\<close>
-            by fastforce
+            using \<open>U \<in> \<U>\<close> by fastforce
           moreover have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> \<Inter>\<U>"
             using PiE_mem Y by fastforce
           ultimately show "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> ?TOP \<inter> \<Inter>\<U>"
@@ -211,24 +204,25 @@
   assumes "openin (product_topology T I) U" "x \<in> U"
   shows "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
 proof -
-  have "generate_topology_on {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}} U"
-    using assms unfolding product_topology_def by (intro openin_topology_generated_by) auto
-  then have "\<And>x. x\<in>U \<Longrightarrow> \<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
+  define IT where "IT \<equiv> \<lambda>X. {i. X i \<noteq> topspace (T i)}"
+  have "generate_topology_on {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite (IT X)} U"
+    using assms unfolding product_topology_def IT_def by (intro openin_topology_generated_by) auto
+  then have "\<And>x. x\<in>U \<Longrightarrow> \<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite (IT X) \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
   proof induction
     case (Int U V x)
     then obtain XU XV where H:
-      "x \<in> Pi\<^sub>E I XU" "(\<forall>i. openin (T i) (XU i))" "finite {i. XU i \<noteq> topspace (T i)}" "Pi\<^sub>E I XU \<subseteq> U"
-      "x \<in> Pi\<^sub>E I XV" "(\<forall>i. openin (T i) (XV i))" "finite {i. XV i \<noteq> topspace (T i)}" "Pi\<^sub>E I XV \<subseteq> V"
-      by auto meson
+      "x \<in> Pi\<^sub>E I XU" "\<And>i. openin (T i) (XU i)" "finite (IT XU)" "Pi\<^sub>E I XU \<subseteq> U"
+      "x \<in> Pi\<^sub>E I XV" "\<And>i. openin (T i) (XV i)" "finite (IT XV)" "Pi\<^sub>E I XV \<subseteq> V"
+      by (meson Int_iff)
     define X where "X = (\<lambda>i. XU i \<inter> XV i)"
     have "Pi\<^sub>E I X \<subseteq> Pi\<^sub>E I XU \<inter> Pi\<^sub>E I XV"
-      unfolding X_def by (auto simp add: PiE_iff)
+      by (auto simp add: PiE_iff X_def)
     then have "Pi\<^sub>E I X \<subseteq> U \<inter> V" using H by auto
     moreover have "\<forall>i. openin (T i) (X i)"
       unfolding X_def using H by auto
-    moreover have "finite {i. X i \<noteq> topspace (T i)}"
-      apply (rule rev_finite_subset[of "{i. XU i \<noteq> topspace (T i)} \<union> {i. XV i \<noteq> topspace (T i)}"])
-      unfolding X_def using H by auto
+    moreover have "finite (IT X)"
+      apply (rule rev_finite_subset[of "IT XU \<union> IT XV"])
+      using H by (auto simp: X_def IT_def)
     moreover have "x \<in> Pi\<^sub>E I X"
       unfolding X_def using H by auto
     ultimately show ?case
@@ -236,16 +230,10 @@
   next
     case (UN K x)
     then obtain k where "k \<in> K" "x \<in> k" by auto
-    with UN have "\<exists>X. x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> k"
-      by simp
-    then obtain X where "x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> k"
-      by blast
-    then have "x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> (\<Union>K)"
-      using \<open>k \<in> K\<close> by auto
-    then show ?case
-      by auto
+    with \<open>k \<in> K\<close> UN show ?case
+      by (meson Sup_upper2)
   qed auto
-  then show ?thesis using \<open>x \<in> U\<close> by auto
+  then show ?thesis using \<open>x \<in> U\<close> IT_def by blast
 qed
 
 lemma product_topology_empty_discrete:
@@ -257,9 +245,7 @@
     arbitrary union_of
           ((finite intersection_of (\<lambda>F. (\<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U)))
            relative_to topspace (product_topology X I))"
-  apply (subst product_topology)
-  apply (simp add: topology_inverse' [OF istopology_subbase])
-  done
+  by (simp add: istopology_subbase product_topology)
 
 lemma subtopology_PiE:
   "subtopology (product_topology X I) (\<Pi>\<^sub>E i\<in>I. (S i)) = product_topology (\<lambda>i. subtopology (X i) (S i)) I"
@@ -350,9 +336,7 @@
     (\<forall>x \<in> S. \<exists>U. finite {i \<in> I. U i \<noteq> topspace(X i)} \<and>
                  (\<forall>i \<in> I. openin (X i) (U i)) \<and> x \<in> Pi\<^sub>E I U \<and> Pi\<^sub>E I U \<subseteq> S)"
   unfolding openin_product_topology arbitrary_union_of_alt product_topology_base_alt topspace_product_topology
-  apply safe
-  apply (drule bspec; blast)+
-  done
+  by (smt (verit, best))
 
 lemma closure_of_product_topology:
   "(product_topology X I) closure_of (PiE I S) = PiE I (\<lambda>i. (X i) closure_of (S i))"
@@ -438,9 +422,7 @@
 
 corollary closedin_product_topology:
    "closedin (product_topology X I) (PiE I S) \<longleftrightarrow> PiE I S = {} \<or> (\<forall>i \<in> I. closedin (X i) (S i))"
-  apply (simp add: PiE_eq PiE_eq_empty_iff closure_of_product_topology flip: closure_of_eq)
-  apply (metis closure_of_empty)
-  done
+  by (smt (verit, best) PiE_eq closedin_empty closure_of_eq closure_of_product_topology)
 
 corollary closedin_product_topology_singleton:
    "f \<in> extensional I \<Longrightarrow> closedin (product_topology X I) {f} \<longleftrightarrow> (\<forall>i \<in> I. closedin (X i) {f i})"
@@ -501,17 +483,13 @@
   also have "... = (\<Inter>i\<in>J. (\<lambda>x. f x i)-`(X i) \<inter> topspace T1) \<inter> (topspace T1)"
     using * \<open>J \<subseteq> I\<close> by auto
   also have "openin T1 (...)"
-    apply (rule openin_INT)
-    apply (simp add: \<open>finite J\<close>)
-    using H(2) assms(1) \<open>J \<subseteq> I\<close> by auto
+    using H(2) \<open>J \<subseteq> I\<close> \<open>finite J\<close> assms(1) by blast
   ultimately show "openin T1 (f-`U \<inter> topspace T1)" by simp
 next
-  show "f `topspace T1 \<subseteq> \<Union>{Pi\<^sub>E I X |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
-    apply (subst topology_generated_by_topspace[symmetric])
-    apply (subst product_topology_def[symmetric])
-    apply (simp only: topspace_product_topology)
-    apply (auto simp add: PiE_iff)
-    using assms unfolding continuous_map_def by auto
+  have "f ` topspace T1 \<subseteq> topspace (product_topology T I)"
+    using assms continuous_map_def by fastforce
+  then show "f `topspace T1 \<subseteq> \<Union>{Pi\<^sub>E I X |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
+    by (simp add: product_topology_def)
 qed
 
 lemma continuous_map_product_then_coordinatewise [intro]:
@@ -522,8 +500,7 @@
   fix i assume "i \<in> I"
   have "(\<lambda>x. f x i) = (\<lambda>y. y i) o f" by auto
   also have "continuous_map T1 (T i) (...)"
-    apply (rule continuous_map_compose[of _ "product_topology T I"])
-    using assms \<open>i \<in> I\<close> by auto
+    by (metis \<open>i \<in> I\<close> assms continuous_map_compose continuous_map_product_coordinates)
   ultimately show "continuous_map T1 (T i) (\<lambda>x. f x i)"
     by simp
 next
@@ -583,28 +560,14 @@
   assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> open (U i)"
   shows "open {f. \<forall>i\<in>I. f (x i) \<in> U i}"
 proof -
-  define J where "J = x`I"
-  define V where "V = (\<lambda>y. if y \<in> J then \<Inter>{U i|i. i\<in>I \<and> x i = y} else UNIV)"
-  define X where "X = (\<lambda>y. if y \<in> J then V y else UNIV)"
+  define V where "V \<equiv> (\<lambda>y. if y \<in> x`I then \<Inter>{U i|i. i\<in>I \<and> x i = y} else UNIV)"
+  define X where "X \<equiv> (\<lambda>y. if y \<in> x`I then V y else UNIV)"
   have *: "open (X i)" for i
     unfolding X_def V_def using assms by auto
-  have **: "finite {i. X i \<noteq> UNIV}"
-    unfolding X_def V_def J_def using assms(1) by auto
-  have "open (Pi\<^sub>E UNIV X)"
-    by (simp add: "*" "**" open_PiE)
+  then have "open (Pi\<^sub>E UNIV X)"
+    by (simp add: X_def assms(1) open_PiE)
   moreover have "Pi\<^sub>E UNIV X = {f. \<forall>i\<in>I. f (x i) \<in> U i}"
-    apply (auto simp add: PiE_iff) unfolding X_def V_def J_def
-    proof (auto)
-      fix f :: "'a \<Rightarrow> 'b" and i :: 'i
-      assume a1: "i \<in> I"
-      assume a2: "\<forall>i. f i \<in> (if i \<in> x`I then if i \<in> x`I then \<Inter>{U ia |ia. ia \<in> I \<and> x ia = i} else UNIV else UNIV)"
-      have f3: "x i \<in> x`I"
-        using a1 by blast
-      have "U i \<in> {U ia |ia. ia \<in> I \<and> x ia = x i}"
-        using a1 by blast
-      then show "f (x i) \<in> U i"
-        using f3 a2 by (meson Inter_iff)
-    qed
+    by (fastforce simp add: PiE_iff X_def V_def split: if_split_asm)
   ultimately show ?thesis by simp
 qed
 
@@ -620,25 +583,13 @@
   fixes f :: "'a::topological_space \<Rightarrow> 'b \<Rightarrow> 'c::topological_space"
   assumes "\<And>i. continuous_on S (\<lambda>x. f x i)"
   shows "continuous_on S f"
-  using continuous_map_coordinatewise_then_product [of UNIV, where T = "\<lambda>i. euclidean"]
-  by (metis UNIV_I assms continuous_map_iff_continuous euclidean_product_topology)
-
-lemma continuous_on_product_then_coordinatewise_UNIV:
-  assumes "continuous_on UNIV f"
-  shows "continuous_on UNIV (\<lambda>x. f x i)"
-  unfolding continuous_map_iff_continuous2 [symmetric]
-  by (rule continuous_map_product_then_coordinatewise [where I=UNIV]) (use assms in \<open>auto simp: euclidean_product_topology\<close>)
+  by (metis UNIV_I assms continuous_map_iff_continuous euclidean_product_topology
+      continuous_map_coordinatewise_then_product)
 
 lemma continuous_on_product_then_coordinatewise:
   assumes "continuous_on S f"
   shows "continuous_on S (\<lambda>x. f x i)"
-proof -
-  have "continuous_on S ((\<lambda>q. q i) \<circ> f)"
-    by (metis assms continuous_on_compose continuous_on_id
-        continuous_on_product_then_coordinatewise_UNIV continuous_on_subset subset_UNIV)
-  then show ?thesis
-    by auto
-qed
+  by (metis UNIV_I assms continuous_map_iff_continuous continuous_map_product_then_coordinatewise(1) euclidean_product_topology)
 
 lemma continuous_on_coordinatewise_iff:
   fixes f :: "('a \<Rightarrow> real) \<Rightarrow> 'b \<Rightarrow> real"
@@ -679,16 +630,15 @@
   next
     case (Suc N)
     define f::"((nat \<Rightarrow> 'a) \<times> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)"
-      where "f = (\<lambda>(x, b). (\<lambda>i. if i = N then b else x i))"
+      where "f = (\<lambda>(x, b). x(N:=b))"
     have "{x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>Suc N. x i = a)} \<subseteq> f`({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
     proof (auto)
       fix x assume H: "\<forall>i::nat. x i \<in> F" "\<forall>i\<ge>Suc N. x i = a"
-      define y where "y = (\<lambda>i. if i = N then a else x i)"
-      have "f (y, x N) = x"
-        unfolding f_def y_def by auto
-      moreover have "(y, x N) \<in> {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F"
-        unfolding y_def using H \<open>a \<in> F\<close> by auto
-      ultimately show "x \<in> f`({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
+      have "f (x(N:=a), x N) = x"
+        unfolding f_def by auto
+      moreover have "(x(N:=a), x N) \<in> {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F"
+        using H \<open>a \<in> F\<close> by auto
+      ultimately show "x \<in> f ` ({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
         by (metis (no_types, lifting) image_eqI)
     qed
     moreover have "countable ({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
@@ -742,7 +692,7 @@
     by metis
   text \<open>\<open>B i\<close> is a countable basis of neighborhoods of \<open>x\<^sub>i\<close>.\<close>
   define B where "B = (\<lambda>i. (A (x i))`UNIV \<union> {UNIV})"
-  have "countable (B i)" for i unfolding B_def by auto
+  have countB: "countable (B i)" for i unfolding B_def by auto
   have open_B: "\<And>X i. X \<in> B i \<Longrightarrow> open X"
     by (auto simp: B_def A)
   define K where "K = {Pi\<^sub>E UNIV X | X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
@@ -750,14 +700,14 @@
     unfolding K_def B_def by auto
   then have "K \<noteq> {}" by auto
   have "countable {X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
-    apply (rule countable_product_event_const) using \<open>\<And>i. countable (B i)\<close> by auto
+    by (simp add: countB countable_product_event_const)
   moreover have "K = (\<lambda>X. Pi\<^sub>E UNIV X)`{X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
     unfolding K_def by auto
   ultimately have "countable K" by auto
-  have "x \<in> k" if "k \<in> K" for k
+  have I: "x \<in> k" if "k \<in> K" for k
     using that unfolding K_def B_def apply auto using A(1) by auto
-  have "open k" if "k \<in> K" for k
-    using that unfolding K_def  by (blast intro: open_B open_PiE elim: )
+  have II: "open k" if "k \<in> K" for k
+    using that unfolding K_def by (blast intro: open_B open_PiE)
   have Inc: "\<exists>k\<in>K. k \<subseteq> U" if "open U \<and> x \<in> U" for U
   proof -
     have "openin (product_topology (\<lambda>i. euclidean) UNIV) U" "x \<in> U"
@@ -775,31 +725,28 @@
     have *: "\<exists>y. y \<in> B i \<and> y \<subseteq> X i" for i
       unfolding B_def using A(3)[OF H(2)] H(1) by (metis PiE_E UNIV_I UnCI image_iff)
     have **: "Y i \<in> B i \<and> Y i \<subseteq> X i" for i
-      apply (cases "i \<in> I")
-      unfolding Y_def using * that apply (auto)
-      apply (metis (no_types, lifting) someI, metis (no_types, lifting) someI_ex subset_iff)
-      unfolding B_def apply simp
-      unfolding I_def apply auto
-      done
+    proof (cases "i \<in> I")
+      case True
+      then show ?thesis  
+        by (metis (mono_tags, lifting) "*" Nitpick.Eps_psimp Y_def)
+    next
+      case False
+      then show ?thesis by (simp add: B_def I_def Y_def)
+    qed
     have "{i. Y i \<noteq> UNIV} \<subseteq> I"
       unfolding Y_def by auto
-    then have ***: "finite {i. Y i \<noteq> UNIV}"
-      unfolding I_def using H(3) rev_finite_subset by blast
-    have "(\<forall>i. Y i \<in> B i) \<and> finite {i. Y i \<noteq> UNIV}"
-      using ** *** by auto
+    with ** have "(\<forall>i. Y i \<in> B i) \<and> finite {i. Y i \<noteq> UNIV}"
+      using H(3) I_def finite_subset by blast
     then have "Pi\<^sub>E UNIV Y \<in> K"
       unfolding K_def by auto
-
     have "Y i \<subseteq> X i" for i
-      apply (cases "i \<in> I") using ** apply simp unfolding Y_def I_def by auto
-    then have "Pi\<^sub>E UNIV Y \<subseteq> Pi\<^sub>E UNIV X" by auto
-    then have "Pi\<^sub>E UNIV Y \<subseteq> U" using H(4) by auto
+      using "**" by auto
+    then have "Pi\<^sub>E UNIV Y \<subseteq> U"
+      by (metis H(4) PiE_mono subset_trans)
     then show ?thesis using \<open>Pi\<^sub>E UNIV Y \<in> K\<close> by auto
   qed
-
   show "\<exists>L. (\<forall>(i::nat). x \<in> L i \<and> open (L i)) \<and> (\<forall>U. open U \<and> x \<in> U \<longrightarrow> (\<exists>i. L i \<subseteq> U))"
-    apply (rule first_countableI[of K])
-    using \<open>countable K\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> x \<in> k\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> open k\<close> Inc by auto
+    using \<open>countable K\<close> I II Inc by (simp add: first_countableI) 
 qed
 
 proposition product_topology_countable_basis:
@@ -821,7 +768,7 @@
     unfolding K_def using \<open>\<And>U. U \<in> B2 \<Longrightarrow> open U\<close> by auto
 
   have "countable {X. (\<forall>(i::'a). X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
-    apply (rule countable_product_event_const) using \<open>countable B2\<close> by auto
+    using \<open>countable B2\<close>  by (intro countable_product_event_const) auto
   moreover have "K = (\<lambda>X. Pi\<^sub>E UNIV X)`{X. (\<forall>i. X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
     unfolding K_def by auto
   ultimately have ii: "countable K" by auto
@@ -832,9 +779,7 @@
     then have "openin (product_topology (\<lambda>i. euclidean) UNIV) U"
       unfolding open_fun_def by auto
     with product_topology_open_contains_basis[OF this \<open>x \<in> U\<close>]
-    have "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV} \<and> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U"
-      by simp
-    then obtain X where H: "x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i)"
+    obtain X where H: "x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i)"
                            "\<And>i. open (X i)"
                            "finite {i. X i \<noteq> UNIV}"
                            "(\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U"
@@ -845,29 +790,15 @@
     have *: "\<exists>y. y \<in> B2 \<and> y \<subseteq> X i \<and> x i \<in> y" for i
       unfolding B2_def using B \<open>open (X i)\<close> \<open>x i \<in> X i\<close> by (meson UnCI topological_basisE)
     have **: "Y i \<in> B2 \<and> Y i \<subseteq> X i \<and> x i \<in> Y i" for i
-      using someI_ex[OF *]
-      apply (cases "i \<in> I")
-      unfolding Y_def using * apply (auto)
-      unfolding B2_def I_def by auto
+      using someI_ex[OF *] by (simp add: B2_def I_def Y_def)
     have "{i. Y i \<noteq> UNIV} \<subseteq> I"
       unfolding Y_def by auto
-    then have ***: "finite {i. Y i \<noteq> UNIV}"
-      unfolding I_def using H(3) rev_finite_subset by blast
-    have "(\<forall>i. Y i \<in> B2) \<and> finite {i. Y i \<noteq> UNIV}"
-      using ** *** by auto
+    then have "(\<forall>i. Y i \<in> B2) \<and> finite {i. Y i \<noteq> UNIV}"
+      using "**" H(3) I_def finite_subset by blast
     then have "Pi\<^sub>E UNIV Y \<in> K"
       unfolding K_def by auto
-
-    have "Y i \<subseteq> X i" for i
-      apply (cases "i \<in> I") using ** apply simp unfolding Y_def I_def by auto
-    then have "Pi\<^sub>E UNIV Y \<subseteq> Pi\<^sub>E UNIV X" by auto
-    then have "Pi\<^sub>E UNIV Y \<subseteq> U" using H(4) by auto
-
-    have "x \<in> Pi\<^sub>E UNIV Y"
-      using ** by auto
-
-    show "\<exists>V\<in>K. x \<in> V \<and> V \<subseteq> U"
-      using \<open>Pi\<^sub>E UNIV Y \<in> K\<close> \<open>Pi\<^sub>E UNIV Y \<subseteq> U\<close> \<open>x \<in> Pi\<^sub>E UNIV Y\<close> by auto
+    then show "\<exists>V\<in>K. x \<in> V \<and> V \<subseteq> U"
+      by (meson "**" H(4) PiE_I PiE_mono UNIV_I order.trans)
   next
     fix U assume "U \<in> K"
     show "open U"
@@ -878,9 +809,11 @@
 qed
 
 instance "fun" :: (countable, second_countable_topology) second_countable_topology
-  apply standard
-  using product_topology_countable_basis topological_basis_imp_subbasis by auto
-
+proof
+  show "\<exists>B::('a \<Rightarrow> 'b) set set. countable B \<and> open = generate_topology B"
+    using product_topology_countable_basis topological_basis_imp_subbasis 
+    by auto
+qed
 
 subsection\<open>The Alexander subbase theorem\<close>
 
@@ -1106,11 +1039,7 @@
   qed
   ultimately show ?thesis
     by metis
-next
-  case False
-  then show ?thesis
-    by (auto simp: continuous_map_def PiE_def)
-qed
+qed (auto simp: continuous_map_def PiE_def)
 
 
 lemma continuous_map_componentwise_UNIV:
@@ -1147,8 +1076,7 @@
       fix x f
       assume "f \<in> V"
       let ?T = "{a \<in> topspace(Y i).
-                   (\<lambda>j. if j = i then a
-                        else if j \<in> I then f j else undefined) \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>}"
+                   (\<lambda>j\<in>I. f j)(i:=a) \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>}"
       show "\<exists>T. openin (Y i) T \<and> f i \<in> T \<and> T \<subseteq> (\<lambda>f. f i) ` V"
       proof (intro exI conjI)
         show "openin (Y i) ?T"
@@ -1164,8 +1092,8 @@
               by simp (metis IntD1 PiE_iff V \<open>f \<in> V\<close> that)
           qed
           then show "continuous_map (Y i) (product_topology Y I)
-                  (\<lambda>x j. if j = i then x else if j \<in> I then f j else undefined)"
-            by (auto simp: continuous_map_componentwise assms extensional_def)
+                  (\<lambda>x. (\<lambda>j\<in>I. f j)(i:=x))"
+            by (auto simp: continuous_map_componentwise assms extensional_def restrict_def)
         next
           have "openin (product_topology Y I) (\<Pi>\<^sub>E i\<in>I. topspace (Y i))"
             by (metis openin_topspace topspace_product_topology)
@@ -1177,18 +1105,15 @@
               show "\<And>X. X \<in> (\<inter>) (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) ` \<F> \<Longrightarrow> openin (product_topology Y I) X"
                 unfolding openin_product_topology relative_to_def
                 apply (clarify intro!: arbitrary_union_of_inc)
-                apply (rename_tac F)
-                apply (rule_tac x=F in exI)
                 using subsetD [OF \<F>]
-                apply (force intro: finite_intersection_of_inc)
-                done
+                by (metis (mono_tags, lifting) finite_intersection_of_inc mem_Collect_eq topspace_product_topology)
             qed (use \<open>finite \<F>\<close> \<open>\<F> \<noteq> {}\<close> in auto)
           qed
           ultimately show "openin (product_topology Y I) ((\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>)"
             by (auto simp only: Int_Inter_eq split: if_split)
         qed
       next
-        have eqf: "(\<lambda>j. if j = i then f i else if j \<in> I then f j else undefined) = f"
+        have eqf: "(\<lambda>j\<in>I. f j)(i:=f i) = f"
           using PiE_arb V \<open>f \<in> V\<close> by force
         show "f i \<in> ?T"
           using V assms \<open>f \<in> V\<close> by (auto simp: PiE_iff eqf)
@@ -1273,10 +1198,8 @@
   next
     assume ?rhs
     then show ?lhs
-      apply (simp only: openin_product_topology)
-      apply (rule arbitrary_union_of_inc)
-      apply (auto simp: product_topology_base_alt)
-      done
+      unfolding openin_product_topology
+      by (intro arbitrary_union_of_inc) (auto simp: product_topology_base_alt)
   qed
   ultimately show ?thesis
     by simp
@@ -1306,9 +1229,7 @@
         by (simp add: continuous_map_product_projection)
       moreover
       have "\<And>x. x \<in> topspace (X i) \<Longrightarrow> x \<in> (\<lambda>f. f i) ` (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
-        using \<open>i \<in> I\<close> z
-        apply (rule_tac x="\<lambda>j. if j = i then x else if j \<in> I then z j else undefined" in image_eqI, auto)
-        done
+        using \<open>i \<in> I\<close> z by (rule_tac x="z(i:=x)" in image_eqI) auto
       then have "(\<lambda>f. f i) ` (\<Pi>\<^sub>E i\<in>I. topspace (X i)) = topspace (X i)"
         using \<open>i \<in> I\<close> z by auto
       ultimately show "compactin (X i) (topspace (X i))"
@@ -1461,7 +1382,7 @@
             then have "(\<lambda>x. x k) ` {x. x j \<in> V} = UNIV" if "k \<noteq> j" for k
               using that
               apply (clarsimp simp add: set_eq_iff)
-              apply (rule_tac x="\<lambda>m. if m = k then x else f m" in image_eqI, auto)
+              apply (rule_tac x="f(k:=x)" in image_eqI, auto)
               done
             then have "{i. (\<lambda>x. x i) ` C \<subset> topspace (X i)} \<subseteq> {j}"
               using Ceq by auto
@@ -1493,7 +1414,7 @@
               using PiE_ext \<open>g \<in> U\<close> gin by force
           next
             case (insert i M)
-            define f where "f \<equiv> \<lambda>j. if j = i then g i else h j"
+            define f where "f \<equiv> h(i:=g i)"
             have fin: "f \<in> PiE I (topspace \<circ> X)"
               unfolding f_def using gin insert.prems(1) by auto
             have subM: "{j \<in> I. f j \<noteq> g j} \<subseteq> M"
@@ -1506,13 +1427,13 @@
               show ?thesis
               proof (cases "i \<in> I")
                 case True
-                let ?U = "{x \<in> topspace(X i). (\<lambda>j. if j = i then x else h j) \<in> U}"
-                let ?V = "{x \<in> topspace(X i). (\<lambda>j. if j = i then x else h j) \<in> V}"
+                let ?U = "{x \<in> topspace(X i). h(i:=x) \<in> U}"
+                let ?V = "{x \<in> topspace(X i). h(i:=x) \<in> V}"
                 have False
                 proof (rule connected_spaceD [OF cs [OF \<open>i \<in> I\<close>]])
                   have "\<And>k. k \<in> I \<Longrightarrow> continuous_map (X i) (X k) (\<lambda>x. if k = i then x else h k)"
                     using continuous_map_eq_topcontinuous_at insert.prems(1) topcontinuous_at_def by fastforce
-                  then have cm: "continuous_map (X i) (product_topology X I) (\<lambda>x j. if j = i then x else h j)"
+                  then have cm: "continuous_map (X i) (product_topology X I) (\<lambda>x. h(i:=x))"
                     using \<open>i \<in> I\<close> insert.prems(1)
                     by (auto simp: continuous_map_componentwise extensional_def)
                   show "openin (X i) ?U"
@@ -1522,45 +1443,23 @@
                   show "topspace (X i) \<subseteq> ?U \<union> ?V"
                   proof clarsimp
                     fix x
-                    assume "x \<in> topspace (X i)" and "(\<lambda>j. if j = i then x else h j) \<notin> V"
+                    assume "x \<in> topspace (X i)" and "h(i:=x) \<notin> V"
                     with True subUV \<open>h \<in> Pi\<^sub>E I (topspace \<circ> X)\<close>
-                    show "(\<lambda>j. if j = i then x else h j) \<in> U"
-                      by (drule_tac c="(\<lambda>j. if j = i then x else h j)" in subsetD) auto
+                    show "h(i:=x) \<in> U"
+                      by (force dest: subsetD [where c="h(i:=x)"])
                   qed
                   show "?U \<inter> ?V = {}"
                     using disj by blast
                   show "?U \<noteq> {}"
-                    using \<open>U \<noteq> {}\<close> f_def
-                  proof -
-                    have "(\<lambda>j. if j = i then g i else h j) \<in> U"
-                      using \<open>f \<in> U\<close> f_def by blast
-                    moreover have "f i \<in> topspace (X i)"
-                      by (metis PiE_iff True comp_apply fin)
-                    ultimately have "\<exists>b. b \<in> topspace (X i) \<and> (\<lambda>a. if a = i then b else h a) \<in> U"
-                      using f_def by auto
-                    then show ?thesis
-                      by blast
-                  qed
-                  have "(\<lambda>j. if j = i then h i else h j) = h"
-                    by force
-                  moreover have "h i \<in> topspace (X i)"
-                    using True insert.prems(1) by auto
-                  ultimately show "?V \<noteq> {}"
-                    using \<open>h \<in> V\<close>  by force
+                    using True \<open>f \<in> U\<close> f_def gin by auto
+                  show "?V \<noteq> {}"
+                    using True \<open>h \<in> V\<close> V openin_subset by fastforce
                 qed
                 then show ?thesis ..
               next
                 case False
                 show ?thesis
-                proof (cases "h = f")
-                  case True
-                  show ?thesis
-                    by (rule insert.IH [OF insert.prems(1)]) (simp add: True subM)
-                next
-                  case False
-                  then show ?thesis
-                    using gin insert.prems(1) \<open>i \<notin> I\<close> unfolding f_def by fastforce
-                qed
+                  using insert.prems(1) by (metis False gin PiE_E \<open>f \<in> U\<close> f_def fun_upd_triv)
               qed
             next
               case False
@@ -1640,14 +1539,8 @@
   assumes "i \<in> I"
   shows "quotient_map(product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow>
            (topspace(product_topology X I) = {} \<longrightarrow> topspace(X i) = {})"
-         (is "?lhs = ?rhs")
-proof
-  assume ?lhs with assms show ?rhs
-    by (auto simp: continuous_open_quotient_map open_map_product_projection)
-next
-  assume ?rhs with assms show ?lhs
-    by (auto simp: Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection)
-qed
+  by (metis assms image_empty quotient_imp_surjective_map retraction_imp_quotient_map 
+      retraction_map_product_projection)
 
 lemma product_topology_homeomorphic_component:
   assumes "i \<in> I" "\<And>j. \<lbrakk>j \<in> I; j \<noteq> i\<rbrakk> \<Longrightarrow> \<exists>a. topspace(X j) = {a}"
@@ -1676,9 +1569,8 @@
     from that obtain f where f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" 
       by force
     have "?SX f i homeomorphic_space X i"
-      apply (simp add: subtopology_PiE )
-      using product_topology_homeomorphic_component [OF \<open>i \<in> I\<close>, of "\<lambda>j. subtopology (X j) (if j = i then topspace (X i) else {f j})"]
-      using f by fastforce
+      using f product_topology_homeomorphic_component [OF \<open>i \<in> I\<close>, of "\<lambda>j. subtopology (X j) (if j = i then topspace (X i) else {f j})"]
+      by (force simp add: subtopology_PiE)
     then show ?thesis
       using minor [OF f major \<open>i \<in> I\<close>] PQ by auto
   qed
--- a/src/HOL/Analysis/Further_Topology.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Further_Topology.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -2936,7 +2936,8 @@
         have "(Arg2pi x / (2*pi)) \<in> {0..1}" "(Arg2pi y / (2*pi)) \<in> {0..1}"
           using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+
         with eq show ?thesis
-          using \<open>simple_path \<gamma>\<close> Arg2pi_lt_2pi unfolding simple_path_def o_def
+          using \<open>simple_path \<gamma>\<close> Arg2pi_lt_2pi 
+          unfolding simple_path_def loop_free_def o_def
           by (metis eq_divide_eq_1 not_less_iff_gr_or_eq)
       qed
       with xy show "x = y"
--- a/src/HOL/Analysis/Homotopy.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Homotopy.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -4294,7 +4294,7 @@
   have "g 0 \<in> path_image g" "g (1/2) \<in> path_image g"
     by (simp_all add: path_defs)
   moreover have "g 0 \<noteq> g (1/2)"
-    using assms by (fastforce simp add: simple_path_def)
+    using assms by (fastforce simp add: simple_path_def loop_free_def)
   ultimately have "\<forall>a. \<not> path_image g \<subseteq> {a}"
     by blast
   then show ?thesis
--- a/src/HOL/Analysis/Jordan_Curve.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Jordan_Curve.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -203,14 +203,15 @@
       by (metis \<open>pathfinish h = g u\<close> pathfinish_def pathfinish_subpath u(3))
     show "path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h) = {a, b}"
     proof
-      show "path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h) \<subseteq> {a, b}"
-        using v  \<open>pathfinish (subpath v 1 h) = a\<close> \<open>simple_path h\<close>
-          apply (auto simp: simple_path_def path_image_subpath image_iff Ball_def)
-        by (metis (full_types) less_eq_real_def less_irrefl less_le_trans)
+      have "loop_free h"
+        using \<open>simple_path h\<close> simple_path_def by blast
+      then show "path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h) \<subseteq> {a, b}"
+        using v \<open>pathfinish (subpath v 1 h) = a\<close>
+        apply (clarsimp simp add: loop_free_def path_image_subpath Ball_def)
+        by (smt (verit))
       show "{a, b} \<subseteq> path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h)"
         using v \<open>pathstart (subpath 0 v h) = a\<close> \<open>pathfinish (subpath v 1 h) = a\<close>
-        apply (auto simp: path_image_subpath image_iff)
-        by (metis atLeastAtMost_iff order_refl)
+        by (auto simp: path_image_subpath image_iff Bex_def)
     qed
     show "path_image (subpath 0 v h) \<union> path_image (subpath v 1 h) = path_image g"
       using v apply (simp add: path_image_subpath pihg [symmetric])
--- a/src/HOL/Analysis/Lebesgue_Measure.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Lebesgue_Measure.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -145,11 +145,8 @@
         proof cases
           assume "?R"
           with \<open>j \<in> S\<close> psubset.prems have "{u..v} \<subseteq> (\<Union>i\<in>S-{j}. {l i<..< r i})"
-            apply (auto simp: subset_eq Ball_def)
-            apply (metis Diff_iff less_le_trans leD linear singletonD)
-            apply (metis Diff_iff less_le_trans leD linear singletonD)
-            apply (metis order_trans less_le_not_le linear)
-            done
+            apply (simp add: subset_eq Ball_def Bex_def)
+            by (metis order_le_less_trans order_less_le_trans order_less_not_sym)
           with \<open>j \<in> S\<close> have "F v - F u \<le> (\<Sum>i\<in>S - {j}. F (r i) - F (l i))"
             by (intro psubset) auto
           also have "\<dots> \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
@@ -162,19 +159,14 @@
             by (auto simp: not_less)
           let ?S1 = "{i \<in> S. l i < l j}"
           let ?S2 = "{i \<in> S. r i > r j}"
-
+          have *: "?S1 \<inter> ?S2 = {}"
+            using j by (fastforce simp add: disjoint_iff)
           have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<ge> (\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i))"
             using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
             by (intro sum_mono2) (auto intro: less_imp_le)
           also have "(\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i)) =
             (\<Sum>i\<in>?S1. F (r i) - F (l i)) + (\<Sum>i\<in>?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))"
-            using psubset(1) psubset.prems(1) j
-            apply (subst sum.union_disjoint)
-            apply simp_all
-            apply (subst sum.union_disjoint)
-            apply auto
-            apply (metis less_le_not_le)
-            done
+            using psubset(1) by (simp add: * sum.union_disjoint disjoint_iff_not_equal)
           also (xtrans) have "(\<Sum>i\<in>?S1. F (r i) - F (l i)) \<ge> F (l j) - F u"
             using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
             apply (intro psubset.IH psubset)
@@ -202,31 +194,24 @@
     proof
       fix i
       note right_cont_F [of "r i"]
+      then have "\<exists>d>0. F (r i + d) - F (r i) < epsilon / 2 ^ (i + 2)"
+        by (simp add: continuous_at_right_real_increasing egt0)
       thus "\<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
-        apply -
-        apply (subst (asm) continuous_at_right_real_increasing)
-        apply (rule mono_F, assumption)
-        apply (drule_tac x = "epsilon / 2 ^ (i + 2)" in spec)
-        apply (erule impE)
-        using egt0 by (auto simp add: field_simps)
+        by force
     qed
     then obtain delta where
         deltai_gt0: "\<And>i. delta i > 0" and
         deltai_prop: "\<And>i. F (r i + delta i) < F (r i) + epsilon / 2^(i+2)"
       by metis
     have "\<exists>a' > a. F a' - F a < epsilon / 2"
-      apply (insert right_cont_F [of a])
-      apply (subst (asm) continuous_at_right_real_increasing)
-      using mono_F apply force
-      apply (drule_tac x = "epsilon / 2" in spec)
-      using egt0 unfolding mult.commute [of 2] by force
+      using right_cont_F [of a]
+      by (metis continuous_at_right_real_increasing egt0 half_gt_zero less_add_same_cancel1 mono_F)
     then obtain a' where a'lea [arith]: "a' > a" and
       a_prop: "F a' - F a < epsilon / 2"
       by auto
     define S' where "S' = {i. l i < r i}"
-    obtain S :: "nat set" where
-      "S \<subseteq> S'" and finS: "finite S" and
-      Sprop: "{a'..b} \<subseteq> (\<Union>i \<in> S. {l i<..<r i + delta i})"
+    obtain S :: "nat set" where "S \<subseteq> S'" and finS: "finite S" 
+      and Sprop: "{a'..b} \<subseteq> (\<Union>i \<in> S. {l i<..<r i + delta i})"
     proof (rule compactE_image)
       show "compact {a'..b}"
         by (rule compact_Icc)
@@ -236,58 +221,34 @@
       also have "{a <.. b} = (\<Union>i\<in>S'. {l i<..r i})"
         unfolding lr_eq_ab[symmetric] by (fastforce simp add: S'_def intro: less_le_trans)
       also have "\<dots> \<subseteq> (\<Union>i \<in> S'. {l i<..<r i + delta i})"
-        apply (intro UN_mono)
-        apply (auto simp: S'_def)
-        apply (cut_tac i=i in deltai_gt0)
-        apply simp
-        done
+        by (intro UN_mono; simp add: add.commute add_strict_increasing deltai_gt0 subset_iff)
       finally show "{a'..b} \<subseteq> (\<Union>i \<in> S'. {l i<..<r i + delta i})" .
     qed
     with S'_def have Sprop2: "\<And>i. i \<in> S \<Longrightarrow> l i < r i" by auto
-    from finS have "\<exists>n. \<forall>i \<in> S. i \<le> n"
-      by (subst finite_nat_set_iff_bounded_le [symmetric])
-    then obtain n where Sbound [rule_format]: "\<forall>i \<in> S. i \<le> n" ..
-    have "F b - F a' \<le> (\<Sum>i\<in>S. F (r i + delta i) - F (l i))"
-      apply (rule claim2 [rule_format])
-      using finS Sprop apply auto
-      apply (frule Sprop2)
-      apply (subgoal_tac "delta i > 0")
-      apply arith
-      by (rule deltai_gt0)
-    also have "... \<le> (\<Sum>i \<in> S. F(r i) - F(l i) + epsilon / 2^(i+2))"
-      apply (rule sum_mono)
-      apply simp
-      apply (rule order_trans)
-      apply (rule less_imp_le)
-      apply (rule deltai_prop)
-      by auto
-    also have "... = (\<Sum>i \<in> S. F(r i) - F(l i)) +
-        (epsilon / 4) * (\<Sum>i \<in> S. (1 / 2)^i)" (is "_ = ?t + _")
-      by (subst sum.distrib) (simp add: field_simps sum_distrib_left)
-    also have "... \<le> ?t + (epsilon / 4) * (\<Sum> i < Suc n. (1 / 2)^i)"
-      apply (rule add_left_mono)
-      apply (rule mult_left_mono)
-      apply (rule sum_mono2)
-      using egt0 apply auto
-      by (frule Sbound, auto)
-    also have "... \<le> ?t + (epsilon / 2)"
-      apply (rule add_left_mono)
-      apply (subst geometric_sum)
-      apply auto
-      apply (rule mult_left_mono)
-      using egt0 apply auto
-      done
-    finally have aux2: "F b - F a' \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2"
-      by simp
-
+    obtain n where Sbound: "\<And>i. i \<in> S \<Longrightarrow> i \<le> n"
+      using Max_ge finS by blast
     have "F b - F a = (F b - F a') + (F a' - F a)"
       by auto
     also have "... \<le> (F b - F a') + epsilon / 2"
       using a_prop by (intro add_left_mono) simp
     also have "... \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2 + epsilon / 2"
-      apply (intro add_right_mono)
-      apply (rule aux2)
-      done
+    proof -
+      have "F b - F a' \<le> (\<Sum>i\<in>S. F (r i + delta i) - F (l i))"
+        using claim2 l_r Sprop by (simp add: deltai_gt0 finS add.commute add_strict_increasing)
+      also have "... \<le> (\<Sum>i \<in> S. F(r i) - F(l i) + epsilon / 2^(i+2))"
+        by (smt (verit) sum_mono deltai_prop)
+      also have "... = (\<Sum>i \<in> S. F(r i) - F(l i)) +
+        (epsilon / 4) * (\<Sum>i \<in> S. (1 / 2)^i)" (is "_ = ?t + _")
+        by (subst sum.distrib) (simp add: field_simps sum_distrib_left)
+      also have "... \<le> ?t + (epsilon / 4) * (\<Sum> i < Suc n. (1 / 2)^i)"
+        using egt0 Sbound by (intro add_left_mono mult_left_mono sum_mono2) force+
+      also have "... \<le> ?t + (epsilon / 2)"
+        using egt0 by (simp add: geometric_sum add_left_mono mult_left_mono)
+      finally have "F b - F a' \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2"
+        by simp
+      then show ?thesis
+        by linarith
+    qed
     also have "... = (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon"
       by auto
     also have "... \<le> (\<Sum>i\<le>n. F (r i) - F (l i)) + epsilon"
@@ -316,12 +277,8 @@
 
 lemma\<^marker>\<open>tag important\<close> sets_interval_measure [simp, measurable_cong]:
     "sets (interval_measure F) = sets borel"
-  apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc)
-  apply (rule sigma_sets_eqI)
-  apply auto
-  apply (case_tac "a \<le> ba")
-  apply (auto intro: sigma_sets.Empty)
-  done
+  apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc image_def split: prod.split)
+  by (metis greaterThanAtMost_empty nle_le)
 
 lemma space_interval_measure [simp]: "space (interval_measure F) = UNIV"
   by (simp add: interval_measure_def space_extend_measure)
@@ -342,21 +299,16 @@
   show "((\<lambda>a. F b - F a) \<longlongrightarrow> emeasure ?F {a..b}) (at_left a)"
   proof (rule tendsto_at_left_sequentially)
     show "a - 1 < a" by simp
-    fix X assume "\<And>n. X n < a" "incseq X" "X \<longlonglongrightarrow> a"
-    with \<open>a \<le> b\<close> have "(\<lambda>n. emeasure ?F {X n<..b}) \<longlonglongrightarrow> emeasure ?F (\<Inter>n. {X n <..b})"
-      apply (intro Lim_emeasure_decseq)
-      apply (auto simp: decseq_def incseq_def emeasure_interval_measure_Ioc *)
-      apply force
-      apply (subst (asm ) *)
-      apply (auto intro: less_le_trans less_imp_le)
-      done
+    fix X assume X: "\<And>n. X n < a" "incseq X" "X \<longlonglongrightarrow> a"
+    then have "emeasure (interval_measure F) {X n<..b} \<noteq> \<infinity>" for n
+      by (smt (verit) "*" \<open>a \<le> b\<close> ennreal_neq_top infinity_ennreal_def)
+    with X have "(\<lambda>n. emeasure ?F {X n<..b}) \<longlonglongrightarrow> emeasure ?F (\<Inter>n. {X n <..b})"
+      by (intro Lim_emeasure_decseq; force simp: decseq_def incseq_def emeasure_interval_measure_Ioc *)
     also have "(\<Inter>n. {X n <..b}) = {a..b}"
-      using \<open>\<And>n. X n < a\<close>
       apply auto
       apply (rule LIMSEQ_le_const2[OF \<open>X \<longlonglongrightarrow> a\<close>])
-      apply (auto intro: less_imp_le)
-      apply (auto intro: less_le_trans)
-      done
+      using less_eq_real_def apply presburger
+      using X(1) order_less_le_trans by blast
     also have "(\<lambda>n. emeasure ?F {X n<..b}) = (\<lambda>n. F b - F (X n))"
       using \<open>\<And>n. X n < a\<close> \<open>a \<le> b\<close> by (subst *) (auto intro: less_imp_le less_le_trans)
     finally show "(\<lambda>n. F b - F (X n)) \<longlonglongrightarrow> emeasure ?F {a..b}" .
@@ -426,12 +378,7 @@
   by (metis (mono_tags, lifting) IntD1 assms measurable_cong_simp space_restrict_space)
 
 lemma lebesgue_on_UNIV_eq: "lebesgue_on UNIV = lebesgue"
-proof -
-  have "measure_of UNIV (sets lebesgue) (emeasure lebesgue) = lebesgue"
-    by (metis measure_of_of_measure space_borel space_completion space_lborel)
-  then show ?thesis
-    by (auto simp: restrict_space_def)
-qed
+  by (simp add: emeasure_restrict_space measure_eqI)
 
 lemma integral_restrict_Int:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -528,11 +475,7 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable lebesgue"
   shows "f \<in> borel_measurable (lebesgue_on S)"
-  using assms
-  apply (simp add: in_borel_measurable_borel Ball_def)
-  apply (elim all_forward imp_forward asm_rl)
-  apply (force simp: space_restrict_space sets_restrict_space image_iff intro: rev_bexI)
-  done
+  using assms by (smt (verit) measurable_lebesgue_cong measurable_restrict_space1)
 
 lemma borel_measurable_if:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -557,33 +500,25 @@
 lemma borel_measurable_vimage_halfspace_component_lt:
      "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
       (\<forall>a i. i \<in> Basis \<longrightarrow> {x \<in> S. f x \<bullet> i < a} \<in> sets (lebesgue_on S))"
-  apply (rule trans [OF borel_measurable_iff_halfspace_less])
-  apply (fastforce simp add: space_restrict_space)
-  done
+  by (force simp add: space_restrict_space trans [OF borel_measurable_iff_halfspace_less])
 
 lemma borel_measurable_vimage_halfspace_component_ge:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
          (\<forall>a i. i \<in> Basis \<longrightarrow> {x \<in> S. f x \<bullet> i \<ge> a} \<in> sets (lebesgue_on S))"
-  apply (rule trans [OF borel_measurable_iff_halfspace_ge])
-  apply (fastforce simp add: space_restrict_space)
-  done
+  by (force simp add: space_restrict_space trans [OF borel_measurable_iff_halfspace_ge])
 
 lemma borel_measurable_vimage_halfspace_component_gt:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
          (\<forall>a i. i \<in> Basis \<longrightarrow> {x \<in> S. f x \<bullet> i > a} \<in> sets (lebesgue_on S))"
-  apply (rule trans [OF borel_measurable_iff_halfspace_greater])
-  apply (fastforce simp add: space_restrict_space)
-  done
+  by (force simp add: space_restrict_space trans [OF borel_measurable_iff_halfspace_greater])
 
 lemma borel_measurable_vimage_halfspace_component_le:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
          (\<forall>a i. i \<in> Basis \<longrightarrow> {x \<in> S. f x \<bullet> i \<le> a} \<in> sets (lebesgue_on S))"
-  apply (rule trans [OF borel_measurable_iff_halfspace_le])
-  apply (fastforce simp add: space_restrict_space)
-  done
+  by (force simp add: space_restrict_space trans [OF borel_measurable_iff_halfspace_le])
 
 lemma
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -630,16 +565,12 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
          (\<forall>T. closed T \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets (lebesgue_on S))"
-        (is "?lhs = ?rhs")
 proof -
-  have eq: "{x \<in> S. f x \<in> T} = S - {x \<in> S. f x \<in> (- T)}" for T
+  have eq: "{x \<in> S. f x \<in> T} = S - (S \<inter> f -` (- T))" for T
     by auto
   show ?thesis
-    apply (simp add: borel_measurable_vimage_open, safe)
-     apply (simp_all (no_asm) add: eq)
-     apply (intro sets.Diff sets_lebesgue_on_refl, force simp: closed_open)
-    apply (intro sets.Diff sets_lebesgue_on_refl, force simp: open_closed)
-    done
+    unfolding borel_measurable_vimage_open eq
+    by (metis Diff_Diff_Int closed_Compl diff_eq open_Compl sets.Diff sets_lebesgue_on_refl vimage_Compl)
 qed
 
 lemma borel_measurable_vimage_closed_interval:
@@ -689,9 +620,7 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   shows "f \<in> borel_measurable lebesgue \<longleftrightarrow>
          (\<forall>T. T \<in> sets borel \<longrightarrow> {x. f x \<in> T} \<in> sets lebesgue)"
-  apply (intro iffI allI impI lebesgue_measurable_vimage_borel)
-    apply (auto simp: in_borel_measurable_borel vimage_def)
-  done
+  by (smt (verit, best) Collect_cong UNIV_I borel_measurable_vimage_borel lebesgue_on_UNIV_eq)
 
 
 subsection \<^marker>\<open>tag unimportant\<close> \<open>Measurability of continuous functions\<close>
@@ -699,13 +628,9 @@
 lemma continuous_imp_measurable_on_sets_lebesgue:
   assumes f: "continuous_on S f" and S: "S \<in> sets lebesgue"
   shows "f \<in> borel_measurable (lebesgue_on S)"
-proof -
-  have "sets (restrict_space borel S) \<subseteq> sets (lebesgue_on S)"
-    by (simp add: mono_restrict_space subsetI)
-  then show ?thesis
-    by (simp add: borel_measurable_continuous_on_restrict [OF f] borel_measurable_subalgebra 
-                  space_restrict_space)
-qed
+  by (metis borel_measurable_continuous_on_restrict borel_measurable_subalgebra f 
+      lebesgue_on_UNIV_eq mono_restrict_space sets_completionI_sets sets_lborel space_borel 
+      space_lebesgue_on space_restrict_space subsetI)
 
 lemma id_borel_measurable_lebesgue [iff]: "id \<in> borel_measurable lebesgue"
   by (simp add: measurable_completion)
@@ -741,12 +666,7 @@
   fixes l u :: real
   assumes [simp]: "l \<le> u"
   shows "emeasure lborel {l .. u} = u - l"
-proof -
-  have "((\<lambda>f. f 1) -` {l..u} \<inter> space (Pi\<^sub>M {1} (\<lambda>b. interval_measure (\<lambda>x. x)))) = {1::real} \<rightarrow>\<^sub>E {l..u}"
-    by (auto simp: space_PiM)
-  then show ?thesis
-    by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc)
-qed
+  by (simp add: emeasure_interval_measure_Icc lborel_eq_real)
 
 lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \<le> u then u - l else 0)"
   by simp
@@ -781,12 +701,7 @@
 lemma emeasure_lborel_Ioc[simp]:
   assumes [simp]: "l \<le> u"
   shows "emeasure lborel {l <.. u} = ennreal (u - l)"
-proof -
-  have "emeasure lborel {l <.. u} = emeasure lborel {l .. u}"
-    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
-  then show ?thesis
-    by simp
-qed
+  by (simp add: emeasure_interval_measure_Ioc lborel_eq_real)
 
 lemma emeasure_lborel_Ico[simp]:
   assumes [simp]: "l \<le> u"
@@ -830,26 +745,12 @@
   by (auto simp: emeasure_lborel_box_eq)
 
 lemma emeasure_lborel_ball_finite: "emeasure lborel (ball c r) < \<infinity>"
-proof -
-  have "bounded (ball c r)" by simp
-  from bounded_subset_cbox_symmetric[OF this] obtain a where a: "ball c r \<subseteq> cbox (-a) a"
-    by auto
-  hence "emeasure lborel (ball c r) \<le> emeasure lborel (cbox (-a) a)"
-    by (intro emeasure_mono) auto
-  also have "\<dots> < \<infinity>" by (simp add: emeasure_lborel_cbox_eq)
-  finally show ?thesis .
-qed
+  by (metis bounded_ball bounded_subset_cbox_symmetric cbox_borel emeasure_lborel_cbox_finite
+      emeasure_mono order_le_less_trans sets_lborel)
 
 lemma emeasure_lborel_cball_finite: "emeasure lborel (cball c r) < \<infinity>"
-proof -
-  have "bounded (cball c r)" by simp
-  from bounded_subset_cbox_symmetric[OF this] obtain a where a: "cball c r \<subseteq> cbox (-a) a"
-    by auto
-  hence "emeasure lborel (cball c r) \<le> emeasure lborel (cbox (-a) a)"
-    by (intro emeasure_mono) auto
-  also have "\<dots> < \<infinity>" by (simp add: emeasure_lborel_cbox_eq)
-  finally show ?thesis .
-qed
+  by (metis bounded_cball bounded_subset_cbox_symmetric cbox_borel emeasure_lborel_cbox_finite 
+            emeasure_mono order_le_less_trans sets_lborel)
 
 lemma fmeasurable_cbox [iff]: "cbox a b \<in> fmeasurable lborel"
   and fmeasurable_box [iff]: "box a b \<in> fmeasurable lborel"
@@ -924,7 +825,7 @@
   also have "emeasure lborel (\<Union>i. {from_nat_into A i}) = 0"
     by (rule emeasure_UN_eq_0) auto
   finally show ?thesis
-    by (auto simp add: )
+    by simp
 qed
 
 lemma countable_imp_null_set_lborel: "countable A \<Longrightarrow> A \<in> null_sets lborel"
@@ -934,30 +835,18 @@
   by (intro countable_imp_null_set_lborel countable_finite)
 
 lemma insert_null_sets_iff [simp]: "insert a N \<in> null_sets lebesgue \<longleftrightarrow> N \<in> null_sets lebesgue"     
-    (is "?lhs = ?rhs")
-proof
-  assume ?lhs then show ?rhs
-    by (meson completion.complete2 subset_insertI)
-next
-  assume ?rhs then show ?lhs
-  by (simp add: null_sets.insert_in_sets null_setsI)
-qed
+  by (meson completion.complete2 finite.simps finite_imp_null_set_lborel null_sets.insert_in_sets 
+      null_sets_completionI subset_insertI)
 
 lemma insert_null_sets_lebesgue_on_iff [simp]:
   assumes "a \<in> S" "S \<in> sets lebesgue"
   shows "insert a N \<in> null_sets (lebesgue_on S) \<longleftrightarrow> N \<in> null_sets (lebesgue_on S)"     
   by (simp add: assms null_sets_restrict_space)
 
-lemma lborel_neq_count_space[simp]: "lborel \<noteq> count_space (A::('a::ordered_euclidean_space) set)"
-proof
-  assume asm: "lborel = count_space A"
-  have "space lborel = UNIV" by simp
-  hence [simp]: "A = UNIV" by (subst (asm) asm) (simp only: space_count_space)
-  have "emeasure lborel {undefined::'a} = 1"
-      by (subst asm, subst emeasure_count_space_finite) auto
-  moreover have "emeasure lborel {undefined} \<noteq> 1" by simp
-  ultimately show False by contradiction
-qed
+lemma lborel_neq_count_space[simp]: 
+  fixes A :: "('a::ordered_euclidean_space) set"
+  shows "lborel \<noteq> count_space A"
+  by (metis finite.simps finite_imp_null_set_lborel insert_not_empty null_sets_count_space singleton_iff)
 
 lemma mem_closed_if_AE_lebesgue_open:
   assumes "open S" "closed C"
@@ -1004,13 +893,9 @@
   let ?A = "\<lambda>n::nat. box (- (real n *\<^sub>R One)) (real n *\<^sub>R One) :: 'a set"
   show "range ?A \<subseteq> ?E" "(\<Union>i. ?A i) = UNIV"
     unfolding UN_box_eq_UNIV by auto
-
-  { fix i show "emeasure lborel (?A i) \<noteq> \<infinity>" by auto }
-  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
-      apply (auto simp: emeasure_eq emeasure_lborel_box_eq)
-      apply (subst box_eq_empty(1)[THEN iffD2])
-      apply (auto intro: less_imp_le simp: not_le)
-      done }
+  show "emeasure lborel (?A i) \<noteq> \<infinity>" for i by auto 
+  show "emeasure lborel X = emeasure M X" if "X \<in> ?E" for X
+    using that box_eq_empty(1) by (fastforce simp: emeasure_eq emeasure_lborel_box_eq)
 qed
 
 lemma\<^marker>\<open>tag important\<close> lborel_affine_euclidean:
@@ -1102,7 +987,7 @@
     by (auto simp: null_sets_def)
 
   show "T \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
-    by (rule completion.measurable_completion2) (auto simp: eq measurable_completion)
+    by (simp add: completion.measurable_completion2 eq measurable_completion)
 
   have "lebesgue = completion (density (distr lborel borel T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>)))"
     using c by (subst lborel_affine_euclidean[of c t]) (simp_all add: T_def[abs_def])
@@ -1135,7 +1020,7 @@
 
 lemma lebesgue_measurable_scaling[measurable]: "(*\<^sub>R) x \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
 proof cases
-  assume "x = 0"
+  assume "x = 0" 
   then have "(*\<^sub>R) x = (\<lambda>x. 0::'a)"
     by (auto simp: fun_eq_iff)
   then show ?thesis by auto
@@ -1352,8 +1237,7 @@
 lemma
   fixes a::real
   shows lmeasurable_interval [iff]: "{a..b} \<in> lmeasurable" "{a<..<b} \<in> lmeasurable"
-  apply (metis box_real(2) lmeasurable_cbox)
-  by (metis box_real(1) lmeasurable_box)
+  by (metis box_real lmeasurable_box lmeasurable_cbox)+
 
 lemma fmeasurable_compact: "compact S \<Longrightarrow> S \<in> fmeasurable lborel"
   using emeasure_compact_finite[of S] by (intro fmeasurableI) (auto simp: borel_compact)
@@ -1387,14 +1271,7 @@
   by (simp add: bounded_interior lmeasurable_open)
 
 lemma null_sets_cbox_Diff_box: "cbox a b - box a b \<in> null_sets lborel"
-proof -
-  have "emeasure lborel (cbox a b - box a b) = 0"
-    by (subst emeasure_Diff) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq box_subset_cbox)
-  then have "cbox a b - box a b \<in> null_sets lborel"
-    by (auto simp: null_sets_def)
-  then show ?thesis
-    by (auto dest!: AE_not_in)
-qed
+  by (simp add: emeasure_Diff emeasure_lborel_box_eq emeasure_lborel_cbox_eq null_setsI subset_box)
 
 lemma bounded_set_imp_lmeasurable:
   assumes "bounded S" "S \<in> sets lebesgue" shows "S \<in> lmeasurable"
@@ -1423,10 +1300,8 @@
     by (simp add: sigma_sets.Empty)
 next
   case (Compl a)
-  then have "\<Omega> - a \<subseteq> \<Omega>" "a \<subseteq> \<Omega>"
-    by (auto simp: sigma_sets_into_sp [OF \<open>M \<subseteq> Pow \<Omega>\<close>])
-  then show ?case
-    by (auto simp: inj_on_image_set_diff [OF \<open>inj_on f \<Omega>\<close>] assms intro: Compl sigma_sets.Compl)
+  with assms show ?case
+    by (metis inj_on_image_set_diff sigma_sets.Compl sigma_sets_into_sp)
 next
   case (Union a) then show ?case
     by (metis image_UN sigma_sets.simps)
@@ -1475,9 +1350,8 @@
 lemma measurable_translation:
    "S \<in> lmeasurable \<Longrightarrow> ((+) a ` S) \<in> lmeasurable"
   using emeasure_lebesgue_affine [of 1 a S]
-  apply (auto intro: lebesgue_sets_translation simp add: fmeasurable_def cong: image_cong_simp)
-  apply (simp add: ac_simps)
-  done
+  by (smt (verit, best) add.commute ennreal_1 fmeasurable_def image_cong lambda_one 
+          lebesgue_sets_translation mem_Collect_eq power_one scaleR_one)
 
 lemma measurable_translation_subtract:
    "S \<in> lmeasurable \<Longrightarrow> ((\<lambda>x. x - a) ` S) \<in> lmeasurable"
@@ -1508,11 +1382,10 @@
 
   let ?f = "\<lambda>n. root DIM('a) (Suc n)"
 
-  have vimage_eq_image: "(*\<^sub>R) (?f n) -` S = (*\<^sub>R) (1 / ?f n) ` S" for n
-    apply safe
-    subgoal for x by (rule image_eqI[of _ _ "?f n *\<^sub>R x"]) auto
-    subgoal by auto
-    done
+  have "?f n *\<^sub>R x \<in> S \<Longrightarrow> x \<in> (*\<^sub>R) (1 / ?f n) ` S" for x n
+    by (rule image_eqI[of _ _ "?f n *\<^sub>R x"]) auto
+  then have vimage_eq_image: "(*\<^sub>R) (?f n) -` S = (*\<^sub>R) (1 / ?f n) ` S" for n
+    by auto
 
   have eq: "(1 / ?f n) ^ DIM('a) = 1 / Suc n" for n
     by (simp add: field_simps)
@@ -1642,12 +1515,12 @@
                    "emeasure (completion M) A = emeasure M A'"
 proof -
   from AE_notin_null_part[OF A] obtain N where N: "N \<in> null_sets M" "null_part M A \<subseteq> N"
-    unfolding eventually_ae_filter using null_part_null_sets[OF A, THEN null_setsD2, THEN sets.sets_into_space] by auto
+    by (meson assms null_part)
   let ?A' = "main_part M A \<union> N"
   show ?thesis
   proof
     show "A \<subseteq> ?A'"
-      using \<open>null_part M A \<subseteq> N\<close> by (subst main_part_null_part_Un[symmetric, OF A]) auto
+      using \<open>null_part M A \<subseteq> N\<close> assms main_part_null_part_Un by blast
     have "main_part M A \<subseteq> A"
       using assms main_part_null_part_Un by auto
     then have "?A' - A \<subseteq> N"
--- a/src/HOL/Analysis/Linear_Algebra.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Linear_Algebra.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -68,38 +68,17 @@
   fixes a :: "'a::ab_group_add"
   assumes "(\<lambda>x. a + x) ` A = (\<lambda>x. a + x) ` B"
   shows "A = B"
-proof -
-  have "(\<lambda>x. -a + x) ` ((\<lambda>x. a + x) ` A) = (\<lambda>x. - a + x) ` ((\<lambda>x. a + x) ` B)"
-    using assms by auto
-  then show ?thesis
-    using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto
-qed
+  using assms translation_assoc by fastforce
 
 lemma translation_galois:
   fixes a :: "'a::ab_group_add"
   shows "T = ((\<lambda>x. a + x) ` S) \<longleftrightarrow> S = ((\<lambda>x. (- a) + x) ` T)"
-  using translation_assoc[of "-a" a S]
-  apply auto
-  using translation_assoc[of a "-a" T]
-  apply auto
-  done
+  by (metis add.right_inverse group_cancel.rule0 translation_invert translation_assoc)
 
 lemma translation_inverse_subset:
   assumes "((\<lambda>x. - a + x) ` V) \<le> (S :: 'n::ab_group_add set)"
   shows "V \<le> ((\<lambda>x. a + x) ` S)"
-proof -
-  {
-    fix x
-    assume "x \<in> V"
-    then have "x-a \<in> S" using assms by auto
-    then have "x \<in> {a + v |v. v \<in> S}"
-      apply auto
-      apply (rule exI[of _ "x-a"], simp)
-      done
-    then have "x \<in> ((\<lambda>x. a+x) ` S)" by auto
-  }
-  then show ?thesis by auto
-qed
+  by (metis assms subset_image_iff translation_galois)
 
 subsection\<^marker>\<open>tag unimportant\<close> \<open>More interesting properties of the norm\<close>
 
@@ -120,41 +99,25 @@
 qed
 
 lemma vector_eq: "x = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y \<and> y \<bullet> y = x \<bullet> x"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs by simp
-next
-  assume ?rhs
-  then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y \<bullet> y = 0"
-    by simp
-  then have "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0"
-    by (simp add: inner_diff inner_commute)
-  then have "(x - y) \<bullet> (x - y) = 0"
-    by (simp add: field_simps inner_diff inner_commute)
-  then show "x = y" by simp
-qed
+  by (metis (no_types, opaque_lifting) inner_commute inner_diff_right inner_eq_zero_iff right_minus_eq)
 
 lemma norm_triangle_half_r:
-  "norm (y - x1) < e / 2 \<Longrightarrow> norm (y - x2) < e / 2 \<Longrightarrow> norm (x1 - x2) < e"
+  "norm (y - x1) < e/2 \<Longrightarrow> norm (y - x2) < e/2 \<Longrightarrow> norm (x1 - x2) < e"
   using dist_triangle_half_r unfolding dist_norm[symmetric] by auto
 
 lemma norm_triangle_half_l:
-  assumes "norm (x - y) < e / 2"
-    and "norm (x' - y) < e / 2"
+  assumes "norm (x - y) < e/2" and "norm (x' - y) < e/2"
   shows "norm (x - x') < e"
-  using dist_triangle_half_l[OF assms[unfolded dist_norm[symmetric]]]
-  unfolding dist_norm[symmetric] .
+  by (metis assms dist_norm dist_triangle_half_l)
 
 lemma abs_triangle_half_r:
   fixes y :: "'a::linordered_field"
-  shows "abs (y - x1) < e / 2 \<Longrightarrow> abs (y - x2) < e / 2 \<Longrightarrow> abs (x1 - x2) < e"
+  shows "abs (y - x1) < e/2 \<Longrightarrow> abs (y - x2) < e/2 \<Longrightarrow> abs (x1 - x2) < e"
   by linarith
 
 lemma abs_triangle_half_l:
   fixes y :: "'a::linordered_field"
-  assumes "abs (x - y) < e / 2"
-    and "abs (x' - y) < e / 2"
+  assumes "abs (x - y) < e/2" and "abs (x' - y) < e/2"
   shows "abs (x - x') < e"
   using assms by linarith
 
@@ -163,41 +126,15 @@
     and "finite S \<Longrightarrow> sum f (insert x S) = (if x \<in> S then sum f S else f x + sum f S)"
   by (auto simp add: insert_absorb)
 
-lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z"
-proof
-  assume "\<forall>x. x \<bullet> y = x \<bullet> z"
-  then have "\<forall>x. x \<bullet> (y - z) = 0"
-    by (simp add: inner_diff)
-  then have "(y - z) \<bullet> (y - z) = 0" ..
-  then show "y = z" by simp
-qed simp
-
-lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = y"
-proof
-  assume "\<forall>z. x \<bullet> z = y \<bullet> z"
-  then have "\<forall>z. (x - y) \<bullet> z = 0"
-    by (simp add: inner_diff)
-  then have "(x - y) \<bullet> (x - y) = 0" ..
-  then show "x = y" by simp
-qed simp
+lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z" and vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = y"
+  by (metis inner_commute vector_eq)+
 
 subsection \<open>Substandard Basis\<close>
 
 lemma ex_card:
   assumes "n \<le> card A"
   shows "\<exists>S\<subseteq>A. card S = n"
-proof (cases "finite A")
-  case True
-  from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..<card A} A" ..
-  moreover from f \<open>n \<le> card A\<close> have "{..< n} \<subseteq> {..< card A}" "inj_on f {..< n}"
-    by (auto simp: bij_betw_def intro: subset_inj_on)
-  ultimately have "f ` {..< n} \<subseteq> A" "card (f ` {..< n}) = n"
-    by (auto simp: bij_betw_def card_image)
-  then show ?thesis by blast
-next
-  case False
-  with \<open>n \<le> card A\<close> show ?thesis by force
-qed
+  by (meson assms obtain_subset_with_card_n)
 
 lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0)}"
   by (auto simp: subspace_def inner_add_left)
@@ -271,7 +208,7 @@
 
 lemma norm_add_Pythagorean:
   assumes "orthogonal a b"
-    shows "norm(a + b) ^ 2 = norm a ^ 2 + norm b ^ 2"
+    shows "(norm (a + b))\<^sup>2 = (norm a)\<^sup>2 + (norm b)\<^sup>2"
 proof -
   from assms have "(a - (0 - b)) \<bullet> (a - (0 - b)) = a \<bullet> a - (0 - b \<bullet> b)"
     by (simp add: algebra_simps orthogonal_def inner_commute)
@@ -300,12 +237,7 @@
 
 lemma\<^marker>\<open>tag unimportant\<close>  orthogonal_transformation:
   "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v. norm (f v) = norm v)"
-  unfolding orthogonal_transformation_def
-  apply auto
-  apply (erule_tac x=v in allE)+
-  apply (simp add: norm_eq_sqrt_inner)
-  apply (simp add: dot_norm linear_add[symmetric])
-  done
+  by (smt (verit, ccfv_threshold) dot_norm linear_add norm_eq_sqrt_inner orthogonal_transformation_def)
 
 lemma\<^marker>\<open>tag unimportant\<close>  orthogonal_transformation_id [simp]: "orthogonal_transformation (\<lambda>x. x)"
   by (simp add: linear_iff orthogonal_transformation_def)
@@ -431,15 +363,8 @@
 next
   fix h
   assume "\<forall>x y. inner (f x) y = inner x (h y)"
-  then have "\<forall>x y. inner x (g y) = inner x (h y)"
-    using assms by simp
-  then have "\<forall>x y. inner x (g y - h y) = 0"
-    by (simp add: inner_diff_right)
-  then have "\<forall>y. inner (g y - h y) (g y - h y) = 0"
-    by simp
-  then have "\<forall>y. h y = g y"
-    by simp
-  then show "h = g" by (simp add: ext)
+  then show "h = g"
+    by (metis assms ext vector_eq_ldot) 
 qed
 
 text \<open>TODO: The following lemmas about adjoints should hold for any
@@ -524,32 +449,18 @@
     proof -
       from Basis_le_norm[OF that, of x]
       show "norm (?g i) \<le> norm (f i) * norm x"
-        unfolding norm_scaleR  by (metis mult.commute mult_left_mono norm_ge_zero)
+        unfolding norm_scaleR by (metis mult.commute mult_left_mono norm_ge_zero)
     qed
     from sum_norm_le[of _ ?g, OF th]
     show "norm (f x) \<le> ?B * norm x"
-      unfolding th0 sum_distrib_right by metis
+      by (simp add: sum_distrib_right th0)
   qed
 qed
 
 lemma linear_conv_bounded_linear:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
   shows "linear f \<longleftrightarrow> bounded_linear f"
-proof
-  assume "linear f"
-  then interpret f: linear f .
-  show "bounded_linear f"
-  proof
-    have "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
-      using \<open>linear f\<close> by (rule linear_bounded)
-    then show "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
-      by (simp add: mult.commute)
-  qed
-next
-  assume "bounded_linear f"
-  then interpret f: bounded_linear f .
-  show "linear f" ..
-qed
+  by (metis mult.commute bounded_linear_axioms.intro bounded_linear_def linear_bounded)
 
 lemmas linear_linear = linear_conv_bounded_linear[symmetric]
 
@@ -562,17 +473,11 @@
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
   assumes lf: "linear f"
  obtains B where "B > 0" "\<And>x. norm (f x) \<le> B * norm x"
-proof -
-  have "\<exists>B > 0. \<forall>x. norm (f x) \<le> norm x * B"
-    using lf unfolding linear_conv_bounded_linear
-    by (rule bounded_linear.pos_bounded)
-  with that show ?thesis
-    by (auto simp: mult.commute)
-qed
+  by (metis bounded_linear.pos_bounded lf linear_linear mult.commute)
 
 lemma linear_invertible_bounded_below_pos:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
-  assumes "linear f" "linear g" "g \<circ> f = id"
+  assumes "linear f" "linear g" and gf: "g \<circ> f = id"
   obtains B where "B > 0" "\<And>x. B * norm x \<le> norm(f x)"
 proof -
   obtain B where "B > 0" and B: "\<And>x. norm (g x) \<le> B * norm x"
@@ -582,13 +487,8 @@
     show "0 < 1/B"
       by (simp add: \<open>B > 0\<close>)
     show "1/B * norm x \<le> norm (f x)" for x
-    proof -
-      have "1/B * norm x = 1/B * norm (g (f x))"
-        using assms by (simp add: pointfree_idE)
-      also have "\<dots> \<le> norm (f x)"
-        using B [of "f x"] by (simp add: \<open>B > 0\<close> mult.commute pos_divide_le_eq)
-      finally show ?thesis .
-    qed
+      by (smt (verit, ccfv_SIG) B \<open>0 < B\<close> gf comp_apply divide_inverse id_apply inverse_eq_divide 
+              less_divide_eq mult.commute)
   qed
 qed
 
@@ -663,15 +563,10 @@
   fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector"
   assumes bh: "bilinear h"
   shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
-proof -
-  have "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> norm x * norm y * B"
-    using bh [unfolded bilinear_conv_bounded_bilinear]
-    by (rule bounded_bilinear.pos_bounded)
-  then show ?thesis
-    by (simp only: ac_simps)
-qed
+  by (metis mult.assoc bh bilinear_conv_bounded_bilinear bounded_bilinear.pos_bounded mult.commute)
 
-lemma bounded_linear_imp_has_derivative: "bounded_linear f \<Longrightarrow> (f has_derivative f) net"
+lemma bounded_linear_imp_has_derivative: 
+  "bounded_linear f \<Longrightarrow> (f has_derivative f) net"
   by (auto simp add: has_derivative_def linear_diff linear_linear linear_def
       dest: bounded_linear.linear)
 
@@ -723,8 +618,7 @@
   assumes "pairwise orthogonal S"
     and "\<And>y. y \<in> S \<Longrightarrow> orthogonal x y"
   shows "pairwise orthogonal (insert x S)"
-  using assms unfolding pairwise_def
-  by (auto simp add: orthogonal_commute)
+  using assms by (auto simp: pairwise_def orthogonal_commute)
 
 lemma basis_orthogonal:
   fixes B :: "'a::real_inner set"
@@ -735,9 +629,7 @@
 proof (induct rule: finite_induct)
   case empty
   then show ?case
-    apply (rule exI[where x="{}"])
-    apply (auto simp add: pairwise_def)
-    done
+    using pairwise_empty by blast
 next
   case (insert a B)
   note fB = \<open>finite B\<close> and aB = \<open>a \<notin> B\<close>
@@ -748,21 +640,15 @@
   let ?C = "insert ?a C"
   from C(1) have fC: "finite ?C"
     by simp
-  from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)"
-    by (simp add: card_insert_if)
+  have cC: "card ?C \<le> card (insert a B)"
+    using C aB card_insert_if local.insert(1) by fastforce
   {
     fix x k
     have th0: "\<And>(a::'a) b c. a - (b - c) = c + (a - b)"
       by (simp add: field_simps)
     have "x - k *\<^sub>R (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x)) \<in> span C \<longleftrightarrow> x - k *\<^sub>R a \<in> span C"
-      apply (simp only: scaleR_right_diff_distrib th0)
-      apply (rule span_add_eq)
-      apply (rule span_scale)
-      apply (rule span_sum)
-      apply (rule span_scale)
-      apply (rule span_base)
-      apply assumption
-      done
+      unfolding scaleR_right_diff_distrib th0
+      by (intro span_add_eq span_scale span_sum span_base)
   }
   then have SC: "span ?C = span (insert a B)"
     unfolding set_eq_iff span_breakdown_eq C(3)[symmetric] by auto
@@ -773,15 +659,14 @@
       by blast
     have fth: "finite (C - {y})"
       using C by simp
-    have "orthogonal ?a y"
+    have "y \<noteq> 0 \<Longrightarrow> \<forall>x\<in>C - {y}. x \<bullet> a * (x \<bullet> y) / (x \<bullet> x) = 0"
+      using \<open>pairwise orthogonal C\<close>
+      by (metis Cy DiffE div_0 insertCI mult_zero_right orthogonal_def pairwise_insert)
+    then have "orthogonal ?a y"
       unfolding orthogonal_def
       unfolding inner_diff inner_sum_left right_minus_eq
       unfolding sum.remove [OF \<open>finite C\<close> \<open>y \<in> C\<close>]
-      apply (clarsimp simp add: inner_commute[of y a])
-      apply (rule sum.neutral)
-      apply clarsimp
-      apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
-      using \<open>y \<in> C\<close> by auto
+      by (auto simp add: sum.neutral inner_commute[of y a])
   }
   with \<open>pairwise orthogonal C\<close> have CPO: "pairwise orthogonal ?C"
     by (rule pairwise_orthogonal_insert)
@@ -792,8 +677,7 @@
 
 lemma orthogonal_basis_exists:
   fixes V :: "('a::euclidean_space) set"
-  shows "\<exists>B. independent B \<and> B \<subseteq> span V \<and> V \<subseteq> span B \<and>
-  (card B = dim V) \<and> pairwise orthogonal B"
+  shows "\<exists>B. independent B \<and> B \<subseteq> span V \<and> V \<subseteq> span B \<and> (card B = dim V) \<and> pairwise orthogonal B"
 proof -
   from basis_exists[of V] obtain B where
     B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "card B = dim V"
@@ -807,18 +691,15 @@
     by (metis span_superset span_mono subset_trans)
   from span_mono[OF B(3)] C have SVC: "span V \<subseteq> span C"
     by (simp add: span_span)
-  from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB
-  have iC: "independent C"
-    by (simp)
   from C fB have "card C \<le> dim V"
     by simp
   moreover have "dim V \<le> card C"
     using span_card_ge_dim[OF CSV SVC C(1)]
     by simp
-  ultimately have CdV: "card C = dim V"
+  ultimately have "card C = dim V"
     using C(1) by simp
-  from C B CSV CdV iC show ?thesis
-    by auto
+  with C B CSV show ?thesis
+    by (metis SVC card_eq_dim dim_span)
 qed
 
 text \<open>Low-dimensional subset is in a hyperplane (weak orthogonal complement).\<close>
@@ -831,22 +712,15 @@
   from sU obtain a where a: "a \<notin> span S"
     by blast
   from orthogonal_basis_exists obtain B where
-    B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B"
-    "card B = dim S" "pairwise orthogonal B"
+    B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B" "card B = dim S" "pairwise orthogonal B"
     by blast
   from B have fB: "finite B" "card B = dim S"
     using independent_bound by auto
-  from span_mono[OF B(2)] span_mono[OF B(3)]
   have sSB: "span S = span B"
-    by (simp add: span_span)
+    by (simp add: B span_eq)
   let ?a = "a - sum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B"
   have "sum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B \<in> span S"
-    unfolding sSB
-    apply (rule span_sum)
-    apply (rule span_scale)
-    apply (rule span_base)
-    apply assumption
-    done
+    by (simp add: sSB span_base span_mul span_sum)
   with a have a0:"?a  \<noteq> 0"
     by auto
   have "?a \<bullet> x = 0" if "x\<in>span B" for x
@@ -861,23 +735,19 @@
         by blast
       have fth: "finite (B - {x})"
         using fB by simp
-      have "?a \<bullet> x = 0"
+      have "(\<Sum>b\<in>B - {x}. a \<bullet> b * (b \<bullet> x) / (b \<bullet> b)) = 0" if "x \<noteq> 0"
+        by (smt (verit) B' B(5) DiffD2 divide_eq_0_iff inner_real_def inner_zero_right insertCI orthogonal_def pairwise_insert sum.neutral)
+      then have "?a \<bullet> x = 0"
         apply (subst B')
         using fB fth
         unfolding sum_clauses(2)[OF fth]
-        apply simp unfolding inner_simps
-        apply (clarsimp simp add: inner_add inner_sum_left)
-        apply (rule sum.neutral, rule ballI)
-        apply (simp only: inner_commute)
-        apply (auto simp add: x field_simps
-          intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])
-        done
+        by (auto simp add: inner_add_left inner_diff_left inner_sum_left)
     }
     then show "?a \<bullet> x = 0" if "x \<in> B" for x
       using that by blast
     qed
-  with a0 show ?thesis
-    unfolding sSB by (auto intro: exI[where x="?a"])
+  with a0 sSB show ?thesis
+    by blast
 qed
 
 lemma span_not_univ_subset_hyperplane:
@@ -890,19 +760,7 @@
   fixes S :: "'a::euclidean_space set"
   assumes d: "dim S < DIM('a)"
   shows "\<exists>a::'a. a \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
-proof -
-  {
-    assume "span S = UNIV"
-    then have "dim (span S) = dim (UNIV :: ('a) set)"
-      by simp
-    then have "dim S = DIM('a)"
-      by (metis Euclidean_Space.dim_UNIV dim_span)
-    with d have False by arith
-  }
-  then have th: "span S \<noteq> UNIV"
-    by blast
-  from span_not_univ_subset_hyperplane[OF th] show ?thesis .
-qed
+  using d dim_eq_full nless_le span_not_univ_subset_hyperplane by blast
 
 lemma linear_eq_stdbasis:
   fixes f :: "'a::euclidean_space \<Rightarrow> _"
@@ -910,8 +768,7 @@
     and lg: "linear g"
     and fg: "\<And>b. b \<in> Basis \<Longrightarrow> f b = g b"
   shows "f = g"
-  using linear_eq_on_span[OF lf lg, of Basis] fg
-  by auto
+  using linear_eq_on_span[OF lf lg, of Basis] fg by auto
 
 
 text \<open>Similar results for bilinear functions.\<close>
@@ -932,15 +789,9 @@
         span_add Ball_def
       intro: bilinear_ladd[OF bf])
   have sfg: "\<And>x. x \<in> B \<Longrightarrow> subspace {a. f x a = g x a}"
-    apply (auto simp add: subspace_def)
-    using bf bg unfolding bilinear_def linear_iff
-      apply (auto simp add: span_zero bilinear_rzero[OF bf] bilinear_rzero[OF bg]
-        span_add Ball_def
-      intro: bilinear_ladd[OF bf])
-    done
+    by (auto simp: subspace_def bf bg bilinear_rzero bilinear_radd bilinear_rmul)
   have "\<forall>y\<in> span C. f x y = g x y" if "x \<in> span B" for x
-    apply (rule span_induct [OF that sp])
-    using fg sfg span_induct by blast
+    using span_induct [OF that sp] fg sfg span_induct by blast
   then show ?thesis
     using SB TC assms by auto
 qed
@@ -972,8 +823,7 @@
   fixes x :: "'a::euclidean_space"
   shows "finite {\<bar>x \<bullet> i\<bar> |i. i \<in> Basis}"
     and "{\<bar>x \<bullet> i\<bar> |i. i \<in> Basis} \<noteq> {}"
-  unfolding infnorm_set_image
-  by auto
+  unfolding infnorm_set_image by auto
 
 lemma infnorm_pos_le:
   fixes x :: "'a::euclidean_space"
@@ -1010,16 +860,7 @@
   by (metis infnorm_neg minus_diff_eq)
 
 lemma absdiff_infnorm: "\<bar>infnorm x - infnorm y\<bar> \<le> infnorm (x - y)"
-proof -
-  have *: "\<And>(nx::real) n ny. nx \<le> n + ny \<Longrightarrow> ny \<le> n + nx \<Longrightarrow> \<bar>nx - ny\<bar> \<le> n"
-    by arith
-  show ?thesis
-  proof (rule *)
-    from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
-    show "infnorm x \<le> infnorm (x - y) + infnorm y" "infnorm y \<le> infnorm (x - y) + infnorm x"
-      by (simp_all add: field_simps infnorm_neg)
-  qed
-qed
+  by (smt (verit, del_insts) diff_add_cancel infnorm_sub infnorm_triangle)
 
 lemma real_abs_infnorm: "\<bar>infnorm x\<bar> = infnorm x"
   using infnorm_pos_le[of x] by arith
@@ -1065,9 +906,9 @@
     by (simp add: zero_le_mult_iff infnorm_pos_le)
   have "x \<bullet> x \<le> (\<Sum>b\<in>Basis. x \<bullet> b * (x \<bullet> b))"
     by (metis euclidean_inner order_refl)
-  also have "... \<le> DIM('a) * \<bar>infnorm x\<bar>\<^sup>2"
+  also have "\<dots> \<le> DIM('a) * \<bar>infnorm x\<bar>\<^sup>2"
     by (rule sum_bounded_above) (metis Basis_le_infnorm abs_le_square_iff power2_eq_square real_abs_infnorm)
-  also have "... \<le> (sqrt DIM('a) * infnorm x)\<^sup>2"
+  also have "\<dots> \<le> (sqrt DIM('a) * infnorm x)\<^sup>2"
     by (simp add: power_mult_distrib)
   finally show "x \<bullet> x \<le> (sqrt DIM('a) * infnorm x)\<^sup>2" .
 qed
@@ -1091,11 +932,11 @@
   then show ?thesis
     by auto
 next
-  case False
+  case False 
   from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"]
   have "?rhs \<longleftrightarrow>
       (norm y * (norm y * norm x * norm x - norm x * (x \<bullet> y)) -
-        norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) =  0)"
+        norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) = 0)"
     using False unfolding inner_simps
     by (auto simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)
   also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)"
@@ -1108,20 +949,9 @@
 lemma norm_cauchy_schwarz_abs_eq:
   "\<bar>x \<bullet> y\<bar> = norm x * norm y \<longleftrightarrow>
     norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm x *\<^sub>R y = - norm y *\<^sub>R x"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof -
-  have th: "\<And>(x::real) a. a \<ge> 0 \<Longrightarrow> \<bar>x\<bar> = a \<longleftrightarrow> x = a \<or> x = - a"
-    by arith
-  have "?rhs \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm (- x) *\<^sub>R y = norm y *\<^sub>R (- x)"
-    by simp
-  also have "\<dots> \<longleftrightarrow> (x \<bullet> y = norm x * norm y \<or> (- x) \<bullet> y = norm x * norm y)"
-    unfolding norm_cauchy_schwarz_eq[symmetric]
-    unfolding norm_minus_cancel norm_scaleR ..
-  also have "\<dots> \<longleftrightarrow> ?lhs"
-    unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] inner_simps
-    by auto
-  finally show ?thesis ..
-qed
+  using norm_cauchy_schwarz_eq [symmetric, of x y]
+  using norm_cauchy_schwarz_eq [symmetric, of "-x" y] Cauchy_Schwarz_ineq2 [of x y]
+  by auto
 
 lemma norm_triangle_eq:
   fixes x y :: "'a::real_inner"
@@ -1137,9 +967,7 @@
   have "norm (x + y) = norm x + norm y \<longleftrightarrow> (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2"
     by simp
   also have "\<dots> \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x"
-    unfolding norm_cauchy_schwarz_eq[symmetric]
-    unfolding power2_norm_eq_inner inner_simps
-    by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)
+    by (smt (verit, best) dot_norm inner_real_def inner_simps norm_cauchy_schwarz_eq power2_eq_square)
   finally show ?thesis .
 qed
 
@@ -1147,11 +975,7 @@
   fixes x y z :: "'a::real_inner"
   shows "dist x z = dist x y + dist y z \<longleftrightarrow>
     norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)"
-proof -
-  have *: "x - y + (y - z) = x - z" by auto
-  show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *]
-    by (auto simp:norm_minus_commute)
-qed
+  by (metis (no_types, lifting) add_diff_eq diff_add_cancel dist_norm norm_triangle_eq)
 
 subsection \<open>Collinearity\<close>
 
@@ -1163,7 +987,7 @@
 proof
   assume ?lhs
   then show ?rhs
-    unfolding collinear_def by (metis Groups.add_ac(2) diff_add_cancel)
+    unfolding collinear_def by (metis add.commute diff_add_cancel)
 next
   assume ?rhs
   then obtain u v where *: "\<And>x. x \<in> S \<Longrightarrow> \<exists>c. x = u + c *\<^sub>R v"
@@ -1192,8 +1016,7 @@
       using \<open>v \<noteq> 0\<close> by blast
   qed
   then show ?thesis
-    apply (clarsimp simp: collinear_def)
-    by (metis scaleR_zero_right vector_fraction_eq_iff)
+    by (metis collinear_def)
 qed
 
 lemma collinear_subset: "\<lbrakk>collinear T; S \<subseteq> T\<rbrakk> \<Longrightarrow> collinear S"
@@ -1206,9 +1029,7 @@
   by (simp add: collinear_def)
 
 lemma collinear_2 [iff]: "collinear {x, y}"
-  apply (simp add: collinear_def)
-  apply (rule exI[where x="x - y"])
-  by (metis minus_diff_eq scaleR_left.minus scaleR_one)
+  by (simp add: collinear_def) (metis minus_diff_eq scaleR_left.minus scaleR_one)
 
 lemma collinear_lemma: "collinear {0, x, y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *\<^sub>R x)"
   (is "?lhs \<longleftrightarrow> ?rhs")
@@ -1237,15 +1058,8 @@
     then obtain c where c: "y = c *\<^sub>R x"
       using False by blast
     show ?lhs
-      unfolding collinear_def c
-      apply (rule exI[where x=x])
-      apply auto
-          apply (rule exI[where x="- 1"], simp)
-         apply (rule exI[where x= "-c"], simp)
-        apply (rule exI[where x=1], simp)
-       apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib)
-      apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib)
-      done
+      apply (simp add: collinear_def c)
+      by (metis (mono_tags, lifting) scaleR_left.minus scaleR_left_diff_distrib scaleR_one)
   qed
 qed
 
@@ -1291,11 +1105,8 @@
   fixes x y :: "'a::real_inner"
   assumes "norm (x + y) = norm x + norm y"
   shows "collinear{0,x,y}"
-proof (cases "x = 0 \<or> y = 0")
-  case False
-  with assms show ?thesis
-    by (meson norm_cauchy_schwarz_abs_eq norm_cauchy_schwarz_equal norm_triangle_eq)
-qed (use collinear_lemma in blast)
+  using assms norm_cauchy_schwarz_abs_eq norm_cauchy_schwarz_equal norm_triangle_eq 
+  by blast
 
 
 subsection\<open>Properties of special hyperplanes\<close>
@@ -1315,7 +1126,7 @@
   proof -
     have "x = (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b)"
       by (simp add: euclidean_representation)
-    also have "... = (\<Sum>b \<in> Basis - {k}. (x \<bullet> b) *\<^sub>R b)"
+    also have "\<dots> = (\<Sum>b \<in> Basis - {k}. (x \<bullet> b) *\<^sub>R b)"
       by (auto simp: sum.remove [of _ k] inner_commute assms that)
     finally have "x = (\<Sum>b\<in>Basis - {k}. (x \<bullet> b) *\<^sub>R b)" .
     then show ?thesis
@@ -1331,11 +1142,7 @@
 lemma dim_special_hyperplane:
   fixes k :: "'n::euclidean_space"
   shows "k \<in> Basis \<Longrightarrow> dim {x. k \<bullet> x = 0} = DIM('n) - 1"
-apply (simp add: special_hyperplane_span)
-apply (rule dim_unique [OF subset_refl])
-apply (auto simp: independent_substdbasis)
-apply (metis member_remove remove_def span_base)
-done
+  by (metis Diff_subset card_Diff_singleton indep_card_eq_dim_span independent_substdbasis special_hyperplane_span)
 
 proposition dim_hyperplane:
   fixes a :: "'a::euclidean_space"
@@ -1358,20 +1165,17 @@
     using \<open>independent B\<close> independent_bound by blast
   have "UNIV \<subseteq> span (insert a B)"
   proof fix y::'a
-    obtain r z where z: "y = r *\<^sub>R a + z" "a \<bullet> z = 0"
-      apply (rule_tac r="(a \<bullet> y) / (a \<bullet> a)" and z = "y - ((a \<bullet> y) / (a \<bullet> a)) *\<^sub>R a" in that)
-      using assms
-      by (auto simp: algebra_simps)
-    show "y \<in> span (insert a B)"
-      by (metis (mono_tags, lifting) z Bsub span_eq_iff
-         add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB)
+    obtain r z where "y = r *\<^sub>R a + z" "a \<bullet> z = 0"
+      by (metis add.commute diff_add_cancel vector_sub_project_orthogonal)
+    then show "y \<in> span (insert a B)"
+      by (metis (mono_tags, lifting) Bsub add_diff_cancel_left'
+          mem_Collect_eq span0 span_breakdown_eq span_eq subspB)
   qed
-  then have dima: "DIM('a) = dim(insert a B)"
+  then have "DIM('a) = dim(insert a B)"
     by (metis independent_Basis span_Basis dim_eq_card top.extremum_uniqueI)
   then show ?thesis
-    by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \<open>a \<notin> span B\<close> ind card0
-        card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE
-        subspB)
+    by (metis One_nat_def \<open>a \<notin> span B\<close> \<open>finite B\<close> card0 card_insert_disjoint 
+        diff_Suc_Suc diff_zero dim_eq_card_independent ind span_base)
 qed
 
 lemma lowdim_eq_hyperplane:
@@ -1379,14 +1183,10 @@
   assumes "dim S = DIM('a) - 1"
   obtains a where "a \<noteq> 0" and "span S = {x. a \<bullet> x = 0}"
 proof -
-  have dimS: "dim S < DIM('a)"
-    by (simp add: assms)
-  then obtain b where b: "b \<noteq> 0" "span S \<subseteq> {a. b \<bullet> a = 0}"
-    using lowdim_subset_hyperplane [of S] by fastforce
-  show ?thesis
-    apply (rule that[OF b(1)])
-    apply (rule subspace_dim_equal)
-    by (auto simp: assms b dim_hyperplane subspace_hyperplane)
+  obtain b where b: "b \<noteq> 0" "span S \<subseteq> {a. b \<bullet> a = 0}"
+    by (metis DIM_positive assms diff_less zero_less_one lowdim_subset_hyperplane)
+  then show ?thesis
+    by (metis assms dim_hyperplane dim_span dim_subset subspace_dim_equal subspace_hyperplane subspace_span that)
 qed
 
 lemma dim_eq_hyperplane:
@@ -1409,10 +1209,9 @@
       using a by (force simp: span_explicit)
     then have "a \<bullet> a = a \<bullet> (\<Sum>v\<in>T. U v *\<^sub>R v)"
       by simp
-    also have "... = 0"
+    also have "\<dots> = 0"
       apply (simp add: inner_sum_right)
-      apply (rule comm_monoid_add_class.sum.neutral)
-      by (metis "0" DiffE \<open>T \<subseteq> S - {a}\<close> mult_not_zero singletonI subsetCE \<open>a \<in> S\<close>)
+      by (smt (verit) "0" DiffE \<open>T \<subseteq> S - {a}\<close> in_mono insertCI mult_not_zero sum.neutral that(1))
     finally show ?thesis
       using \<open>0 \<notin> S\<close> \<open>a \<in> S\<close> by auto
   qed
@@ -1424,14 +1223,8 @@
   fixes S :: "'a::euclidean_space set"
   assumes "pairwise orthogonal S"
     shows "finite S"
-proof -
-  have "independent (S - {0})"
-    apply (rule pairwise_orthogonal_independent)
-     apply (metis Diff_iff assms pairwise_def)
-    by blast
-  then show ?thesis
-    by (meson independent_imp_finite infinite_remove)
-qed
+  by (metis Set.set_insert assms finite_insert independent_bound pairwise_insert 
+            pairwise_orthogonal_independent)
 
 lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}"
   by (simp add: subspace_def orthogonal_clauses)
@@ -1457,7 +1250,7 @@
   proof -
     have "a \<bullet> x = (\<Sum>y\<in>S. if y = x then y \<bullet> a else 0)"
       by (simp add: \<open>finite S\<close> inner_commute that)
-    also have "... =  (\<Sum>b\<in>S. b \<bullet> a * (b \<bullet> x) / (b \<bullet> b))"
+    also have "\<dots> =  (\<Sum>b\<in>S. b \<bullet> a * (b \<bullet> x) / (b \<bullet> b))"
       apply (rule sum.cong [OF refl], simp)
       by (meson S orthogonal_def pairwise_def that)
    finally show ?thesis
@@ -1487,20 +1280,15 @@
       (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute
         pairwise_orthogonal_insert span_clauses)
   have orthS: "\<And>x. x \<in> S \<Longrightarrow> a' \<bullet> x = 0"
-    apply (simp add: a'_def)
-    using Gram_Schmidt_step [OF \<open>pairwise orthogonal S\<close>]
-    apply (force simp: orthogonal_def inner_commute span_superset [THEN subsetD])
-    done
+    using Gram_Schmidt_step a'_def insert.prems orthogonal_commute orthogonal_def span_base by blast
   have "span (S \<union> insert a' U) = span (insert a' (S \<union> T))"
     using spanU by simp
-  also have "... = span (insert a (S \<union> T))"
-    apply (rule eq_span_insert_eq)
-    apply (simp add: a'_def span_neg span_sum span_base span_mul)
-    done
-  also have "... = span (S \<union> insert a T)"
+  also have "\<dots> = span (insert a (S \<union> T))"
+    by (simp add: a'_def span_neg span_sum span_base span_mul eq_span_insert_eq)
+  also have "\<dots> = span (S \<union> insert a T)"
     by simp
   finally show ?case
-    by (rule_tac x="insert a' U" in exI) (use orthU in auto)
+    using orthU by blast
 qed
 
 
@@ -1524,14 +1312,12 @@
   obtains U where "U \<inter> (insert 0 S) = {}" "pairwise orthogonal (S \<union> U)"
                   "span (S \<union> U) = span (S \<union> T)"
 proof -
-  obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)"
+  obtain U where U: "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)"
     using orthogonal_extension assms by blast
-  then show ?thesis
-    apply (rule_tac U = "U - (insert 0 S)" in that)
-      apply blast
-     apply (force simp: pairwise_def)
-    apply (metis Un_Diff_cancel Un_insert_left span_redundant span_zero)
-    done
+  moreover have "pairwise orthogonal (S \<union> (U - insert 0 S))"
+    by (smt (verit, best) Un_Diff_Int Un_iff U pairwise_def)
+  ultimately show ?thesis
+    by (metis Diff_disjoint Un_Diff_cancel Un_insert_left inf_commute span_insert_0 that)
 qed
 
 subsection\<open>Decomposing a vector into parts in orthogonal subspaces\<close>
@@ -1542,27 +1328,14 @@
   fixes S :: "'a :: euclidean_space set"
   assumes "subspace S"
   obtains B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S"
-proof -
-  obtain B where "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S"
-    using basis_exists by blast
-  with orthogonal_extension [of "{}" B]
-  show ?thesis
-    by (metis Un_empty_left assms pairwise_empty span_superset span_subspace that)
-qed
+  by (metis assms basis_orthogonal basis_subspace_exists span_eq)
 
 lemma orthogonal_basis_subspace:
   fixes S :: "'a :: euclidean_space set"
   assumes "subspace S"
   obtains B where "0 \<notin> B" "B \<subseteq> S" "pairwise orthogonal B" "independent B"
                   "card B = dim S" "span B = S"
-proof -
-  obtain B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S"
-    using assms orthogonal_spanningset_subspace by blast
-  then show ?thesis
-    apply (rule_tac B = "B - {0}" in that)
-    apply (auto simp: indep_card_eq_dim_span pairwise_subset pairwise_orthogonal_independent elim: pairwise_subset)
-    done
-qed
+  by (metis assms dependent_zero orthogonal_basis_exists span_eq span_eq_iff)
 
 proposition orthonormal_basis_subspace:
   fixes S :: "'a :: euclidean_space set"
@@ -1609,7 +1382,7 @@
   obtain B where "B \<subseteq> span S" and orthB: "pairwise orthogonal B"
              and "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
              and "independent B" "card B = dim S" "span B = span S"
-    by (rule orthonormal_basis_subspace [of "span S", OF subspace_span]) (auto)
+    by (metis dim_span orthonormal_basis_subspace subspace_span)
   with assms obtain u where spanBT: "span B \<subseteq> span T" and "u \<notin> span B" "u \<in> span T"
     by auto
   obtain C where orthBC: "pairwise orthogonal (B \<union> C)" and spanBC: "span (B \<union> C) = span (B \<union> {u})"
@@ -1630,9 +1403,8 @@
     then obtain x where "x \<in> C" "x \<noteq> 0" "x \<notin> B"
       by blast
     then have "x \<in> span T"
-      by (metis (no_types, lifting) Un_insert_right Un_upper2 \<open>u \<in> span T\<close> spanBT spanBC
-          \<open>u \<in> span T\<close> insert_subset span_superset span_mono
-          span_span subsetCE subset_trans sup_bot.comm_neutral)
+      by (smt (verit, ccfv_SIG) Set.set_insert  \<open>u \<in> span T\<close> empty_subsetI insert_subset 
+          le_sup_iff spanBC spanBT span_mono span_span span_superset subset_trans)
     moreover have "orthogonal x y" if "y \<in> span B" for y
       using that
     proof (rule span_induct)
@@ -1652,8 +1424,7 @@
   obtains x where "x \<noteq> 0" "\<And>y. y \<in> span S \<Longrightarrow> orthogonal x y"
 proof -
   have "span S \<subset> UNIV"
-  by (metis (mono_tags) UNIV_I assms inner_eq_zero_iff less_le lowdim_subset_hyperplane
-      mem_Collect_eq top.extremum_strict top.not_eq_extremum)
+    by (metis assms dim_eq_full order_less_imp_not_less top.not_eq_extremum)
   with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis
     by (auto)
 qed
@@ -1683,11 +1454,8 @@
   have orth: "orthogonal (x - ?a) w" if "w \<in> span S" for w
     by (simp add: Gram_Schmidt_step \<open>pairwise orthogonal T\<close> \<open>span T = span S\<close>
         orthogonal_commute that)
-  show ?thesis
-    apply (rule_tac y = "?a" and z = "x - ?a" in that)
-      apply (meson \<open>T \<subseteq> span S\<close> span_scale span_sum subsetCE)
-     apply (fact orth, simp)
-    done
+  with that[of ?a "x-?a"] \<open>T \<subseteq> span S\<close> show ?thesis
+    by (simp add: span_mul span_sum subsetD)
 qed
 
 lemma orthogonal_subspace_decomp_unique:
@@ -1702,15 +1470,15 @@
   moreover have "\<And>a b. \<lbrakk>a \<in> span S; b \<in> span T\<rbrakk> \<Longrightarrow> orthogonal a b"
     by (meson orth orthogonal_commute orthogonal_to_span)
   ultimately have "0 = x' - x"
-    by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self)
+    using assms
+    by (metis add.commute add_diff_cancel_right' diff_right_commute orthogonal_self span_diff)
   with assms show ?thesis by auto
 qed
 
 lemma vector_in_orthogonal_spanningset:
   fixes a :: "'a::euclidean_space"
   obtains S where "a \<in> S" "pairwise orthogonal S" "span S = UNIV"
-  by (metis UNIV_I Un_iff empty_iff insert_subset orthogonal_extension pairwise_def
-      pairwise_orthogonal_insert span_UNIV subsetI subset_antisym)
+  by (metis UnI1 Un_UNIV_right insertI1 orthogonal_extension pairwise_singleton span_UNIV)
 
 lemma vector_in_orthogonal_basis:
   fixes a :: "'a::euclidean_space"
@@ -1755,17 +1523,15 @@
       using \<open>pairwise orthogonal S\<close> by (auto simp: pairwise_def orthogonal_def)
     show "\<And>x. x \<in> (\<lambda>x. x /\<^sub>R norm x) ` S \<Longrightarrow> norm x = 1"
       using \<open>0 \<notin> S\<close> by (auto simp: field_split_simps)
-    then show "independent ?S"
+    then show ind: "independent ?S"
       by (metis \<open>pairwise orthogonal ((\<lambda>x. x /\<^sub>R norm x) ` S)\<close> norm_zero pairwise_orthogonal_independent zero_neq_one)
     have "inj_on (\<lambda>x. x /\<^sub>R norm x) S"
       unfolding inj_on_def
       by (metis (full_types) S(1) \<open>0 \<notin> S\<close> inverse_nonzero_iff_nonzero norm_eq_zero orthogonal_scaleR orthogonal_self pairwise_def)
     then show "card ?S = DIM('a)"
       by (simp add: card_image S)
-    show "span ?S = UNIV"
-      by (metis (no_types) \<open>0 \<notin> S\<close> \<open>finite S\<close> \<open>span S = UNIV\<close>
-          field_class.field_inverse_zero inverse_inverse_eq less_irrefl span_image_scale
-          zero_less_norm_iff)
+    then show "span ?S = UNIV"
+      by (metis ind dim_eq_card dim_eq_full)
   qed
 qed
 
@@ -1786,12 +1552,10 @@
     by (auto simp add: span_Un image_def)
   also have "dim \<dots> = dim {x + y |x y. x \<in> span A \<and> y \<in> span B}"
     by (auto intro!: arg_cong [where f=dim])
-  also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B} + dim(span A \<inter> span B)"
-    by (auto simp: dest: 0)
-  also have "... = dim (span A) + dim (span B)"
-    by (rule dim_sums_Int) (auto)
-  also have "... = dim A + dim B"
-    by (simp)
+  also have "\<dots> = dim {x + y |x y. x \<in> span A \<and> y \<in> span B} + dim(span A \<inter> span B)"
+    by (auto dest: 0)
+  also have "\<dots> = dim A + dim B"
+    using dim_sums_Int by fastforce
   finally show ?thesis .
 qed
 
@@ -1810,17 +1574,16 @@
     proof -
       obtain y z where "x = y + z" "y \<in> span A" and orth: "\<And>w. w \<in> span A \<Longrightarrow> orthogonal z w"
         using orthogonal_subspace_decomp_exists [of A x] that by auto
+      moreover
       have "y \<in> span B"
         using \<open>y \<in> span A\<close> assms(3) span_mono by blast
-      then have "z \<in> {a \<in> B. \<forall>x. x \<in> A \<longrightarrow> orthogonal x a}"
-        apply simp
-        using \<open>x = y + z\<close> assms(1) assms(2) orth orthogonal_commute span_add_eq
-          span_eq_iff that by blast
+      ultimately have "z \<in> B \<and> (\<forall>x. x \<in> A \<longrightarrow> orthogonal x z)"
+        using assms by (metis orthogonal_commute span_add_eq span_eq_iff that)
       then have z: "z \<in> span {y \<in> B. \<forall>x\<in>A. orthogonal x y}"
-        by (meson span_superset subset_iff)
+        by (simp add: span_base)
       then show ?thesis
-        apply (auto simp: span_Un image_def  \<open>x = y + z\<close> \<open>y \<in> span A\<close>)
-        using \<open>y \<in> span A\<close> add.commute by blast
+        by (smt (verit, best) \<open>x = y + z\<close> \<open>y \<in> span A\<close> le_sup_iff span_add_eq span_subspace_induct 
+            span_superset subset_iff subspace_span)
     qed
     show "span B \<subseteq> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)"
       by (rule span_minimal) (auto intro: * span_minimal)
@@ -1845,12 +1608,8 @@
 qed
 
 lemma linear_continuous_at:
-  assumes "bounded_linear f"
-  shows "continuous (at a) f"
-  unfolding continuous_at using assms
-  apply (rule bounded_linear.tendsto)
-  apply (rule tendsto_ident_at)
-  done
+  "bounded_linear f \<Longrightarrow>continuous (at a) f"
+  by (simp add: bounded_linear.isUCont isUCont_isCont)
 
 lemma linear_continuous_within:
   "bounded_linear f \<Longrightarrow> continuous (at x within s) f"
@@ -1869,21 +1628,7 @@
     using linear_bounded_pos [OF \<open>linear h\<close>] by blast
   show ?thesis
     unfolding tendsto_iff
-  proof (intro allI impI)
-    show "\<forall>\<^sub>F x in F. dist (h (f x)) (h l) < e" if "e > 0" for e
-    proof -
-      have "\<forall>\<^sub>F x in F. dist (f x) l < e/B"
-        by (simp add: \<open>0 < B\<close> assms(1) tendstoD that)
-      then show ?thesis
-        unfolding dist_norm
-      proof (rule eventually_mono)
-        show "norm (h (f x) - h l) < e" if "norm (f x - l) < e / B" for x
-          using that B
-          apply (simp add: field_split_simps)
-          by (metis \<open>linear h\<close> le_less_trans linear_diff)
-      qed
-    qed
-  qed
+      by (simp add: assms bounded_linear.tendsto linear_linear tendstoD)
 qed
 
 lemma linear_continuous_compose:
--- a/src/HOL/Analysis/Measure_Space.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Measure_Space.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -26,9 +26,9 @@
   proof (rule SUP_eqI)
     fix y :: ennreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
     from this[of "Suc i"] show "f i \<le> y" by auto
-  qed (insert assms, simp)
+  qed (use assms in simp)
   ultimately show ?thesis using assms
-    by (subst suminf_eq_SUP) (auto simp: indicator_def)
+    by (simp add: suminf_eq_SUP)
 qed
 
 lemma suminf_indicator:
@@ -71,11 +71,8 @@
         by (induct n)  (auto simp add: binaryset_def f)
     qed
   moreover
-  have "... \<longlonglongrightarrow> f A + f B" by (rule tendsto_const)
-  ultimately
-  have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
-    by metis
-  hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
+  have "\<dots> \<longlonglongrightarrow> f A + f B" by (rule tendsto_const)
+  ultimately have "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
     by simp
   thus ?thesis by (rule LIMSEQ_offset [where k=2])
 qed
@@ -83,7 +80,7 @@
 lemma binaryset_sums:
   assumes f: "f {} = 0"
   shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
-    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
+  using LIMSEQ_binaryset f sums_def by blast
 
 lemma suminf_binaryset_eq:
   fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
@@ -148,7 +145,7 @@
   by (auto simp add: increasing_def)
 
 lemma countably_additiveI[case_names countably]:
-  "(\<And>A. range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> M \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i))
+  "(\<And>A. \<lbrakk>range A \<subseteq> M; disjoint_family A; (\<Union>i. A i) \<in> M\<rbrakk> \<Longrightarrow> (\<Sum>i. f(A i)) = f(\<Union>i. A i))
   \<Longrightarrow> countably_additive M f"
   by (simp add: countably_additive_def)
 
@@ -198,9 +195,9 @@
   assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
   then have "y - x \<in> M" by auto
   then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono zero_le)
-  also have "... = f (x \<union> (y-x))" using addf
-    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
-  also have "... = f y"
+  also have "\<dots> = f (x \<union> (y-x))"
+    by (metis addf Diff_disjoint \<open>y - x \<in> M\<close> additiveD xy(1))
+  also have "\<dots> = f y"
     by (metis Un_Diff_cancel Un_absorb1 xy(3))
   finally show "f x \<le> f y" by simp
 qed
@@ -222,9 +219,12 @@
   also have "\<dots> = f (A x) + f ((\<Union>i\<in>F. A i) - A x)"
     using f(2) by (rule additiveD) (insert in_M, auto)
   also have "\<dots> \<le> f (A x) + f (\<Union>i\<in>F. A i)"
-    using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono)
-  also have "\<dots> \<le> f (A x) + (\<Sum>i\<in>F. f (A i))" using insert by (auto intro: add_left_mono)
-  finally show "f (\<Union>i\<in>(insert x F). A i) \<le> (\<Sum>i\<in>(insert x F). f (A i))" using insert by simp
+    using additive_increasing[OF f] in_M subs 
+    by (simp add: increasingD)
+  also have "\<dots> \<le> f (A x) + (\<Sum>i\<in>F. f (A i))" 
+    using insert by (auto intro: add_left_mono)
+  finally show "f (\<Union>i\<in>(insert x F). A i) \<le> (\<Sum>i\<in>(insert x F). f (A i))"
+    by (simp add: insert)
 qed
 
 lemma (in ring_of_sets) countably_additive_additive:
@@ -239,10 +239,8 @@
   hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
          (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
          f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
-    using ca
-    by (simp add: countably_additive_def)
-  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
-         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
+    using ca by (simp add: countably_additive_def)
+  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow> f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
     by (simp add: range_binaryset_eq UN_binaryset_eq)
   thus "f (x \<union> y) = f x + f y" using posf x y
     by (auto simp add: Un suminf_binaryset_eq positive_def)
@@ -259,8 +257,8 @@
   fix N
   note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
   have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
-    using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
-  also have "... \<le> f \<Omega>" using space_closed A
+    using A by (intro additive_sum [OF f ad]) (auto simp: disj_N)
+  also have "\<dots> \<le> f \<Omega>" using space_closed A
     by (intro increasingD[OF inc] finite_UN) auto
   finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp
 qed (insert f A, auto simp: positive_def)
@@ -272,16 +270,10 @@
 proof (rule countably_additiveI)
   fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F"
 
-  have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
-  from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
-
-  have inj_f: "inj_on f {i. F i \<noteq> {}}"
-  proof (rule inj_onI, simp)
-    fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
-    then have "f i \<in> F i" "f j \<in> F j" using f by force+
-    with disj * show "i = j" by (auto simp: disjoint_family_on_def)
-  qed
-  have "finite (\<Union>i. F i)"
+  have "\<forall>i. F i \<noteq> {} \<longrightarrow> (\<exists>x. x \<in> F i)" by auto
+  then obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by metis
+
+  have finU: "finite (\<Union>i. F i)"
     by (metis F(2) assms(1) infinite_super sets_into_space)
 
   have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
@@ -290,8 +282,11 @@
   proof (rule finite_imageD)
     from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
     then show "finite (f`{i. F i \<noteq> {}})"
-      by (rule finite_subset) fact
-  qed fact
+      by (simp add: finU finite_subset)
+    show inj_f: "inj_on f {i. F i \<noteq> {}}" 
+      using f disj
+      by (simp add: inj_on_def disjoint_family_on_def disjoint_iff) metis
+  qed 
   ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
     by (rule finite_subset)
 
@@ -323,8 +318,7 @@
   have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
     by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
   moreover have "(\<lambda>n. (\<Sum>i<n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
-    using f(1)[unfolded positive_def] dA
-    by (auto intro!: summable_LIMSEQ)
+    by (simp add: summable_LIMSEQ)
   from LIMSEQ_Suc[OF this]
   have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
     unfolding lessThan_Suc_atMost .
@@ -332,23 +326,21 @@
     using disjointed_additive[OF f A(1,2)] .
   ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" by simp
 next
-  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
+  assume cont[rule_format]: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
   fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "disjoint_family A" "(\<Union>i. A i) \<in> M"
   have *: "(\<Union>n. (\<Union>i<n. A i)) = (\<Union>i. A i)" by auto
-  have "(\<lambda>n. f (\<Union>i<n. A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
-  proof (unfold *[symmetric], intro cont[rule_format])
-    show "range (\<lambda>i. \<Union>i<i. A i) \<subseteq> M" "(\<Union>i. \<Union>i<i. A i) \<in> M"
-      using A * by auto
-  qed (force intro!: incseq_SucI)
+  have "range (\<lambda>i. \<Union>i<i. A i) \<subseteq> M" "(\<Union>i. \<Union>i<i. A i) \<in> M"
+    using A * by auto
+  then have "(\<lambda>n. f (\<Union>i<n. A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
+    unfolding *[symmetric] by (force intro!: cont incseq_SucI)+
   moreover have "\<And>n. f (\<Union>i<n. A i) = (\<Sum>i<n. f (A i))"
     using A
-    by (intro additive_sum[OF f, of _ A, symmetric])
-       (auto intro: disjoint_family_on_mono[where B=UNIV])
+    by (intro additive_sum[OF f, symmetric]) (auto intro: disjoint_family_on_mono)
   ultimately
   have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)"
     unfolding sums_def by simp
-  from sums_unique[OF this]
-  show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp
+  then show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)"
+    by (metis sums_unique)
 qed
 
 lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
@@ -357,13 +349,15 @@
   shows "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))
      \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0)"
 proof safe
-  assume cont: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
-  with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
+  assume cont[rule_format]: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))"
+  fix A :: "nat \<Rightarrow> 'a set" 
+  assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
+  with cont[of A] show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
     using \<open>positive M f\<close>[unfolded positive_def] by auto
 next
-  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>"
+  assume cont[rule_format]: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
+  fix A :: "nat \<Rightarrow> 'a set" 
+  assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>"
 
   have f_mono: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
     using additive_increasing[OF f] unfolding increasing_def by simp
@@ -378,23 +372,19 @@
     using A by (auto intro!: f_mono)
   then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>"
     using A by (auto simp: top_unique)
-  { fix i
-    have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono)
-    then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>"
-      using A by (auto simp: top_unique) }
-  note f_fin = this
+  have f_fin: "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>" for i
+    using A by (metis Diff Diff_subset f_mono infinity_ennreal_def range_subsetD top_unique)
   have "(\<lambda>i. f (A i - (\<Inter>i. A i))) \<longlonglongrightarrow> 0"
-  proof (intro cont[rule_format, OF _ decseq _ f_fin])
+  proof (intro cont[ OF _ decseq _ f_fin])
     show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
       using A by auto
   qed
-  from INF_Lim[OF decseq_f this]
-  have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
+  with INF_Lim decseq_f have "(INF n. f (A n - (\<Inter>i. A i))) = 0" by metis
   moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
     by auto
   ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
     using A(4) f_fin f_Int_fin
-    by (subst INF_ennreal_add_const) (auto simp: decseq_f)
+    using INF_ennreal_add_const by presburger
   moreover {
     fix n
     have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
@@ -424,7 +414,8 @@
     also have "((\<Union>i. A i) - A i) \<union> A i = (\<Union>i. A i)"
       by auto
     finally have "f ((\<Union>i. A i) - A i) = f (\<Union>i. A i) - f (A i)"
-      using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) }
+      using assms f by fastforce
+  }
   moreover have "\<forall>\<^sub>F i in sequentially. f (A i) \<le> f (\<Union>i. A i)"
     using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\<Union>i. A i"] A
     by (auto intro!: always_eventually simp: subset_eq)
@@ -502,17 +493,11 @@
   by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le)
 
 lemma emeasure_Diff:
-  assumes finite: "emeasure M B \<noteq> \<infinity>"
-  and [measurable]: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
+  assumes "emeasure M B \<noteq> \<infinity>"
+  and "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
   shows "emeasure M (A - B) = emeasure M A - emeasure M B"
-proof -
-  have "(A - B) \<union> B = A" using \<open>B \<subseteq> A\<close> by auto
-  then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
-  also have "\<dots> = emeasure M (A - B) + emeasure M B"
-    by (subst plus_emeasure[symmetric]) auto
-  finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
-    using finite by simp
-qed
+  by (smt (verit, best) add_diff_self_ennreal assms emeasure_Un emeasure_mono 
+      ennreal_add_left_cancel le_iff_sup)
 
 lemma emeasure_compl:
   "s \<in> sets M \<Longrightarrow> emeasure M s \<noteq> \<infinity> \<Longrightarrow> emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
@@ -563,8 +548,7 @@
   also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
     using A finite * by (simp, subst emeasure_Diff) auto
   finally show ?thesis
-    by (rule ennreal_minus_cancel[rotated 3])
-       (insert finite A, auto intro: INF_lower emeasure_mono)
+    by (smt (verit, best) Inf_lower diff_diff_ennreal le_MI finite range_eqI)
 qed
 
 lemma INF_emeasure_decseq':
@@ -580,7 +564,7 @@
   have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))"
   proof (rule INF_eq)
     show "\<exists>j\<in>UNIV. emeasure M (A (j + i)) \<le> emeasure M (A i')" for i'
-      by (intro bexI[of _ i'] emeasure_mono decseqD[OF \<open>decseq A\<close>] A) auto
+      by (meson A \<open>decseq A\<close> decseq_def emeasure_mono iso_tuple_UNIV_I nat_le_iff_add)
   qed auto
   also have "\<dots> = emeasure M (INF n. (A (n + i)))"
     using A \<open>decseq A\<close> fin by (intro INF_emeasure_decseq) (auto simp: decseq_def less_top)
@@ -674,12 +658,12 @@
   assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
   assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> P \<le> lfp F \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
   shows "emeasure (M s) {x\<in>space N. lfp F x} = lfp f s"
-proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric])
+proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and f=F , symmetric])
   fix C assume "incseq C" "\<And>i. Measurable.pred N (C i)"
   then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (SUP i. C i) x}) = (SUP i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
-    unfolding SUP_apply[abs_def]
+    unfolding SUP_apply
     by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
-qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont)
+qed (auto simp add: iter le_fun_def SUP_apply intro!: meas cont)
 
 lemma emeasure_subadditive_finite:
   "finite I \<Longrightarrow> A ` I \<subseteq> sets M \<Longrightarrow> emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
@@ -1118,7 +1102,7 @@
 
 text \<open>The next lemma is convenient to combine with a lemma whose conclusion is of the
 form \<open>AE x in M. P x = Q x\<close>: for such a lemma, there is no \<open>[symmetric]\<close> variant,
-but using \<open>AE_symmetric[OF...]\<close> will replace it.\<close>
+but using \<open>AE_symmetric[OF\<dots>]\<close> will replace it.\<close>
 
 (* depricated replace by laws about eventually *)
 lemma
@@ -1645,11 +1629,9 @@
   have "emeasure M (A \<union> B) \<noteq> \<infinity>"
     using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique)
   then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
-    using emeasure_subadditive[OF measurable] fin
-    apply simp
-    apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure)
-    apply (auto simp flip: ennreal_plus)
-    done
+    unfolding measure_def
+    by (metis emeasure_subadditive[OF measurable] fin   enn2real_mono enn2real_plus 
+        ennreal_add_less_top infinity_ennreal_def less_top)
 qed
 
 lemma measure_subadditive_finite:
@@ -1672,13 +1654,13 @@
   assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
   shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
 proof -
-  from fin have **: "\<And>i. emeasure M (A i) \<noteq> top"
-    using ennreal_suminf_lessD[of "\<lambda>i. emeasure M (A i)"] by (simp add: less_top)
-  { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
-      using emeasure_subadditive_countably[OF A] .
-    also have "\<dots> < \<infinity>"
-      using fin by (simp add: less_top)
-    finally have "emeasure M (\<Union>i. A i) \<noteq> top" by simp }
+  have **: "\<And>i. emeasure M (A i) \<noteq> top"
+    using fin ennreal_suminf_lessD[of "\<lambda>i. emeasure M (A i)"] by (simp add: less_top)
+  have ge0: "(\<Sum>i. Sigma_Algebra.measure M (A i)) \<ge> 0"
+    using fin emeasure_eq_ennreal_measure[OF **]
+    by (metis infinity_ennreal_def measure_nonneg suminf_cong suminf_nonneg summable_suminf_not_top)
+  have "emeasure M (\<Union>i. A i) \<noteq> top"
+    by (metis A emeasure_subadditive_countably fin infinity_ennreal_def neq_top_trans)
   then have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
     by (rule emeasure_eq_ennreal_measure[symmetric])
   also have "\<dots> \<le> (\<Sum>i. emeasure M (A i))"
@@ -1687,11 +1669,7 @@
     using fin unfolding emeasure_eq_ennreal_measure[OF **]
     by (subst suminf_ennreal) (auto simp: **)
   finally show ?thesis
-    apply (rule ennreal_le_iff[THEN iffD1, rotated])
-    apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top)
-    using fin
-    apply (simp add: emeasure_eq_ennreal_measure[OF **])
-    done
+    using ge0 ennreal_le_iff by blast
 qed
 
 lemma measure_Un_null_set: "A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> measure M (A \<union> B) = measure M A"
@@ -2033,16 +2011,16 @@
     apply (rule Lim_emeasure_incseq) unfolding B_def by (auto simp add: SUP_subset_mono incseq_def)
   moreover have "emeasure M (B N) \<le> ennreal (\<Sum>n. measure M (A n))" for N
   proof -
-    have *: "(\<Sum>n\<in>{..<N}. measure M (A n)) \<le> (\<Sum>n. measure M (A n))"
+    have *: "(\<Sum>n<N. measure M (A n)) \<le> (\<Sum>n. measure M (A n))"
       using assms(3) measure_nonneg sum_le_suminf by blast
 
-    have "emeasure M (B N) \<le> (\<Sum>n\<in>{..<N}. emeasure M (A n))"
+    have "emeasure M (B N) \<le> (\<Sum>n<N. emeasure M (A n))"
       unfolding B_def by (rule emeasure_subadditive_finite, auto)
-    also have "... = (\<Sum>n\<in>{..<N}. ennreal(measure M (A n)))"
+    also have "\<dots> = (\<Sum>n<N. ennreal(measure M (A n)))"
       using assms(2) by (simp add: emeasure_eq_ennreal_measure less_top)
-    also have "... = ennreal (\<Sum>n\<in>{..<N}. measure M (A n))"
+    also have "\<dots> = ennreal (\<Sum>n<N. measure M (A n))"
       by auto
-    also have "... \<le> ennreal (\<Sum>n. measure M (A n))"
+    also have "\<dots> \<le> ennreal (\<Sum>n. measure M (A n))"
       using * by (auto simp: ennreal_leI)
     finally show ?thesis by simp
   qed
@@ -2069,9 +2047,9 @@
       have I: "(\<Union>k\<in>{n..}. A k) = (\<Union>k. A (k+n))" by (auto, metis le_add_diff_inverse2, fastforce)
       have "emeasure M (limsup A) \<le> emeasure M (\<Union>k\<in>{n..}. A k)"
         by (rule emeasure_mono, auto simp add: limsup_INF_SUP)
-      also have "... = emeasure M (\<Union>k. A (k+n))"
+      also have "\<dots> = emeasure M (\<Union>k. A (k+n))"
         using I by auto
-      also have "... \<le> (\<Sum>k. measure M (A (k+n)))"
+      also have "\<dots> \<le> (\<Sum>k. measure M (A (k+n)))"
         apply (rule emeasure_union_summable)
         using assms summable_ignore_initial_segment[OF assms(3), of n] by auto
       finally show ?thesis by simp
@@ -2089,12 +2067,8 @@
 proof -
   have "AE x in M. x \<notin> limsup A"
     using borel_cantelli_limsup1[OF assms] unfolding eventually_ae_filter by auto
-  moreover
-  {
-    fix x assume "x \<notin> limsup A"
-    then obtain N where "x \<notin> (\<Union>n\<in>{N..}. A n)" unfolding limsup_INF_SUP by blast
-    then have "eventually (\<lambda>n. x \<notin> A n) sequentially" using eventually_sequentially by auto
-  }
+  moreover have "\<forall>\<^sub>F n in sequentially. x \<notin> A n" if "x \<notin> limsup A" for x
+    using that  by (auto simp: limsup_INF_SUP eventually_sequentially)
   ultimately show ?thesis by auto
 qed
 
@@ -2310,12 +2284,11 @@
   assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
   assumes bound: "\<And>P. f P \<le> f (\<lambda>s. emeasure (M s) (space (M s)))"
   shows "emeasure (M s) {x\<in>space N. gfp F x} = gfp f s"
-proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and
-    P="Measurable.pred N", symmetric])
+proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and P="Measurable.pred N", symmetric])
   interpret finite_measure "M s" for s by fact
   fix C assume "decseq C" "\<And>i. Measurable.pred N (C i)"
   then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (INF i. C i) x}) = (INF i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
-    unfolding INF_apply[abs_def]
+    unfolding INF_apply
     by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
 next
   show "f x \<le> (\<lambda>s. emeasure (M s) {x \<in> space N. F top x})" for x
@@ -2843,20 +2816,8 @@
     show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)"
       by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI)
   qed (simp add: sets_restrict_space)
-  then show ?thesis
-    apply -
-    apply (erule bexE)
-    subgoal for Y
-      apply (intro bexI[of _ Y] conjI ballI conjI)
-         apply (simp_all add: sets_restrict_space emeasure_restrict_space)
-         apply safe
-      subgoal for X Z
-        by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1)
-      subgoal for X Z
-        by (erule ballE[of _ _ X])  (auto simp add: Int_absorb1 ac_simps)
-      apply auto
-      done
-    done
+  with assms show ?thesis
+    by (metis Int_subset_iff emeasure_restrict_space sets.Int_space_eq2 sets_restrict_space_iff space_restrict_space)
 qed
 
 text\<^marker>\<open>tag important\<close> \<open>
@@ -2898,11 +2859,7 @@
 end
 
 proposition le_measure: "sets M = sets N \<Longrightarrow> M \<le> N \<longleftrightarrow> (\<forall>A\<in>sets M. emeasure M A \<le> emeasure N A)"
-  apply -
-  apply (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq)
-  subgoal for X
-    by (cases "X \<in> sets M") (auto simp: emeasure_notin_sets)
-  done
+  by (metis emeasure_neq_0_sets le_fun_def le_measure_iff order_class.order_eq_iff sets_eq_imp_space_eq)
 
 definition\<^marker>\<open>tag important\<close> sup_measure' :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
 "sup_measure' A B =
@@ -2930,7 +2887,7 @@
       then have [measurable]: "\<And>i. X i \<in> sets A" and "disjoint_family X"
         by auto
       have disjoint: "disjoint_family (\<lambda>i. X i \<inter> Y)" "disjoint_family (\<lambda>i. X i - Y)" for Y
-        by (auto intro: disjoint_family_on_bisimulation [OF \<open>disjoint_family X\<close>, simplified])
+        using "1"(2) disjoint_family_subset by fastforce+
       have "(\<Sum>i. ?S (X i)) = (SUP Y\<in>sets A. \<Sum>i. ?d (X i) Y)"
       proof (rule ennreal_suminf_SUP_eq_directed)
         fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \<in> sets A" "b \<in> sets A"
@@ -2938,13 +2895,7 @@
         proof cases
           assume "emeasure A (X i) = top \<or> emeasure B (X i) = top"
           then show ?thesis
-          proof
-            assume "emeasure A (X i) = top" then show ?thesis
-              by (intro bexI[of _ "X i"]) auto
-          next
-            assume "emeasure B (X i) = top" then show ?thesis
-              by (intro bexI[of _ "{}"]) auto
-          qed
+            by force
         next
           assume finite: "\<not> (emeasure A (X i) = top \<or> emeasure B (X i) = top)"
           then have "\<exists>Y\<in>sets A. Y \<subseteq> X i \<and> (\<forall>C\<in>sets A. C \<subseteq> Y \<longrightarrow> B C \<le> A C) \<and> (\<forall>C\<in>sets A. C \<subseteq> X i \<longrightarrow> C \<inter> Y = {} \<longrightarrow> A C \<le> B C)"
@@ -2955,7 +2906,7 @@
             by auto
 
           show ?thesis
-          proof (intro bexI[of _ Y] ballI conjI)
+          proof (intro bexI ballI conjI)
             fix a assume [measurable]: "a \<in> sets A"
             have *: "(X i \<inter> a \<inter> Y \<union> (X i \<inter> a - Y)) = X i \<inter> a" "(X i - a) \<inter> Y \<union> (X i - a - Y) = X i \<inter> - a"
               for a Y by auto
@@ -2975,36 +2926,22 @@
           and C: "\<And>a. a \<in> sets A \<Longrightarrow> ?d (X i) a \<le> ?d (X i) (C i)" for i
           by metis
         have *: "X i \<inter> (\<Union>i. C i) = X i \<inter> C i" for i
-        proof safe
-          fix x j assume "x \<in> X i" "x \<in> C j"
-          moreover have "i = j \<or> X i \<inter> X j = {}"
-            using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def)
-          ultimately show "x \<in> C i"
-            using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto
-        qed auto
-        have **: "X i \<inter> - (\<Union>i. C i) = X i \<inter> - C i" for i
-        proof safe
-          fix x j assume "x \<in> X i" "x \<notin> C i" "x \<in> C j"
-          moreover have "i = j \<or> X i \<inter> X j = {}"
-            using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def)
-          ultimately show False
-            using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto
-        qed auto
-        show "\<exists>c\<in>sets A. \<forall>i\<in>J. ?d (X i) a \<le> ?d (X i) c \<and> ?d (X i) b \<le> ?d (X i) c"
-          apply (intro bexI[of _ "\<Union>i. C i"])
-          unfolding * **
-          apply (intro C ballI conjI)
-          apply auto
-          done
+          using \<open>disjoint_family X\<close> \<open>\<And>i. C i \<subseteq> X i\<close>
+          by (simp add: disjoint_family_on_def disjoint_iff_not_equal set_eq_iff) (metis subsetD)
+        then have **: "X i \<inter> - (\<Union>i. C i) = X i \<inter> - C i" for i by blast
+        moreover have "(\<Union>i. C i) \<in> sets A"
+          by fastforce
+        ultimately show "\<exists>c\<in>sets A. \<forall>i\<in>J. ?d (X i) a \<le> ?d (X i) c \<and> ?d (X i) b \<le> ?d (X i) c"
+          by (metis "*" C \<open>a \<in> sets A\<close> \<open>b \<in> sets A\<close>)
       qed
       also have "\<dots> = ?S (\<Union>i. X i)"
-        apply (simp only: UN_extend_simps(4))
-        apply (rule arg_cong [of _ _ Sup])
-        apply (rule image_cong)
-         apply (fact refl)
-        using disjoint
-        apply (auto simp add: suminf_add [symmetric] Diff_eq [symmetric] image_subset_iff suminf_emeasure simp del: UN_simps)
-        done
+      proof -
+        have "\<And>Y. Y \<in> sets A \<Longrightarrow> (\<Sum>i. emeasure A (X i \<inter> Y) + emeasure B (X i \<inter> -Y)) 
+                              = emeasure A (\<Union>i. X i \<inter> Y) + emeasure B (\<Union>i. X i \<inter> -Y)"
+          using disjoint
+          by (auto simp flip: suminf_add Diff_eq simp add: image_subset_iff suminf_emeasure)
+        then show ?thesis by force
+      qed
       finally show "(\<Sum>i. ?S (X i)) = ?S (\<Union>i. X i)" .
     qed
   qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const)
@@ -3290,10 +3227,7 @@
         show "(\<Sum>n. ?\<mu> i (F n)) = ?\<mu> i (\<Union>(F ` UNIV))"
         proof cases
           assume "i \<noteq> {}" with i ** show ?thesis
-            apply (intro suminf_emeasure \<open>disjoint_family F\<close>)
-            apply (subst sets_sup_measure_F[OF _ _ sets_eq])
-            apply auto
-            done
+            by (smt (verit, best) "1"(2) Measure_Space.sets_sup_measure_F assms(1) mem_Collect_eq subset_eq suminf_cong suminf_emeasure)
         qed simp
       qed
     qed
@@ -3559,33 +3493,27 @@
 
 subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Supremum of a set of \<open>\<sigma>\<close>-algebras\<close>
 
-lemma space_Sup_eq_UN: "space (Sup M) = (\<Union>x\<in>M. space x)"
-  unfolding Sup_measure_def
-  apply (intro Sup_lexord[where P="\<lambda>x. space x = _"])
-  apply (subst space_Sup_measure'2)
-  apply auto []
-  apply (subst space_measure_of[OF UN_space_closed])
-  apply auto
-  done
+lemma space_Sup_eq_UN: "space (Sup M) = (\<Union>x\<in>M. space x)" (is "?L=?R")
+proof
+  show "?L \<subseteq> ?R"
+    using Sup_lexord[where P="\<lambda>x. space x = _"]
+    apply (clarsimp simp: Sup_measure_def)
+    by (smt (verit) Sup_lexord_def UN_E mem_Collect_eq space_Sup_measure'2 space_measure_of_conv)
+qed (use Sup_upper le_measureD1 in fastforce)
+
 
 lemma sets_Sup_eq:
   assumes *: "\<And>m. m \<in> M \<Longrightarrow> space m = X" and "M \<noteq> {}"
   shows "sets (Sup M) = sigma_sets X (\<Union>x\<in>M. sets x)"
   unfolding Sup_measure_def
-  apply (rule Sup_lexord1)
-  apply fact
-  apply (simp add: assms)
+proof (rule Sup_lexord1 [OF \<open>M \<noteq> {}\<close>])
+  show "sets (Sup_lexord sets Sup_measure' (\<lambda>U. sigma (\<Union> (space ` U)) (\<Union> (sets ` U))) M)
+      = sigma_sets X (\<Union> (sets ` M))"
   apply (rule Sup_lexord)
-  subgoal premises that for a S
-    unfolding that(3) that(2)[symmetric]
-    using that(1)
-    apply (subst sets_Sup_measure'2)
-    apply (intro arg_cong2[where f=sigma_sets])
-    apply (auto simp: *)
-    done
-  apply (subst sets_measure_of[OF UN_space_closed])
-  apply (simp add:  assms)
-  done
+  apply (metis (mono_tags, lifting) "*" empty_iff mem_Collect_eq sets.sigma_sets_eq sets_Sup_measure')
+  by (metis "*" SUP_eq_const UN_space_closed assms(2) sets_measure_of)
+qed (use * in blast)
+
 
 lemma in_sets_Sup: "(\<And>m. m \<in> M \<Longrightarrow> space m = X) \<Longrightarrow> m \<in> M \<Longrightarrow> A \<in> sets m \<Longrightarrow> A \<in> sets (Sup M)"
   by (subst sets_Sup_eq[where X=X]) auto
@@ -3605,16 +3533,11 @@
 qed
 
 lemma sets_SUP_cong:
-  assumes eq: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (SUP i\<in>I. M i) = sets (SUP i\<in>I. N i)"
+  assumes eq: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i)" 
+  shows "sets (SUP i\<in>I. M i) = sets (SUP i\<in>I. N i)"
   unfolding Sup_measure_def
   using eq eq[THEN sets_eq_imp_space_eq]
-  apply (intro Sup_lexord_rel[where R="\<lambda>x y. sets x = sets y"])
-  apply simp
-  apply simp
-  apply (simp add: sets_Sup_measure'2)
-  apply (intro arg_cong2[where f="\<lambda>x y. sets (sigma x y)"])
-  apply auto
-  done
+  by (intro Sup_lexord_rel[where R="\<lambda>x y. sets x = sets y"], simp_all add: sets_Sup_measure'2)
 
 lemma sets_Sup_in_sets:
   assumes "M \<noteq> {}"
@@ -3648,18 +3571,15 @@
   from M obtain m where "m \<in> M" by auto
   have space_eq: "\<And>n. n \<in> M \<Longrightarrow> space n = space m"
     by (intro const_space \<open>m \<in> M\<close>)
+  have eq: "sets (sigma (\<Union> (space ` M)) (\<Union> (sets ` M))) = sets (Sup M)"
+    by (metis M SUP_eq_const UN_space_closed sets_Sup_eq sets_measure_of space_eq)
   have "f \<in> measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m))"
   proof (rule measurable_measure_of)
     show "f \<in> space N \<rightarrow> \<Union>(space ` M)"
       using measurable_space[OF f] M by auto
   qed (auto intro: measurable_sets f dest: sets.sets_into_space)
   also have "measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)) = measurable N (Sup M)"
-    apply (intro measurable_cong_sets refl)
-    apply (subst sets_Sup_eq[OF space_eq M])
-    apply simp
-    apply (subst sets_measure_of[OF UN_space_closed])
-    apply (simp add: space_eq M)
-    done
+    using eq measurable_cong_sets by blast
   finally show ?thesis .
 qed
 
@@ -3674,15 +3594,11 @@
 proof -
   { fix a m assume "a \<in> sigma_sets \<Omega> m" "m \<in> M"
     then have "a \<in> sigma_sets \<Omega> (\<Union>M)"
-     by induction (auto intro: sigma_sets.intros(2-)) }
+      by induction (auto intro: sigma_sets.intros(2-)) }
+  then have "sigma_sets \<Omega> (\<Union> (sigma_sets \<Omega> ` M)) = sigma_sets \<Omega> (\<Union> M)"
+    by (smt (verit, best) UN_iff Union_iff sigma_sets.Basic sigma_sets_eqI)
   then show "sets (SUP m\<in>M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))"
-    apply (subst sets_Sup_eq[where X="\<Omega>"])
-    apply (auto simp add: M) []
-    apply auto []
-    apply (simp add: space_measure_of_conv M Union_least)
-    apply (rule sigma_sets_eqI)
-    apply auto
-    done
+    by (subst sets_Sup_eq) (fastforce simp add: M Union_least)+
 qed
 
 lemma Sup_sigma:
@@ -3694,9 +3610,8 @@
   show "sigma \<Omega> (\<Union>M) \<le> (SUP m\<in>M. sigma \<Omega> m)"
   proof (intro less_eq_measure.intros(3))
     show "space (sigma \<Omega> (\<Union>M)) = space (SUP m\<in>M. sigma \<Omega> m)"
-      "sets (sigma \<Omega> (\<Union>M)) = sets (SUP m\<in>M. sigma \<Omega> m)"
-      using sets_Sup_sigma[OF assms] sets_Sup_sigma[OF assms, THEN sets_eq_imp_space_eq]
-      by auto
+         "sets (sigma \<Omega> (\<Union>M)) = sets (SUP m\<in>M. sigma \<Omega> m)"
+      by (auto simp add: M sets_Sup_sigma sets_eq_imp_space_eq space_measure_of_conv)
   qed (simp add: emeasure_sigma le_fun_def)
   fix m assume "m \<in> M" then show "sigma \<Omega> m \<le> sigma \<Omega> (\<Union>M)"
     by (subst sigma_le_iff) (auto simp add: M *)
@@ -3709,29 +3624,17 @@
 lemma sets_vimage_Sup_eq:
   assumes *: "M \<noteq> {}" "f \<in> X \<rightarrow> Y" "\<And>m. m \<in> M \<Longrightarrow> space m = Y"
   shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m \<in> M. vimage_algebra X f m)"
-  (is "?IS = ?SI")
+  (is "?L = ?R")
 proof
-  show "?IS \<subseteq> ?SI"
-    apply (intro sets_image_in_sets measurable_Sup2)
-    apply (simp add: space_Sup_eq_UN *)
-    apply (simp add: *)
-    apply (intro measurable_Sup1)
-    apply (rule imageI)
-    apply assumption
-    apply (rule measurable_vimage_algebra1)
-    apply (auto simp: *)
-    done
-  show "?SI \<subseteq> ?IS"
+  have "\<And>m. m \<in> M \<Longrightarrow> f \<in> Sup (vimage_algebra X f ` M) \<rightarrow>\<^sub>M m"
+    using assms
+    by (smt (verit, del_insts) Pi_iff imageE image_eqI measurable_Sup1
+            measurable_vimage_algebra1 space_vimage_algebra)
+  then show "?L \<subseteq> ?R"
+     by (intro sets_image_in_sets measurable_Sup2) (simp_all add: space_Sup_eq_UN *)
+  show "?R \<subseteq> ?L"
     apply (intro sets_Sup_in_sets)
-    apply (auto simp: *) []
-    apply (auto simp: *) []
-    apply (elim imageE)
-    apply simp
-    apply (rule sets_image_in_sets)
-    apply simp
-    apply (simp add: measurable_def)
-    apply (simp add: * space_Sup_eq_UN sets_vimage_algebra2)
-    apply (auto intro: in_sets_Sup[OF *(3)])
+    apply (force simp add: * space_Sup_eq_UN sets_vimage_algebra2 intro: in_sets_Sup)+
     done
 qed
 
@@ -3765,13 +3668,8 @@
 
 lemma measurable_iff_sets:
   "f \<in> measurable M N \<longleftrightarrow> (f \<in> space M \<rightarrow> space N \<and> sets (vimage_algebra (space M) f N) \<subseteq> sets M)"
-proof -
-  have *: "{f -` A \<inter> space M |A. A \<in> sets N} \<subseteq> Pow (space M)"
-    by auto
-  show ?thesis
     unfolding measurable_def
-    by (auto simp add: vimage_algebra_def sigma_le_sets[OF *])
-qed
+    by (smt (verit, ccfv_threshold) mem_Collect_eq sets_vimage_algebra sigma_sets_le_sets_iff subset_eq)
 
 lemma sets_vimage_algebra_space: "X \<in> sets (vimage_algebra X f M)"
   using sets.top[of "vimage_algebra X f M"] by simp
@@ -3786,7 +3684,7 @@
   moreover assume "\<forall>y\<in>sets N. f -` y \<inter> space M \<in> sets M" note this[THEN bspec, of A]
   ultimately show "f -` A \<inter> space M' \<in> sets M'"
     using assms by auto
-qed (insert N M, auto)
+qed (use N M in auto)
 
 lemma measurable_Sup_measurable:
   assumes f: "f \<in> space N \<rightarrow> A"
@@ -3802,7 +3700,7 @@
   shows "sigma_sets \<Omega>' a \<subseteq> M"
 proof
   show "x \<in> M" if x: "x \<in> sigma_sets \<Omega>' a" for x
-    using x by (induct rule: sigma_sets.induct) (insert a, auto)
+    using x by (induct rule: sigma_sets.induct) (use a in auto)
 qed
 
 lemma in_sets_SUP: "i \<in> I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> space (M i) = Y) \<Longrightarrow> X \<in> sets (M i) \<Longrightarrow> X \<in> sets (SUP i\<in>I. M i)"
@@ -3824,20 +3722,13 @@
   "sets M \<le> sets N \<Longrightarrow> sets (vimage_algebra X f M) \<subseteq> sets (vimage_algebra X f N)"
   using sets.top[of "sigma X {f -` A \<inter> X |A. A \<in> sets N}"]
   unfolding vimage_algebra_def
-  apply (subst (asm) space_measure_of)
-  apply auto []
-  apply (subst sigma_le_sets)
-  apply auto
-  done
+  by (smt (verit, del_insts) space_measure_of sigma_le_sets Pow_iff inf_le2 mem_Collect_eq subset_eq)
 
 lemma mono_restrict_space: "sets M \<le> sets N \<Longrightarrow> sets (restrict_space M X) \<subseteq> sets (restrict_space N X)"
   unfolding sets_restrict_space by (rule image_mono)
 
 lemma sets_eq_bot: "sets M = {{}} \<longleftrightarrow> M = bot"
-  apply safe
-  apply (intro measure_eqI)
-  apply auto
-  done
+  by (metis measure_eqI emeasure_empty sets_bot singletonD)
 
 lemma sets_eq_bot2: "{{}} = sets M \<longleftrightarrow> M = bot"
   using sets_eq_bot[of M] by blast
@@ -3847,10 +3738,8 @@
   "countable {x. measure M {x} \<noteq> 0}"
 proof cases
   assume "measure M (space M) = 0"
-  with bounded_measure measure_le_0_iff have "{x. measure M {x} \<noteq> 0} = {}"
-    by auto
-  then show ?thesis
-    by simp
+  then show ?thesis 
+    by (metis (mono_tags, lifting) bounded_measure measure_le_0_iff Collect_empty_eq countable_empty) 
 next
   let ?M = "measure M (space M)" and ?m = "\<lambda>x. measure M {x}"
   assume "?M \<noteq> 0"
@@ -3862,7 +3751,7 @@
     fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
     then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X"
       by (metis infinite_arbitrarily_large)
-    from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x"
+    then have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x"
       by auto
     { fix x assume "x \<in> X"
       from \<open>?M \<noteq> 0\<close> *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff)
--- a/src/HOL/Analysis/Path_Connected.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Path_Connected.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -13,30 +13,32 @@
 subsection \<open>Paths and Arcs\<close>
 
 definition\<^marker>\<open>tag important\<close> path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
-  where "path g \<longleftrightarrow> continuous_on {0..1} g"
+  where "path g \<equiv> continuous_on {0..1} g"
 
 definition\<^marker>\<open>tag important\<close> pathstart :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
-  where "pathstart g = g 0"
+  where "pathstart g \<equiv> g 0"
 
 definition\<^marker>\<open>tag important\<close> pathfinish :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
-  where "pathfinish g = g 1"
+  where "pathfinish g \<equiv> g 1"
 
 definition\<^marker>\<open>tag important\<close> path_image :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a set"
-  where "path_image g = g ` {0 .. 1}"
+  where "path_image g \<equiv> g ` {0 .. 1}"
 
 definition\<^marker>\<open>tag important\<close> reversepath :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
-  where "reversepath g = (\<lambda>x. g(1 - x))"
+  where "reversepath g \<equiv> (\<lambda>x. g(1 - x))"
 
 definition\<^marker>\<open>tag important\<close> joinpaths :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a"
     (infixr "+++" 75)
-  where "g1 +++ g2 = (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
+  where "g1 +++ g2 \<equiv> (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
+
+definition\<^marker>\<open>tag important\<close> loop_free :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
+  where "loop_free g \<equiv> \<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
 
 definition\<^marker>\<open>tag important\<close> simple_path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
-  where "simple_path g \<longleftrightarrow>
-     path g \<and> (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
+  where "simple_path g \<equiv> path g \<and> loop_free g"
 
 definition\<^marker>\<open>tag important\<close> arc :: "(real \<Rightarrow> 'a :: topological_space) \<Rightarrow> bool"
-  where "arc g \<longleftrightarrow> path g \<and> inj_on g {0..1}"
+  where "arc g \<equiv> path g \<and> inj_on g {0..1}"
 
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Invariance theorems\<close>
@@ -71,12 +73,8 @@
   from linear_injective_left_inverse [OF assms]
   obtain h where h: "linear h" "h \<circ> f = id"
     by blast
-  then have g: "g = h \<circ> (f \<circ> g)"
-    by (metis comp_assoc id_comp)
-  show ?thesis
-    unfolding path_def
-    using h assms
-    by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear)
+  with assms show ?thesis
+    by (metis comp_assoc id_comp linear_continuous_on linear_linear path_continuous_image)
 qed
 
 lemma pathstart_translation: "pathstart((\<lambda>x. a + x) \<circ> g) = a + pathstart g"
@@ -110,21 +108,32 @@
 lemma joinpaths_linear_image: "linear f \<Longrightarrow> (f \<circ> g1) +++ (f \<circ> g2) = f \<circ> (g1 +++ g2)"
   by (rule ext) (simp add: joinpaths_def)
 
+lemma loop_free_translation_eq:
+  fixes g :: "real \<Rightarrow> 'a::euclidean_space"
+  shows "loop_free((\<lambda>x. a + x) \<circ> g) = loop_free g"
+  by (simp add: loop_free_def)
+
 lemma simple_path_translation_eq:
   fixes g :: "real \<Rightarrow> 'a::euclidean_space"
   shows "simple_path((\<lambda>x. a + x) \<circ> g) = simple_path g"
-  by (simp add: simple_path_def path_translation_eq)
+  by (simp add: simple_path_def loop_free_translation_eq path_translation_eq)
+
+lemma loop_free_linear_image_eq:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "linear f" "inj f"
+    shows "loop_free(f \<circ> g) = loop_free g"
+  using assms inj_on_eq_iff [of f] by (auto simp: loop_free_def)
 
 lemma simple_path_linear_image_eq:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "linear f" "inj f"
     shows "simple_path(f \<circ> g) = simple_path g"
-  using assms inj_on_eq_iff [of f]
-  by (auto simp: path_linear_image_eq simple_path_def path_translation_eq)
+  using assms
+  by (simp add: loop_free_linear_image_eq path_linear_image_eq simple_path_def)
 
 lemma arc_translation_eq:
   fixes g :: "real \<Rightarrow> 'a::euclidean_space"
-  shows "arc((\<lambda>x. a + x) \<circ> g) = arc g"
+  shows "arc((\<lambda>x. a + x) \<circ> g) \<longleftrightarrow> arc g"
   by (auto simp: arc_def inj_on_def path_translation_eq)
 
 lemma arc_linear_image_eq:
@@ -161,8 +170,11 @@
 lemma continuous_on_path: "path f \<Longrightarrow> t \<subseteq> {0..1} \<Longrightarrow> continuous_on t f"
   using continuous_on_subset path_def by blast
 
+lemma inj_on_imp_loop_free: "inj_on g {0..1} \<Longrightarrow> loop_free g"
+  by (simp add: inj_onD loop_free_def)
+
 lemma arc_imp_simple_path: "arc g \<Longrightarrow> simple_path g"
-  by (simp add: arc_def inj_on_def simple_path_def)
+  by (simp add: arc_def inj_on_imp_loop_free simple_path_def)
 
 lemma arc_imp_path: "arc g \<Longrightarrow> path g"
   using arc_def by blast
@@ -173,9 +185,11 @@
 lemma simple_path_imp_path: "simple_path g \<Longrightarrow> path g"
   using simple_path_def by blast
 
+lemma loop_free_cases: "loop_free g \<Longrightarrow> inj_on g {0..1} \<or> pathfinish g = pathstart g"
+  by (force simp: inj_on_def loop_free_def pathfinish_def pathstart_def)
+
 lemma simple_path_cases: "simple_path g \<Longrightarrow> arc g \<or> pathfinish g = pathstart g"
-  unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def
-  by force
+  using arc_def loop_free_cases simple_path_def by blast
 
 lemma simple_path_imp_arc: "simple_path g \<Longrightarrow> pathfinish g \<noteq> pathstart g \<Longrightarrow> arc g"
   using simple_path_cases by auto
@@ -250,12 +264,9 @@
 lemma path_reversepath [simp]: "path (reversepath g) \<longleftrightarrow> path g"
 proof -
   have *: "\<And>g. path g \<Longrightarrow> path (reversepath g)"
-    unfolding path_def reversepath_def
-    apply (rule continuous_on_compose[unfolded o_def, of _ "\<lambda>x. 1 - x"])
-    apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"])
-    done
-  show ?thesis
-    using "*" by force
+    by (metis cancel_comm_monoid_add_class.diff_cancel continuous_on_compose 
+        continuous_on_op_minus diff_zero image_diff_atLeastAtMost path_def reversepath_o)
+  then show ?thesis by force
 qed
 
 lemma arc_reversepath:
@@ -270,10 +281,12 @@
     using assms  by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **)
 qed
 
+lemma loop_free_reversepath:
+  assumes "loop_free g" shows "loop_free(reversepath g)"
+  using assms by (simp add: reversepath_def loop_free_def Ball_def) (smt (verit))
+
 lemma simple_path_reversepath: "simple_path g \<Longrightarrow> simple_path (reversepath g)"
-  apply (simp add: simple_path_def)
-  apply (force simp: reversepath_def)
-  done
+  by (simp add: loop_free_reversepath simple_path_def)
 
 lemmas reversepath_simps =
   path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath
@@ -362,8 +375,7 @@
   by auto
 
 lemma subset_path_image_join:
-  assumes "path_image g1 \<subseteq> s"
-    and "path_image g2 \<subseteq> s"
+  assumes "path_image g1 \<subseteq> s" and "path_image g2 \<subseteq> s"
   shows "path_image (g1 +++ g2) \<subseteq> s"
   using path_image_join_subset[of g1 g2] and assms
   by auto
@@ -394,8 +406,7 @@
 qed
 
 lemma not_in_path_image_join:
-  assumes "x \<notin> path_image g1"
-    and "x \<notin> path_image g2"
+  assumes "x \<notin> path_image g1" and "x \<notin> path_image g2"
   shows "x \<notin> path_image (g1 +++ g2)"
   using assms and path_image_join_subset[of g1 g2]
   by auto
@@ -421,8 +432,11 @@
    \<Longrightarrow>  t \<in> {0..1} \<Longrightarrow> (p +++ q) t = (p' +++ q') t"
   by (auto simp: joinpaths_def)
 
+lemma loop_free_inj_on: "loop_free g \<Longrightarrow> inj_on g {0<..<1}"
+  by (force simp: inj_on_def loop_free_def)
+
 lemma simple_path_inj_on: "simple_path g \<Longrightarrow> inj_on g {0<..<1}"
-  by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def)
+  using loop_free_inj_on simple_path_def by auto
 
 
 subsection\<^marker>\<open>tag unimportant\<close>\<open>Simple paths with the endpoints removed\<close>
@@ -435,8 +449,8 @@
     using less_eq_real_def by (auto simp: path_image_def pathstart_def pathfinish_def)
   show "?rhs \<subseteq> ?lhs"
     using assms 
-    apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def)
-    using less_eq_real_def zero_le_one by blast+
+    apply (simp add: image_subset_iff path_image_def pathstart_def pathfinish_def simple_path_def loop_free_def Ball_def)
+    by (smt (verit))
 qed
 
 lemma connected_simple_path_endless:
@@ -470,21 +484,16 @@
 lemma continuous_on_joinpaths:
   assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2"
     shows "continuous_on {0..1} (g1 +++ g2)"
-proof -
-  have "{0..1::real} = {0..1/2} \<union> {1/2..1}"
-    by auto
-  then show ?thesis
-    using assms by (metis path_def path_join)
-qed
+  using assms path_def path_join by blast
 
 lemma path_join_imp: "\<lbrakk>path g1; path g2; pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> path(g1 +++ g2)"
   by simp
 
-lemma simple_path_join_loop:
+lemma arc_join:
   assumes "arc g1" "arc g2"
-          "pathfinish g1 = pathstart g2"  "pathfinish g2 = pathstart g1"
-          "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
-  shows "simple_path(g1 +++ g2)"
+          "pathfinish g1 = pathstart g2"
+          "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
+    shows "arc(g1 +++ g2)"
 proof -
   have injg1: "inj_on g1 {0..1}"
     using assms
@@ -492,6 +501,32 @@
   have injg2: "inj_on g2 {0..1}"
     using assms
     by (simp add: arc_def)
+  have g11: "g1 1 = g2 0"
+   and sb:  "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g2 0}"
+    using assms
+    by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def)
+  { fix x and y::real
+    assume xy: "g2 (2 * x - 1) = g1 (2 * y)" "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1"
+    then have "g1 (2 * y) = g2 0"
+      using sb by force
+    then have False
+      using xy inj_onD injg2 by fastforce
+   } note * = this
+  have "inj_on (g1 +++ g2) {0..1}"
+    using inj_onD [OF injg1] inj_onD [OF injg2] *
+    by (simp add: inj_on_def joinpaths_def Ball_def) (smt (verit))
+  then show ?thesis
+    using arc_def assms path_join_imp by blast
+qed
+
+lemma simple_path_join_loop:
+  assumes "arc g1" "arc g2"
+          "pathfinish g1 = pathstart g2"  "pathfinish g2 = pathstart g1"
+          "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
+        shows "simple_path(g1 +++ g2)"
+proof -
+  have injg1: "inj_on g1 {0..1}" and injg2: "inj_on g2 {0..1}"
+    using assms by (auto simp add: arc_def)
   have g12: "g1 1 = g2 0"
    and g21: "g2 1 = g1 0"
    and sb:  "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g1 0, g2 0}"
@@ -516,49 +551,12 @@
         using g2_eq g12 inj_onD [OF injg2] atLeastAtMost_iff xy(1) xy(4) by fastforce
       with xy show False by auto
     qed
-  } note * = this
-  { fix x and y::real
-    assume xy: "g1 (2 * x) = g2 (2 * y - 1)" "y \<le> 1" "0 \<le> x" "\<not> y * 2 \<le> 1" "x * 2 \<le> 1" 
-    then have "x = 0 \<and> y = 1"
-      using * xy by force
-   } note ** = this
-  show ?thesis
-    using assms
-    apply (simp add: arc_def simple_path_def)
-    apply (auto simp: joinpaths_def split: if_split_asm 
-                dest!: * ** dest: inj_onD [OF injg1] inj_onD [OF injg2])
-    done
-qed
-
-lemma arc_join:
-  assumes "arc g1" "arc g2"
-          "pathfinish g1 = pathstart g2"
-          "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
-    shows "arc(g1 +++ g2)"
-proof -
-  have injg1: "inj_on g1 {0..1}"
-    using assms
-    by (simp add: arc_def)
-  have injg2: "inj_on g2 {0..1}"
-    using assms
-    by (simp add: arc_def)
-  have g11: "g1 1 = g2 0"
-   and sb:  "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g2 0}"
-    using assms
-    by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def)
-  { fix x and y::real
-    assume xy: "g2 (2 * x - 1) = g1 (2 * y)" "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1"
-    then have "g1 (2 * y) = g2 0"
-      using sb by force
-    then have False
-      using xy inj_onD injg2 by fastforce
-   } note * = this
-  show ?thesis
-    using assms
-    apply (simp add: arc_def inj_on_def)
-    apply (auto simp: joinpaths_def arc_imp_path split: if_split_asm 
-                dest: * *[OF sym] inj_onD [OF injg1] inj_onD [OF injg2])
-    done
+  } note * = this 
+  have "loop_free(g1 +++ g2)"
+    using inj_onD [OF injg1] inj_onD [OF injg2] *
+    by (simp add: loop_free_def joinpaths_def Ball_def) (smt (verit))
+  then show ?thesis
+    by (simp add: arc_imp_path assms simple_path_def)
 qed
 
 lemma reversepath_joinpaths:
@@ -625,7 +623,7 @@
 proof -
   have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk>
                \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
-    using assms by (simp add: simple_path_def)
+    using assms by (simp add: simple_path_def loop_free_def)
   have "path g1"
     using assms path_join simple_path_imp_path by blast
   moreover have "inj_on g1 {0..1}"
@@ -646,7 +644,7 @@
     fix x y
     assume "g2 x = g2 y" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1"
     then show "x = y"
-      using * [of "(x + 1) / 2" "(y + 1) / 2"]
+      using * [of "(x+1) / 2" "(y+1) / 2"]
       by (force simp: joinpaths_def split_ifs field_split_simps)
   qed
   ultimately have "arc g2"
@@ -672,19 +670,10 @@
            arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
            (is "?lhs = ?rhs")
 proof
-  assume ?lhs
-  then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path)
-  then have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk>
-               \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
-    using assms by (simp add: simple_path_def)
-  have False if "g1 0 = g2 u" "0 \<le> u" "u \<le> 1" for u
-    using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \<open>?lhs\<close>]
-    by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps)
-  then have n1: "pathstart g1 \<notin> path_image g2"
-    unfolding pathstart_def path_image_def
-    using atLeastAtMost_iff by blast
-  show ?rhs using \<open>?lhs\<close>
-    using \<open>simple_path (g1 +++ g2)\<close> assms n1 simple_path_joinE by auto
+  assume ?lhs then show ?rhs 
+    using reversepath_simps assms
+    by (smt (verit, ccfv_threshold) Int_commute arc_distinct_ends arc_imp_simple_path arc_reversepath 
+            in_mono insertE pathfinish_join reversepath_joinpaths simple_path_joinE subsetI)
 next
   assume ?rhs then show ?lhs
     using assms
@@ -694,8 +683,7 @@
 lemma arc_join_eq_alt:
         "pathfinish g1 = pathstart g2
         \<Longrightarrow> (arc(g1 +++ g2) \<longleftrightarrow>
-             arc g1 \<and> arc g2 \<and>
-             path_image g1 \<inter> path_image g2 = {pathstart g2})"
+             arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 = {pathstart g2})"
 using pathfinish_in_path_image by (fastforce simp: arc_join_eq)
 
 
@@ -791,11 +779,10 @@
   assumes "path g" "u \<in> {0..1}" "v \<in> {0..1}"
     shows "path(subpath u v g)"
 proof -
-  have "continuous_on {0..1} (g \<circ> (\<lambda>x. ((v-u) * x+ u)))"
-    using assms
-    apply (intro continuous_intros; simp add: image_affinity_atLeastAtMost [where c=u])
-    apply (auto simp: path_def continuous_on_subset)
-    done
+  have "continuous_on {u..v} g" "continuous_on {v..u} g"
+    using assms continuous_on_path by fastforce+
+  then have "continuous_on {0..1} (g \<circ> (\<lambda>x. ((v-u) * x+ u)))"
+    by (intro continuous_intros; simp add: image_affinity_atLeastAtMost [where c=u])
   then show ?thesis
     by (simp add: path_def subpath_def)
 qed
@@ -847,7 +834,7 @@
   then have p: "path (\<lambda>x. g ((v - u) * x + u))"
         and sim: "(\<And>x y. \<lbrakk>x\<in>{0..1}; y\<in>{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\<rbrakk>
                   \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
-    by (auto simp: simple_path_def subpath_def)
+    by (auto simp: simple_path_def loop_free_def subpath_def)
   { fix x y
     assume "x \<in> closed_segment u v" "y \<in> closed_segment u v" "g x = g y"
     then have "x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u"
@@ -871,45 +858,13 @@
   have [simp]: "\<And>x. u + x * v = v + x * u \<longleftrightarrow> u=v \<or> x=1"
     by algebra
   show ?lhs using psp ne
-    unfolding simple_path_def subpath_def
+    unfolding simple_path_def loop_free_def subpath_def
     by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2)
 qed
 
 lemma arc_subpath_eq:
   "arc(subpath u v g) \<longleftrightarrow> path(subpath u v g) \<and> u\<noteq>v \<and> inj_on g (closed_segment u v)"
-    (is "?lhs = ?rhs")
-proof 
-  assume ?lhs
-  then have p: "path (\<lambda>x. g ((v - u) * x + u))"
-        and sim: "(\<And>x y. \<lbrakk>x\<in>{0..1}; y\<in>{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\<rbrakk>
-                  \<Longrightarrow> x = y)"
-    by (auto simp: arc_def inj_on_def subpath_def)
-  { fix x y
-    assume "x \<in> closed_segment u v" "y \<in> closed_segment u v" "g x = g y"
-    then have "x = y"
-      using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p
-      by (cases "v = u")
-        (simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost,
-           simp add: field_simps)
-  } moreover
-  have "path(subpath u v g) \<and> u\<noteq>v"
-    using sim [of "1/3" "2/3"] p
-    by (auto simp: subpath_def)
-  ultimately show ?rhs
-    unfolding inj_on_def
-    by metis
-next
-  assume ?rhs
-  then
-  have d1: "\<And>x y. \<lbrakk>g x = g y; u \<le> x; x \<le> v; u \<le> y; y \<le> v\<rbrakk> \<Longrightarrow> x = y"
-   and d2: "\<And>x y. \<lbrakk>g x = g y; v \<le> x; x \<le> u; v \<le> y; y \<le> u\<rbrakk> \<Longrightarrow> x = y"
-   and ne: "u < v \<or> v < u"
-   and psp: "path (subpath u v g)"
-    by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost)
-  show ?lhs using psp ne
-    unfolding arc_def subpath_def inj_on_def
-    by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2)
-qed
+  by (smt (verit, best) arc_simple_path closed_segment_commute ends_in_segment(2) inj_on_def pathfinish_subpath pathstart_subpath simple_path_subpath_eq)
 
 
 lemma simple_path_subpath:
@@ -917,7 +872,7 @@
   shows "simple_path(subpath u v g)"
   using assms
   apply (simp add: simple_path_subpath_eq simple_path_imp_path)
-  apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce)
+  apply (simp add: simple_path_def loop_free_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce)
   done
 
 lemma arc_simple_path_subpath:
@@ -930,7 +885,7 @@
 
 lemma arc_simple_path_subpath_interior:
     "\<lbrakk>simple_path g; u \<in> {0..1}; v \<in> {0..1}; u \<noteq> v; \<bar>u-v\<bar> < 1\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
-  by (force simp: simple_path_def intro: arc_simple_path_subpath)
+  by (force simp: simple_path_def loop_free_def intro: arc_simple_path_subpath)
 
 lemma path_image_subpath_subset:
     "\<lbrakk>u \<in> {0..1}; v \<in> {0..1}\<rbrakk> \<Longrightarrow> path_image(subpath u v g) \<subseteq> path_image g"
@@ -1065,17 +1020,8 @@
     assumes S: "closed S" and g: "path g" and g0: "pathstart g \<in> S" and g1: "pathfinish g \<notin> S"
     obtains h where "path h" "pathstart h = pathstart g" "path_image h \<subseteq> path_image g \<inter> S"
                     "pathfinish h \<in> frontier S"
-proof -
-  obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \<subseteq> path_image g"
-                    "path_image h - {pathfinish h} \<subseteq> interior S"
-                    "pathfinish h \<in> frontier S"
-    using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto
-  show ?thesis
-  proof
-    show "path_image h \<subseteq> path_image g \<inter> S"
-      using assms h interior_subset [of S] by (auto simp: frontier_def)
-  qed (use h in auto)
-qed
+  by (smt (verit, del_insts) Diff_iff Int_iff S closure_closed exists_path_subpath_to_frontier 
+      frontier_def g g0 g1 interior_subset singletonD subset_eq)
 
 
 subsection \<open>Shift Path to Start at Some Given Point\<close>
@@ -1121,8 +1067,7 @@
   have *: "{0 .. 1} = {0 .. 1-a} \<union> {1-a .. 1}"
     using assms(3) by auto
   have **: "\<And>x. x + a = 1 \<Longrightarrow> g (x + a - 1) = g (x + a)"
-    using assms(2)[unfolded pathfinish_def pathstart_def]
-    by auto
+    by (smt (verit, best) assms(2) pathfinish_def pathstart_def)
   show ?thesis
     unfolding path_def shiftpath_def *
   proof (rule continuous_on_closed_Un)
@@ -1175,20 +1120,21 @@
     by (auto simp: image_iff)
 qed
 
+lemma loop_free_shiftpath:
+  assumes "loop_free g" "pathfinish g = pathstart g" and a: "0 \<le> a" "a \<le> 1"
+    shows "loop_free (shiftpath a g)"
+  unfolding loop_free_def
+proof (intro conjI impI ballI)
+  show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
+    if "x \<in> {0..1}" "y \<in> {0..1}" "shiftpath a g x = shiftpath a g y" for x y
+    using that a assms unfolding shiftpath_def loop_free_def
+    by (smt (verit, ccfv_threshold) atLeastAtMost_iff)
+qed
+
 lemma simple_path_shiftpath:
   assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \<le> a" "a \<le> 1"
-    shows "simple_path (shiftpath a g)"
-  unfolding simple_path_def
-proof (intro conjI impI ballI)
-  show "path (shiftpath a g)"
-    by (simp add: assms path_shiftpath simple_path_imp_path)
-  have *: "\<And>x y. \<lbrakk>g x = g y; x \<in> {0..1}; y \<in> {0..1}\<rbrakk> \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
-    using assms by (simp add:  simple_path_def)
-  show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
-    if "x \<in> {0..1}" "y \<in> {0..1}" "shiftpath a g x = shiftpath a g y" for x y
-    using that a unfolding shiftpath_def
-    by (force split: if_split_asm dest!: *)
-qed
+  shows "simple_path (shiftpath a g)"
+  using assms loop_free_shiftpath path_shiftpath simple_path_def by fastforce
 
 
 subsection \<open>Straight-Line Paths\<close>
@@ -1281,14 +1227,7 @@
 
 lemma inj_on_linepath:
   assumes "a \<noteq> b" shows "inj_on (linepath a b) {0..1}"
-proof (clarsimp simp: inj_on_def linepath_def)
-  fix x y
-  assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1"
-  then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)"
-    by (auto simp: algebra_simps)
-  then show "x=y"
-    using assms by auto
-qed
+  using arc_imp_inj_on arc_linepath assms by blast
 
 lemma linepath_le_1:
   fixes a::"'a::linordered_idom" shows "\<lbrakk>a \<le> 1; b \<le> 1; 0 \<le> u; u \<le> 1\<rbrakk> \<Longrightarrow> (1 - u) * a + u * b \<le> 1"
@@ -1303,16 +1242,11 @@
 
 lemma linepath_in_convex_hull:
   fixes x::real
-  assumes a: "a \<in> convex hull S"
-    and b: "b \<in> convex hull S"
-    and x: "0\<le>x" "x\<le>1"
+  assumes "a \<in> convex hull S"
+    and "b \<in> convex hull S"
+    and "0\<le>x" "x\<le>1"
   shows "linepath a b x \<in> convex hull S"
-proof -
-  have "linepath a b x \<in> closed_segment a b"
-    using x by (auto simp flip: linepath_image_01)
-  then show ?thesis
-    using a b convex_contains_segment by blast
-qed
+  by (meson assms atLeastAtMost_iff convex_contains_segment convex_convex_hull linepath_in_path subset_eq)
 
 lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b"
   by (simp add: linepath_def)
@@ -1358,12 +1292,7 @@
 lemma midpoints_in_convex_hull:
   assumes "x \<in> convex hull s" "y \<in> convex hull s"
     shows "midpoint x y \<in> convex hull s"
-proof -
-  have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \<in> convex hull s"
-    by (rule convexD_alt) (use assms in auto)
-  then show ?thesis
-    by (simp add: midpoint_def algebra_simps)
-qed
+  using assms closed_segment_subset_convex_hull csegment_midpoint_subset by blast
 
 lemma not_in_interior_convex_hull_3:
   fixes a :: "complex"
@@ -1381,27 +1310,18 @@
 lemma continuous_IVT_local_extremum:
   fixes f :: "'a::euclidean_space \<Rightarrow> real"
   assumes contf: "continuous_on (closed_segment a b) f"
-      and "a \<noteq> b" "f a = f b"
+      and ab: "a \<noteq> b" "f a = f b"
   obtains z where "z \<in> open_segment a b"
                   "(\<forall>w \<in> closed_segment a b. (f w) \<le> (f z)) \<or>
                    (\<forall>w \<in> closed_segment a b. (f z) \<le> (f w))"
 proof -
   obtain c where "c \<in> closed_segment a b" and c: "\<And>y. y \<in> closed_segment a b \<Longrightarrow> f y \<le> f c"
     using continuous_attains_sup [of "closed_segment a b" f] contf by auto
+  moreover
   obtain d where "d \<in> closed_segment a b" and d: "\<And>y. y \<in> closed_segment a b \<Longrightarrow> f d \<le> f y"
     using continuous_attains_inf [of "closed_segment a b" f] contf by auto
-  show ?thesis
-  proof (cases "c \<in> open_segment a b \<or> d \<in> open_segment a b")
-    case True
-    then show ?thesis
-      using c d that by blast
-  next
-    case False
-    then have "(c = a \<or> c = b) \<and> (d = a \<or> d = b)"
-      by (simp add: \<open>c \<in> closed_segment a b\<close> \<open>d \<in> closed_segment a b\<close> open_segment_def)
-    with \<open>a \<noteq> b\<close> \<open>f a = f b\<close> c d show ?thesis
-      by (rule_tac z = "midpoint a b" in that) (fastforce+)
-  qed
+  ultimately show ?thesis
+    by (smt (verit) UnE ab closed_segment_eq_open empty_iff insert_iff midpoint_in_open_segment that)
 qed
 
 text\<open>An injective map into R is also an open map w.r.T. the universe, and conversely. \<close>
@@ -1483,14 +1403,8 @@
   assumes "path g"
     and "z \<notin> path_image g"
   shows "\<exists>e>0. cball z e \<inter> (path_image g) = {}"
-proof -
-  obtain e where "ball z e \<inter> path_image g = {}" "e > 0"
-    using not_on_path_ball[OF assms] by auto
-  moreover have "cball z (e/2) \<subseteq> ball z e"
-    using \<open>e > 0\<close> by auto
-  ultimately show ?thesis
-    by (rule_tac x="e/2" in exI) auto
-qed
+  by (smt (verit, ccfv_threshold) open_ball assms centre_in_ball inf.orderE inf_assoc
+      inf_bot_right not_on_path_ball open_contains_cball_eq)
 
 subsection \<open>Path component\<close>
 
@@ -1538,8 +1452,7 @@
 lemma path_component_linepath:
     fixes S :: "'a::real_normed_vector set"
     shows "closed_segment a b \<subseteq> S \<Longrightarrow> path_component S a b"
-  unfolding path_component_def
-  by (rule_tac x="linepath a b" in exI, auto)
+  unfolding path_component_def by fastforce
 
 subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Path components as sets\<close>
 
@@ -1605,7 +1518,7 @@
   assume as: "open e1" "open e2" "S \<subseteq> e1 \<union> e2" "e1 \<inter> e2 \<inter> S = {}" "e1 \<inter> S \<noteq> {}" "e2 \<inter> S \<noteq> {}"
   then obtain x1 x2 where obt:"x1 \<in> e1 \<inter> S" "x2 \<in> e2 \<inter> S"
     by auto
-  then obtain g where g: "path g" "path_image g \<subseteq> S" "pathstart g = x1" "pathfinish g = x2"
+  then obtain g where g: "path g" "path_image g \<subseteq> S" and pg: "pathstart g = x1" "pathfinish g = x2"
     using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto
   have *: "connected {0..1::real}"
     by (auto intro!: convex_connected)
@@ -1616,9 +1529,7 @@
     unfolding subset_eq
     by auto
   moreover have "{x \<in> {0..1}. g x \<in> e1} \<noteq> {} \<and> {x \<in> {0..1}. g x \<in> e2} \<noteq> {}"
-    using g(3,4)[unfolded path_defs]
-    using obt
-    by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl)
+    by (smt (verit, ccfv_threshold) IntE atLeastAtMost_iff empty_iff pg mem_Collect_eq obt pathfinish_def pathstart_def)
   ultimately show False
     using *[unfolded connected_local not_ex, rule_format,
       of "{0..1} \<inter> g -` e1" "{0..1} \<inter> g -` e2"]
@@ -1632,18 +1543,8 @@
   assumes "open S"
   shows "open (path_component_set S x)"
   unfolding open_contains_ball
-proof
-  fix y
-  assume as: "y \<in> path_component_set S x"
-  then have "y \<in> S"
-    by (simp add: path_component_mem(2))
-  then obtain e where e: "e > 0" "ball y e \<subseteq> S"
-    using assms openE by blast
-have "\<And>u. dist y u < e \<Longrightarrow> path_component S x u"
-      by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component)
-  then show "\<exists>e > 0. ball y e \<subseteq> path_component_set S x"
-    using \<open>e>0\<close> by auto
-qed
+  by (metis assms centre_in_ball convex_ball convex_imp_path_connected equals0D openE 
+      path_component_eq path_component_eq_empty path_component_maximal)
 
 lemma open_non_path_component:
   fixes S :: "'a::real_normed_vector set"
@@ -1660,14 +1561,8 @@
     show "\<And>x. x \<in> ball y e \<Longrightarrow> x \<in> S"
       using e by blast
     show False if "z \<in> ball y e" "z \<in> path_component_set S x" for z
-    proof -
-      have "y \<in> path_component_set S z"
-        by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1))
-      then have "y \<in> path_component_set S x"
-        using path_component_eq that(2) by blast
-      then show False
-        using y by blast
-    qed
+      by (metis (no_types, lifting) Diff_iff centre_in_ball convex_ball convex_imp_path_connected  
+          path_component_eq path_component_maximal subsetD that y e)
   qed (use e in auto)
 qed
 
@@ -1688,8 +1583,8 @@
       by auto
     ultimately
     show False
-      using \<open>y \<in> S\<close> open_non_path_component[OF assms(1)] open_path_component[OF assms(1)]
-      using assms(2)[unfolded connected_def not_ex, rule_format,
+      using \<open>y \<in> S\<close> open_non_path_component[OF \<open>open S\<close>] open_path_component[OF \<open>open S\<close>]
+      using \<open>connected S\<close>[unfolded connected_def not_ex, rule_format,
         of "path_component_set S x" "S - path_component_set S x"]
       by auto
   qed
@@ -1700,16 +1595,14 @@
     and "path_connected S"
   shows "path_connected (f ` S)"
   unfolding path_connected_def
-proof (rule, rule)
-  fix x' y'
-  assume "x' \<in> f ` S" "y' \<in> f ` S"
-  then obtain x y where x: "x \<in> S" and y: "y \<in> S" and x': "x' = f x" and y': "y' = f y"
-    by auto
-  from x y obtain g where "path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y"
-    using assms(2)[unfolded path_connected_def] by fast
-  then show "\<exists>g. path g \<and> path_image g \<subseteq> f ` S \<and> pathstart g = x' \<and> pathfinish g = y'"
-    unfolding x' y' path_defs
-    by (fastforce intro: continuous_on_compose continuous_on_subset[OF contf])
+proof clarsimp
+  fix x y 
+  assume x: "x \<in> S" and y: "y \<in> S" 
+  with \<open>path_connected S\<close> 
+  show "\<exists>g. path g \<and> path_image g \<subseteq> f ` S \<and> pathstart g = f x \<and> pathfinish g = f y"
+    unfolding path_defs path_connected_def
+    using continuous_on_subset[OF contf]
+    by (smt (verit, ccfv_threshold) continuous_on_compose2 image_eqI image_subset_iff)
 qed
 
 lemma path_connected_translationI:
@@ -1759,25 +1652,9 @@
   assume x: "x \<in> S \<union> T" and y: "y \<in> S \<union> T"
   from assms obtain z where z: "z \<in> S" "z \<in> T"
     by auto
-  show "path_component (S \<union> T) x y"
-    using x y
-  proof safe
-    assume "x \<in> S" "y \<in> S"
-    then show "path_component (S \<union> T) x y"
-      by (meson Un_upper1 \<open>path_connected S\<close> path_component_of_subset path_connected_component)
-  next
-    assume "x \<in> S" "y \<in> T"
-    then show "path_component (S \<union> T) x y"
-      by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component)
-  next
-  assume "x \<in> T" "y \<in> S"
-    then show "path_component (S \<union> T) x y"
-      by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component)
-  next
-    assume "x \<in> T" "y \<in> T"
-    then show "path_component (S \<union> T) x y"
-      by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute)
-  qed
+  with x y show "path_component (S \<union> T) x y"
+    by (smt (verit) assms(1,2) in_mono mem_Collect_eq path_component_eq path_component_maximal 
+        sup.bounded_iff sup.cobounded2 sup_ge1)
 qed
 
 lemma path_connected_UNION:
@@ -1791,7 +1668,7 @@
   then have "path_component (S i) x z" and "path_component (S j) z y"
     using assms by (simp_all add: path_connected_component)
   then have "path_component (\<Union>i\<in>A. S i) x z" and "path_component (\<Union>i\<in>A. S i) z y"
-    using *(1,3) by (auto elim!: path_component_of_subset [rotated])
+    using *(1,3) by (meson SUP_upper path_component_of_subset)+
   then show "path_component (\<Union>i\<in>A. S i) x y"
     by (rule path_component_trans)
 qed
@@ -1821,26 +1698,31 @@
   by (meson path_component_path_image_pathstart path_component_sym path_component_trans)
 
 lemma path_connected_path_component [simp]:
-   "path_connected (path_component_set s x)"
-proof -
-  { fix y z
-    assume pa: "path_component s x y" "path_component s x z"
-    then have pae: "path_component_set s x = path_component_set s y"
-      using path_component_eq by auto
-    have yz: "path_component s y z"
-      using pa path_component_sym path_component_trans by blast
-    then have "\<exists>g. path g \<and> path_image g \<subseteq> path_component_set s x \<and> pathstart g = y \<and> pathfinish g = z"
-      apply (simp add: path_component_def)
-      by (metis pae path_component_maximal path_connected_path_image pathstart_in_path_image)
-  }
-  then show ?thesis
-    by (simp add: path_connected_def)
+  "path_connected (path_component_set S x)"
+proof (clarsimp simp: path_connected_def)
+  fix y z
+  assume pa: "path_component S x y" "path_component S x z"
+  then have pae: "path_component_set S x = path_component_set S y"
+    using path_component_eq by auto
+  obtain g where g: "path g \<and> path_image g \<subseteq> S \<and> pathstart g = y \<and> pathfinish g = z"
+    using pa path_component_sym path_component_trans path_component_def by metis
+  then have "path_image g \<subseteq> path_component_set S x"
+    using pae path_component_maximal path_connected_path_image by blast
+  then show "\<exists>g. path g \<and> path_image g \<subseteq> path_component_set S x \<and>
+                 pathstart g = y \<and> pathfinish g = z"
+    using g by blast
 qed
 
-lemma path_component: "path_component S x y \<longleftrightarrow> (\<exists>t. path_connected t \<and> t \<subseteq> S \<and> x \<in> t \<and> y \<in> t)"
-  apply (intro iffI)
-  apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image)
-  using path_component_of_subset path_connected_component by blast
+lemma path_component: 
+  "path_component S x y \<longleftrightarrow> (\<exists>t. path_connected t \<and> t \<subseteq> S \<and> x \<in> t \<and> y \<in> t)"
+    (is "?lhs = ?rhs")
+proof 
+  assume ?lhs then show ?rhs
+    by (metis path_component_def path_connected_path_image pathfinish_in_path_image pathstart_in_path_image)
+next
+  assume ?rhs then show ?lhs
+    by (meson path_component_of_subset path_connected_component)
+qed
 
 lemma path_component_path_component [simp]:
    "path_component_set (path_component_set S x) x = path_component_set S x"
@@ -1976,9 +1858,7 @@
 
 lemma Union_path_component [simp]:
    "Union {path_component_set S x |x. x \<in> S} = S"
-apply (rule subset_antisym)
-using path_component_subset apply force
-using path_component_refl by auto
+  using path_component_subset path_component_refl by blast
 
 lemma path_component_disjoint:
    "disjnt (path_component_set S a) (path_component_set S b) \<longleftrightarrow>
@@ -2244,27 +2124,21 @@
 
 lemma path_components_of_maximal:
    "\<lbrakk>C \<in> path_components_of X; path_connectedin X S; ~disjnt C S\<rbrakk> \<Longrightarrow> S \<subseteq> C"
-  apply (auto simp: path_components_of_def path_component_of_equiv)
-  using path_component_of_maximal path_connectedin_def apply fastforce
-  by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal)
+  by (smt (verit, ccfv_SIG) disjnt_iff imageE mem_Collect_eq path_component_of_equiv 
+      path_component_of_maximal path_components_of_def)
 
 lemma pairwise_disjoint_path_components_of:
      "pairwise disjnt (path_components_of X)"
   by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv)
 
 lemma complement_path_components_of_Union:
-   "C \<in> path_components_of X
-        \<Longrightarrow> topspace X - C = \<Union>(path_components_of X - {C})"
-  by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of)
+   "C \<in> path_components_of X \<Longrightarrow> topspace X - C = \<Union>(path_components_of X - {C})"
+  by (metis Union_path_components_of bot.extremum ccpo_Sup_singleton diff_Union_pairwise_disjoint 
+        insert_subsetI pairwise_disjoint_path_components_of)
 
 lemma nonempty_path_components_of:
   assumes "C \<in> path_components_of X" shows "C \<noteq> {}"
-proof -
-  have "C \<in> path_component_of_set X ` topspace X"
-    using assms path_components_of_def by blast
-  then show ?thesis
-    using path_component_of_refl by fastforce
-qed
+  by (metis assms imageE path_component_of_eq_empty path_components_of_def)
 
 lemma path_components_of_subset: "C \<in> path_components_of X \<Longrightarrow> C \<subseteq> topspace X"
   by (auto simp: path_components_of_def path_component_of_equiv)
@@ -2324,18 +2198,8 @@
 next
   case False
   have "(path_components_of X = {S}) \<longleftrightarrow> (path_connected_space X \<and> topspace X = S)"
-  proof (intro iffI conjI)
-    assume L: "path_components_of X = {S}"
-    then show "path_connected_space X"
-      by (simp add: path_connected_space_iff_components_eq)
-    show "topspace X = S"
-      by (metis L ccpo_Sup_singleton [of S] Union_path_components_of)
-  next
-    assume R: "path_connected_space X \<and> topspace X = S"
-    then show "path_components_of X = {S}"
-      using ccpo_Sup_singleton [of S]
-      by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set)
-  qed
+    by (metis False Set.set_insert ex_in_conv insert_iff path_component_in_path_components_of 
+        path_connected_space_iff_components_eq path_connected_space_path_component_set)
   with False show ?thesis
     by (simp add: path_components_of_eq_empty subset_singleton_iff)
 qed
@@ -2367,22 +2231,7 @@
 lemma exists_path_component_of_superset:
   assumes S: "path_connectedin X S" and ne: "topspace X \<noteq> {}"
   obtains C where "C \<in> path_components_of X" "S \<subseteq> C"
-proof (cases "S = {}")
-  case True
-  then show ?thesis
-    using ne path_components_of_eq_empty that by fastforce
-next
-  case False
-  then obtain a where "a \<in> S"
-    by blast
-  show ?thesis
-  proof
-    show "Collect (path_component_of X a) \<in> path_components_of X"
-      by (meson \<open>a \<in> S\<close> S subsetD path_component_in_path_components_of path_connectedin_subset_topspace)
-    show "S \<subseteq> Collect (path_component_of X a)"
-      by (simp add: S \<open>a \<in> S\<close> path_component_of_maximal)
-  qed
-qed
+    by (metis S ne ex_in_conv path_component_in_path_components_of path_component_of_maximal path_component_of_subset_topspace subset_eq that)
 
 lemma path_component_of_eq_overlap:
    "path_component_of X x = path_component_of X y \<longleftrightarrow>
@@ -2444,8 +2293,8 @@
     have "Collect (path_component_of Y (f x)) \<subseteq> topspace Y"
       by (simp add: path_component_of_subset_topspace)
     moreover have "g ` Collect(path_component_of Y (f x)) \<subseteq> Collect (path_component_of X (g (f x)))"
-      using g x unfolding homeomorphic_maps_def
-      by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of)
+      using f g x unfolding homeomorphic_maps_def
+      by (metis image_Collect_subsetI image_eqI mem_Collect_eq path_component_of_equiv path_component_of_maximal path_connectedin_continuous_map_image path_connectedin_path_component_of)
     ultimately show "Collect (path_component_of Y (f x)) \<subseteq> f ` Collect (path_component_of X x)"
       using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff
       by metis
@@ -2460,7 +2309,6 @@
 lemma homeomorphic_map_path_components_of:
   assumes "homeomorphic_map X Y f"
   shows "path_components_of Y = (image f) ` (path_components_of X)"
-    (is "?lhs = ?rhs")
   unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric]
   using assms homeomorphic_map_path_component_of by fastforce
 
@@ -2497,8 +2345,7 @@
       by (simp add: inner_commute)
   qed
   obtain S :: "'a set" where "S \<subseteq> Basis" and "card S = Suc (Suc 0)"
-    using ex_card[OF assms]
-    by auto
+    using obtain_subset_with_card_n[OF assms] by (force simp add: eval_nat_numeral)
   then obtain b0 b1 :: 'a where "b0 \<in> Basis" and "b1 \<in> Basis" and "b0 \<noteq> b1"
     unfolding card_Suc_eq by auto
   then have "a + b0 - b1 \<in> ?A \<inter> ?B"
@@ -2528,11 +2375,11 @@
 proof (cases r "0::real" rule: linorder_cases)
   case less
   then show ?thesis
-    by (simp)
+    by simp
 next
   case equal
   then show ?thesis
-    by (simp)
+    by simp
 next
   case greater
   then have eq: "(sphere (0::'a) r) = (\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a})"
@@ -2713,10 +2560,13 @@
   case empty
   show ?case using \<open>connected S\<close> by simp
 next
-  case (insert x F)
-  then have "connected (S-F)" by auto
-  moreover have "open (S - F)" using finite_imp_closed[OF \<open>finite F\<close>] \<open>open S\<close> by auto
-  ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto
+  case (insert x T)
+  then have "connected (S-T)" 
+    by auto
+  moreover have "open (S - T)" 
+    using finite_imp_closed[OF \<open>finite T\<close>] \<open>open S\<close> by auto
+  ultimately have "connected (S - T - {x})" 
+    using connected_open_delete[OF _ _ 2] by auto
   thus ?case by (metis Diff_insert)
 qed
 
@@ -2756,13 +2606,8 @@
   fixes a :: "'a :: euclidean_space"
   assumes "DIM('a) = 1" and "r > 0"
   obtains x y where "sphere a r = {x,y} \<and> dist x y = 2*r"
-proof -
-  have "sphere a r = (+) a ` sphere 0 r"
-    by (metis add.right_neutral sphere_translation)
-  then show ?thesis
-    using sphere_1D_doubleton_zero [OF assms]
-    by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that)
-qed
+  using sphere_1D_doubleton_zero [OF assms] dist_add_cancel image_empty image_insert
+  by (metis (no_types, opaque_lifting) add.right_neutral sphere_translation)
 
 lemma psubset_sphere_Compl_connected:
   fixes S :: "'a::euclidean_space set"
@@ -2780,8 +2625,7 @@
   moreover have "connected {x. dist a x \<le> r \<and> x \<notin> S}"
     using assms
     by (force intro: connected_intermediate_closure [of "ball a r"])
-  moreover
-  have "connected {x. r \<le> dist a x \<and> x \<notin> S}"
+  moreover have "connected {x. r \<le> dist a x \<and> x \<notin> S}"
   proof (rule connected_intermediate_closure [of "- cball a r"])
     show "{x. r \<le> dist a x \<and> x \<notin> S} \<subseteq> closure (- cball a r)"
       using interior_closure by (force intro: connected_complement_bounded_convex)
@@ -2872,12 +2716,8 @@
   assumes x: "x \<in> S" and S: "open S"
   shows "x \<in> closure (connected_component_set S y) \<longleftrightarrow>  x \<in> connected_component_set S y"
 proof -
-  { assume "x \<in> closure (connected_component_set S y)"
-    moreover have "x \<in> connected_component_set S x"
-      using x by simp
-    ultimately have "x \<in> connected_component_set S y"
-      using S by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component)
-  }
+  have "x islimpt connected_component_set S y \<Longrightarrow> connected_component S y x"
+    by (metis (no_types, lifting) S connected_component_eq connected_component_refl islimptE mem_Collect_eq open_connected_component x)
   then show ?thesis
     by (auto simp: closure_def)
 qed
@@ -2903,10 +2743,8 @@
   from connectedD [OF \<open>connected S\<close> 1 2 4 3]
   have "S \<inter> \<Union>(B-{T}) = {}"
     by (auto simp: Int_commute \<open>S \<inter> T \<noteq> {}\<close>)
-  with \<open>T \<in> B\<close> have "S \<subseteq> T"
-    using "3" by auto
-  show ?thesis
-    using \<open>S \<inter> \<Union>(B - {T}) = {}\<close> \<open>S \<subseteq> T\<close> \<open>T \<in> B\<close> that by auto
+  with \<open>T \<in> B\<close> 3 that show ?thesis
+    by (metis IntI UnE empty_iff subsetD subsetI)
 qed
 
 lemma connected_disjoint_Union_open_subset:
@@ -2935,7 +2773,7 @@
       and SB: "\<And>S. S \<in> B \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
       and eq [simp]: "\<Union>A = \<Union>B"
     shows "A = B"
-by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms)
+by (metis subset_antisym connected_disjoint_Union_open_subset assms)
 
 proposition components_open_unique:
  fixes S :: "'a::real_normed_vector set"
@@ -2998,11 +2836,9 @@
   then have *: "\<And>x. B \<le> norm x \<Longrightarrow> x \<in> S"
     by (force simp: ball_def dist_norm)
   obtain x' where x': "connected_component S x x'" "norm x' > B"
-    using bo [unfolded bounded_def dist_norm, simplified, rule_format]
-    by (metis diff_zero norm_minus_commute not_less)
+    using B(1) bo(1) bounded_pos by force
   obtain y' where y': "connected_component S y y'" "norm y' > B"
-    using bo [unfolded bounded_def dist_norm, simplified, rule_format]
-    by (metis diff_zero norm_minus_commute not_less)
+    using B(1) bo(2) bounded_pos by force
   have x'y': "connected_component S x' y'"
     unfolding connected_component_def
   proof (intro exI conjI)
@@ -3010,11 +2846,8 @@
       using assms by (auto intro: connected_complement_bounded_convex)
   qed (use x' y' dist_norm * in auto)
   show ?thesis
-  proof (rule connected_component_eq)
-    show "x \<in> connected_component_set S y"
       using x' y' x'y'
-      by (metis (no_types) connected_component_eq_eq connected_component_in mem_Collect_eq)
-  qed
+      by (metis connected_component_eq mem_Collect_eq)
 qed
 
 lemma cobounded_unbounded_components:
@@ -3088,12 +2921,8 @@
   fixes u::real
   assumes "x \<ge> B" "y \<ge> B" "0 \<le> u" "u \<le> 1"
   shows "(1 - u) * x + u * y \<ge> B"
-proof -
-  obtain dx dy where "dx \<ge> 0" "dy \<ge> 0" "x = B + dx" "y = B + dy"
-    using assms by auto (metis add.commute diff_add_cancel)
-  with \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> show ?thesis
-    by (simp add: add_increasing2 mult_left_le field_simps)
-qed
+  by (smt (verit) assms convex_bound_le ge_iff_diff_ge_0 minus_add_distrib 
+      mult_minus_right neg_le_iff_le)
 
 lemma cobounded_outside:
   fixes S :: "'a :: real_normed_vector set"
@@ -3151,7 +2980,7 @@
 lemma connected_outside:
     fixes S :: "'a::euclidean_space set"
     assumes "bounded S" "2 \<le> DIM('a)"
-      shows "connected(outside S)"
+    shows "connected(outside S)"
   apply (clarsimp simp add: connected_iff_connected_component outside)
   apply (rule_tac S="connected_component_set (- S) x" in connected_component_of_subset)
   apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq)
@@ -3159,9 +2988,14 @@
 
 lemma outside_connected_component_lt:
   "outside S = {x. \<forall>B. \<exists>y. B < norm(y) \<and> connected_component (- S) x y}"
-  apply (auto simp: outside bounded_def dist_norm)
-   apply (metis diff_0 norm_minus_cancel not_less)
-  by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6))
+proof -
+  have "\<And>x B. x \<in> outside S \<Longrightarrow> \<exists>y. B < norm y \<and> connected_component (- S) x y"
+    by (metis boundedI linorder_not_less mem_Collect_eq outside)
+  moreover
+  have "\<And>x. \<forall>B. \<exists>y. B < norm y \<and> connected_component (- S) x y \<Longrightarrow> x \<in> outside S"
+    by (metis bounded_pos linorder_not_less mem_Collect_eq outside)
+  ultimately show ?thesis by auto
+qed
 
 lemma outside_connected_component_le:
   "outside S = {x. \<forall>B. \<exists>y. B \<le> norm(y) \<and> connected_component (- S) x y}"
@@ -3175,14 +3009,15 @@
 proof -
   obtain B::real where B: "0 < B" and Bno: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B"
     using S [simplified bounded_pos] by auto
-  { fix y::'a and z::'a
-    assume yz: "B < norm z" "B < norm y"
+  have cyz: "connected_component (- S) y z" 
+    if yz: "B < norm z" "B < norm y" for y::'a and z::'a
+  proof -
     have "connected_component (- cball 0 B) y z"
       using assms yz
       by (force simp: dist_norm intro: connected_componentI [OF _ subset_refl] connected_complement_bounded_convex)
-    then have "connected_component (- S) y z"
+    then show ?thesis
       by (metis connected_component_of_subset Bno Compl_anti_mono mem_cball_0 subset_iff)
-  } note cyz = this
+  qed
   show ?thesis
     apply (auto simp: outside bounded_pos)
     apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le)
@@ -3211,9 +3046,8 @@
 lemma inside_subset:
   assumes "connected U" and "\<not> bounded U" and "T \<union> U = - S"
   shows "inside S \<subseteq> T"
-  apply (auto simp: inside_def)
-  by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal
-      Compl_iff Un_iff assms subsetI)
+  using bounded_subset [of "connected_component_set (- S) _" U] assms
+  by (metis (no_types, lifting) ComplI Un_iff connected_component_maximal inside_def mem_Collect_eq subsetI)
 
 lemma frontier_not_empty:
   fixes S :: "'a :: real_normed_vector set"
@@ -3264,23 +3098,12 @@
       by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono)
   next
     case False
-    have 1: "closed_segment x y \<inter> T \<noteq> {}" 
-      using \<open>y \<in> T\<close> by blast
-    have 2: "closed_segment x y - T \<noteq> {}"
-      using False by blast
+    have \<section>: "closed_segment x y \<inter> T \<noteq> {}" "closed_segment x y - T \<noteq> {}"
+      using \<open>y \<in> T\<close> False by blast+
     obtain c where "c \<in> closed_segment x y" "c \<in> frontier T"
-       using False connected_Int_frontier [OF connected_segment 1 2] by auto
-    then show ?thesis
-    proof -
-      have "norm (y - x) < e"
-        by (metis dist_norm \<open>dist y x < e\<close>)
-      moreover have "norm (c - x) \<le> norm (y - x)"
-        by (simp add: \<open>c \<in> closed_segment x y\<close> segment_bound(1))
-      ultimately have "norm (c - x) < e"
-        by linarith
-      then show ?thesis
-        by (metis (no_types) \<open>c \<in> frontier T\<close> dist_norm that(1))
-    qed
+       using False connected_Int_frontier [OF connected_segment \<section>] by auto
+     with that show ?thesis
+       by (smt (verit) dist_norm segment_bound1)
   qed
   then show ?thesis
     by (fastforce simp add: frontier_def closure_approachable)
@@ -3289,8 +3112,7 @@
 lemma frontier_Union_subset:
   fixes F :: "'a::real_normed_vector set set"
   shows "finite F \<Longrightarrow> frontier(\<Union>F) \<subseteq> (\<Union>t \<in> F. frontier t)"
-by (rule order_trans [OF frontier_Union_subset_closure])
-   (auto simp: closure_subset_eq)
+  by (metis closed_UN closure_closed frontier_Union_subset_closure frontier_closed)
 
 lemma frontier_of_components_subset:
   fixes S :: "'a::real_normed_vector set"
@@ -3329,7 +3151,7 @@
     show "- frontier C \<subseteq> C \<union> - closure C"
       by (simp add: \<open>open C\<close> closed_Compl frontier_closures)
     then show "- closure C \<inter> - frontier C \<noteq> {}"
-      by (metis (no_types, lifting) C Compl_subset_Compl_iff \<open>frontier C \<subset> S\<close> compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb)
+      by (metis C Compl_Diff_eq Un_Int_eq(4) Un_commute \<open>frontier C \<subset> S\<close> \<open>open C\<close> compl_le_compl_iff frontier_def in_components_subset interior_eq leD sup_bot.right_neutral)
   qed
   then show False
     using \<open>connected (- frontier C)\<close> by blast
@@ -3368,10 +3190,8 @@
       using connected_component_in by auto
   }
   then show ?thesis
-    apply (auto simp: inside_def frontier_def)
-    apply (rule classical)
-    apply (rule bounded_subset [OF assms], blast)
-    done
+    using bounded_subset [OF assms]
+    by (metis (no_types, lifting) Diff_iff frontier_def inside_def mem_Collect_eq subsetI)
 qed
 
 lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)"
@@ -3483,17 +3303,7 @@
     fixes S :: "'a::real_normed_vector set"
     assumes "bounded S"
     shows  "outside(frontier S) \<subseteq> - closure S"
-  unfolding outside_inside boolean_algebra_class.compl_le_compl_iff
-proof -
-  { assume "interior S \<subseteq> inside (frontier S)"
-    hence "interior S \<union> inside (frontier S) = inside (frontier S)"
-      by (simp add: subset_Un_eq)
-    then have "closure S \<subseteq> frontier S \<union> inside (frontier S)"
-      using frontier_def by auto
-  }
-  then show "closure S \<subseteq> frontier S \<union> inside (frontier S)"
-    using interior_inside_frontier [OF assms] by blast
-qed
+  using assms frontier_def interior_inside_frontier outside_inside by fastforce
 
 lemma outside_frontier_eq_complement_closure:
   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
@@ -3557,18 +3367,7 @@
     fixes S :: "'a::real_normed_vector set"
     assumes "closed S"
       shows "frontier(inside S) \<subseteq> S"
-proof -
-  have "closure (inside S) \<inter> - inside S = closure (inside S) - interior (inside S)"
-    by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside)
-  moreover have "- inside S \<inter> - outside S = S"
-    by (metis (no_types) compl_sup double_compl inside_Un_outside)
-  moreover have "closure (inside S) \<subseteq> - outside S"
-    by (metis (no_types) assms closure_inside_subset union_with_inside)
-  ultimately have "closure (inside S) - interior (inside S) \<subseteq> S"
-    by blast
-  then show ?thesis
-    by (simp add: frontier_def open_inside interior_open)
-qed
+  using assms closure_inside_subset frontier_closures frontier_disjoint_eq open_inside by fastforce
 
 lemma closure_outside_subset:
     fixes S :: "'a::real_normed_vector set"
@@ -3624,10 +3423,8 @@
 
 lemma outside_bounded_nonempty:
   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
-    assumes "bounded S" shows "outside S \<noteq> {}"
-  by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel
-                   Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball
-                   double_complement order_refl outside_convex outside_def)
+  assumes "bounded S" shows "outside S \<noteq> {}"
+  using assms unbounded_outside by force
 
 lemma outside_compact_in_open:
     fixes S :: "'a :: {real_normed_vector,perfect_space} set"
@@ -3701,14 +3498,7 @@
     then obtain r where "0 < r" and r: "S \<union> T \<subseteq> ball 0 r"
       using bounded_subset_ballD by blast
     have outst: "outside S \<inter> outside T \<noteq> {}"
-    proof -
-      have "- ball 0 r \<subseteq> outside S"
-        by (meson convex_ball le_supE outside_subset_convex r)
-      moreover have "- ball 0 r \<subseteq> outside T"
-        by (meson convex_ball le_supE outside_subset_convex r)
-      ultimately show ?thesis
-        by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap)
-    qed
+      by (metis bounded_Un bounded_subset bst cobounded_outside disjoint_eq_subset_Compl unbounded_outside)
     have "S \<inter> T = {}" using assms
       by (metis disjoint_iff_not_equal inside_no_overlap subsetCE)
     moreover have "outside S \<inter> inside T \<noteq> {}"
@@ -3735,27 +3525,29 @@
     using that 
   proof
     assume "a \<in> S" then show ?thesis
-      by (rule_tac x=a in exI, rule_tac x="{a}" in exI, simp)
+      using cons by blast
   next
     assume a: "a \<in> inside S"
     then have ain: "a \<in> closure (inside S)"
       by (simp add: closure_def)
-    show ?thesis
-      apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside S"])
-        apply (simp_all add: ain b)
-      subgoal for h
-        apply (rule_tac x="pathfinish h" in exI)
-        apply (simp add: subsetD [OF frontier_inside_subset[OF S]])
-        apply (rule_tac x="path_image h" in exI)
-        apply (simp add: pathfinish_in_path_image connected_path_image, auto)
-        by (metis Diff_single_insert S frontier_inside_subset insert_iff interior_subset subsetD)
-      done
+    obtain h where h: "path h" "pathstart h = a" 
+                   "path_image h - {pathfinish h} \<subseteq> interior (inside S)"
+                   "pathfinish h \<in> frontier (inside S)"
+      using ain b
+      by (metis exists_path_subpath_to_frontier path_linepath pathfinish_linepath pathstart_linepath)
+    moreover
+    have h1S: "pathfinish h \<in> S"  
+      using S h frontier_inside_subset by blast
+    moreover 
+    have "path_image h \<subseteq> S \<union> inside S"
+      using IntD1 S h1S h interior_eq open_inside by fastforce
+    ultimately show ?thesis by blast
   qed
   show ?thesis
     apply (simp add: connected_iff_connected_component)
     apply (clarsimp simp add: connected_component_def dest!: *)
     subgoal for x y u u' T t'
-      by (rule_tac x="(S \<union> T \<union> t')" in exI) (auto intro!: connected_Un cons)
+      by (rule_tac x = "S \<union> T \<union> t'" in exI) (auto intro!: connected_Un cons)
     done
 qed
 
@@ -3777,16 +3569,19 @@
     assume a: "a \<in> outside S"
     then have ain: "a \<in> closure (outside S)"
       by (simp add: closure_def)
-    show ?thesis
-      apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside S"])
-        apply (simp_all add: ain b)
-      subgoal for h
-      apply (rule_tac x="pathfinish h" in exI)
-        apply (simp add: subsetD [OF frontier_outside_subset[OF S]])
-      apply (rule_tac x="path_image h" in exI)
-      apply (simp add: pathfinish_in_path_image connected_path_image, auto)
-        by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image S subsetCE)
-      done
+    obtain h where h: "path h" "pathstart h = a" 
+                   "path_image h - {pathfinish h} \<subseteq> interior (outside S)"
+                   "pathfinish h \<in> frontier (outside S)"
+      using ain b
+      by (metis exists_path_subpath_to_frontier path_linepath pathfinish_linepath pathstart_linepath)
+    moreover 
+    have h1S: "pathfinish h \<in> S"
+      using S frontier_outside_subset h(4) by blast
+    moreover 
+    have "path_image h \<subseteq> S \<union> outside S"
+      using IntD1 S h1S h interior_eq open_outside by fastforce
+    ultimately show ?thesis
+      by blast
   qed
   show ?thesis
     apply (simp add: connected_iff_connected_component)
@@ -3799,9 +3594,13 @@
 lemma inside_inside_eq_empty [simp]:
     fixes S :: "'a :: {real_normed_vector, perfect_space} set"
     assumes S: "closed S" and cons: "connected S"
-      shows "inside (inside S) = {}"
-  by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un
-           inside_complement_unbounded_connected_empty unbounded_outside union_with_outside)
+    shows "inside (inside S) = {}"
+proof -
+  have "connected (- inside S)"
+    by (metis S connected_with_outside cons union_with_outside)
+  then show ?thesis
+    by (metis bounded_Un inside_complement_unbounded_connected_empty unbounded_outside union_with_outside)
+qed
 
 lemma inside_in_components:
      "inside S \<in> components (- S) \<longleftrightarrow> connected(inside S) \<and> inside S \<noteq> {}" (is "?lhs = ?rhs")
@@ -3829,9 +3628,10 @@
 lemma bounded_unique_outside:
   fixes S :: "'a :: euclidean_space set"
   assumes "bounded S" "DIM('a) \<ge> 2"
-  shows "(c \<in> components (- S) \<and> \<not> bounded c \<longleftrightarrow> c = outside S)" 
+  shows "(c \<in> components (- S) \<and> \<not> bounded c) \<longleftrightarrow> c = outside S" 
   using assms
-  by (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside)
+  by (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty
+      outside_in_components unbounded_outside)
 
 
 subsection\<open>Condition for an open map's image to contain a ball\<close>
@@ -3869,21 +3669,20 @@
     then have "(\<xi> \<circ> K) \<longlonglongrightarrow> f y"
       by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially)
     with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto
-    have rle: "r \<le> norm (f y - f a)"
+    have "r \<le> norm (f y - f a)"
     proof (rule le_no)
       show "y \<in> frontier S"
         using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y)
     qed
-    have **: "(b \<inter> (- S) \<noteq> {} \<and> b - (- S) \<noteq> {} \<Longrightarrow> b \<inter> f \<noteq> {})
-              \<Longrightarrow> (b \<inter> S \<noteq> {}) \<Longrightarrow> b \<inter> f = {} \<Longrightarrow> b \<subseteq> S" 
-             for b f and S :: "'b set"
-      by blast
-    have \<section>: "\<And>y. \<lbrakk>norm (f a - y) < r; y \<in> frontier (f ` S)\<rbrakk> \<Longrightarrow> False"
-      by (metis dw_le norm_minus_commute not_less order_trans rle wy)
-    show ?thesis
-      apply (rule ** [OF connected_Int_frontier [where t = "f`S", OF connected_ball]])
-        (*such a horrible mess*)
-      using \<open>a \<in> S\<close> \<open>0 < r\<close> by (auto simp: disjoint_iff_not_equal dist_norm dest: \<section>)
+    then have "\<And>y. \<lbrakk>norm (f a - y) < r; y \<in> frontier (f ` S)\<rbrakk> \<Longrightarrow> False"
+      by (metis dw_le norm_minus_commute not_less order_trans wy)
+    then have "ball (f a) r \<inter> frontier (f ` S) = {}"
+      by (metis disjoint_iff_not_equal dist_norm mem_ball)
+    moreover
+    have "ball (f a) r \<inter> f ` S \<noteq> {}"
+      using \<open>a \<in> S\<close> \<open>0 < r\<close> centre_in_ball by blast
+    ultimately show ?thesis
+      by (meson connected_Int_frontier connected_ball diff_shunt_var)
 qed
 
 
@@ -3892,8 +3691,7 @@
 lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)"
 proof -
   have "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V"
-    if "x \<noteq> y"
-    for x y :: 'a
+    if "x \<noteq> y" for x y :: 'a
   proof (intro exI conjI)
     let ?r = "dist x y / 2"
     have [simp]: "?r > 0"
@@ -4140,7 +3938,6 @@
 lemma path_image_rectpath_cbox_minus_box:
   assumes "Re a \<le> Re b" "Im a \<le> Im b"
   shows   "path_image (rectpath a b) = cbox a b - box a b"
-  using assms by (auto simp: path_image_rectpath in_cbox_complex_iff
-                             in_box_complex_iff)
+  using assms by (auto simp: path_image_rectpath in_cbox_complex_iff in_box_complex_iff)
 
 end
--- a/src/HOL/Analysis/Tagged_Division.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Tagged_Division.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -300,15 +300,15 @@
       done
   }
   ultimately show ?l
-    unfolding division_of_def cbox_sing by auto
+    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_sing division_ofD(2) division_ofD(3) subset_singletonD that)
+    by (metis \<open>s division_of cbox a a\<close> cbox_idem division_ofD(2) division_ofD(3) subset_singletonD that)
   moreover have "s \<noteq> {}"
     using \<open>s division_of cbox a a\<close> by auto
   ultimately show ?r
-    unfolding cbox_sing by auto
+    unfolding cbox_idem by auto
 qed
 
 lemma elementary_empty: obtains p where "p division_of {}"
@@ -2316,7 +2316,7 @@
     "S \<subseteq> \<Union>\<D>"
 proof -
   have aibi: "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" and normab: "0 < norm(b - a)"
-    using \<open>box a b \<noteq> {}\<close> box_eq_empty box_sing by fastforce+
+    using \<open>box a b \<noteq> {}\<close> box_eq_empty box_idem by fastforce+
   let ?K0 = "\<lambda>(n, f::'a\<Rightarrow>nat).
                     cbox (\<Sum>i \<in> Basis. (a \<bullet> i + (f i / 2^n) * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)
                          (\<Sum>i \<in> Basis. (a \<bullet> i + ((f i + 1) / 2^n) * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)"
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -127,12 +127,10 @@
         by (force simp: SOME_Basis dist_norm)
     next
       case False
-      have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))"
-        by (simp add: algebra_simps)
-      also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))"
+      have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))"
         by (simp add: algebra_simps)
       also from \<open>a \<noteq> a'\<close> have "... = \<bar>- norm (a - a') - r\<bar>"
-        by simp (simp add: field_simps)
+        by (simp add: divide_simps)
       finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \<bar>norm (a - a') + r\<bar>"
         by linarith
       from \<open>a \<noteq> a'\<close> show ?thesis
@@ -167,8 +165,7 @@
       have False if "norm (a - a') + r \<ge> r'"
       proof -
         from that have "\<bar>r' - norm (a - a')\<bar> \<le> r"
-          by (simp split: abs_split)
-            (metis \<open>0 \<le> r\<close> \<open>?lhs\<close> centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq)
+          by (smt (verit, best) \<open>0 \<le> r\<close> \<open>?lhs\<close> ball_subset_cball cball_subset_cball_iff dist_norm order_trans)
         then show ?thesis
           using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \<open>?lhs\<close>] \<open>a \<noteq> a'\<close>
           apply (simp add: dist_norm)
@@ -222,8 +219,8 @@
     assume ?lhs
     then have "0 < r'"
       using False by metric
-    then have "(cball a r \<subseteq> cball a' r')"
-      by (metis False\<open>?lhs\<close> closure_ball closure_mono not_less)
+    then have "cball a r \<subseteq> cball a' r'"
+      by (metis False \<open>?lhs\<close> closure_ball closure_mono not_less)
     then show ?rhs
       using False cball_subset_cball_iff by fastforce
   qed metric
@@ -233,65 +230,17 @@
 lemma ball_eq_ball_iff:
   fixes x :: "'a :: euclidean_space"
   shows "ball x d = ball y e \<longleftrightarrow> d \<le> 0 \<and> e \<le> 0 \<or> x=y \<and> d=e"
-        (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs
-  proof (cases "d \<le> 0 \<or> e \<le> 0")
-    case True
-      with \<open>?lhs\<close> show ?rhs
-        by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric])
-  next
-    case False
-    with \<open>?lhs\<close> show ?rhs
-      apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps)
-      apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans)
-      apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero)
-      done
-  qed
-next
-  assume ?rhs then show ?lhs
-    by (auto simp: set_eq_subset ball_subset_ball_iff)
-qed
+  by (smt (verit, del_insts) ball_empty ball_subset_cball_iff dist_norm norm_pths(2))
 
 lemma cball_eq_cball_iff:
   fixes x :: "'a :: euclidean_space"
   shows "cball x d = cball y e \<longleftrightarrow> d < 0 \<and> e < 0 \<or> x=y \<and> d=e"
-        (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs
-  proof (cases "d < 0 \<or> e < 0")
-    case True
-      with \<open>?lhs\<close> show ?rhs
-        by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric])
-  next
-    case False
-    with \<open>?lhs\<close> show ?rhs
-      apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps)
-      apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans)
-      apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero)
-      done
-  qed
-next
-  assume ?rhs then show ?lhs
-    by (auto simp: set_eq_subset cball_subset_cball_iff)
-qed
+  by (smt (verit, ccfv_SIG) cball_empty cball_subset_cball_iff dist_norm norm_pths(2) zero_le_dist)
 
 lemma ball_eq_cball_iff:
   fixes x :: "'a :: euclidean_space"
   shows "ball x d = cball y e \<longleftrightarrow> d \<le> 0 \<and> e < 0" (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs
-    apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps)
-    apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist)
-    apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le)
-    using \<open>?lhs\<close> ball_eq_empty cball_eq_empty apply blast+
-    done
-next
-  assume ?rhs then show ?lhs by auto
-qed
+  by (smt (verit) ball_eq_empty ball_subset_cball_iff cball_eq_empty cball_subset_ball_iff order.refl)
 
 lemma cball_eq_ball_iff:
   fixes x :: "'a :: euclidean_space"
@@ -308,9 +257,8 @@
   obtain e2 where "0 < e2" and "\<forall>x\<in>X. x \<noteq> p \<longrightarrow> e2 \<le> dist p x"
     using finite_set_avoid[OF \<open>finite X\<close>,of p] by auto
   hence "\<forall>w\<in>ball p (min e1 e2). w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" using e1_b by auto
-  thus "\<exists>e>0. \<forall>w\<in>ball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> X)" using \<open>e2>0\<close> \<open>e1>0\<close>
-    apply (rule_tac x="min e1 e2" in exI)
-    by auto
+  thus "\<exists>e>0. \<forall>w\<in>ball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> X)" 
+    using \<open>e2>0\<close> \<open>e1>0\<close> by (rule_tac x="min e1 e2" in exI) auto
 qed
 
 lemma finite_cball_avoid:
@@ -391,9 +339,7 @@
   by (force simp: cbox_Pair_eq)
 
 lemma cbox_Complex_eq: "cbox (Complex a c) (Complex b d) = (\<lambda>(x,y). Complex x y) ` (cbox a b \<times> cbox c d)"
-  apply (auto simp: cbox_def Basis_complex_def)
-  apply (rule_tac x = "(Re x, Im x)" in image_eqI)
-  using complex_eq by auto
+  by (force simp: cbox_def Basis_complex_def)
 
 lemma cbox_Pair_eq_0: "cbox (a, c) (b, d) = {} \<longleftrightarrow> cbox a b = {} \<or> cbox c d = {}"
   by (force simp: cbox_Pair_eq)
@@ -425,22 +371,14 @@
   define e' where "e' = e / (2 * sqrt (real (DIM ('a))))"
   then have e: "e' > 0"
     using assms by (auto)
-  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i")
-  proof
-    fix i
-    from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e
-    show "?th i" by auto
-  qed
-  from choice[OF this] obtain a where
-    a: "\<forall>xa. a xa \<in> \<rat> \<and> a xa < x \<bullet> xa \<and> x \<bullet> xa - a xa < e'" ..
-  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i")
-  proof
-    fix i
-    from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e
-    show "?th i" by auto
-  qed
-  from choice[OF this] obtain b where
-    b: "\<forall>xa. b xa \<in> \<rat> \<and> x \<bullet> xa < b xa \<and> b xa - x \<bullet> xa < e'" ..
+  have "\<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" for i
+    using Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e by force
+  then obtain a where
+    a: "\<And>u. a u \<in> \<rat> \<and> a u < x \<bullet> u \<and> x \<bullet> u - a u < e'" by metis
+  have "\<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" for i
+    using Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e by force
+  then obtain b where
+    b: "\<And>u. b u \<in> \<rat> \<and> x \<bullet> u < b u \<and> b u - x \<bullet> u < e'" by metis
   let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i"
   show ?thesis
   proof (rule exI[of _ ?a], rule exI[of _ ?b], safe)
@@ -454,10 +392,8 @@
       assume i: "i \<in> Basis"
       have "a i < y\<bullet>i \<and> y\<bullet>i < b i"
         using * i by (auto simp: box_def)
-      moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'"
-        using a by auto
-      moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'"
-        using b by auto
+      moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" "x\<bullet>i < b i" "b i - x\<bullet>i < e'"
+        using a b by auto
       ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'"
         by auto
       then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
@@ -471,7 +407,7 @@
       using \<open>0 < e\<close> by simp
     finally show "y \<in> ball x e"
       by (auto simp: ball_def)
-  qed (insert a b, auto simp: box_def)
+  qed (use a b in \<open>auto simp: box_def\<close>)
 qed
 
 lemma open_UNION_box:
@@ -527,22 +463,14 @@
   define e' where "e' = e / (2 * sqrt (real (DIM ('a))))"
   then have e: "e' > 0"
     using assms by auto
-  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i")
-  proof
-    fix i
-    from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e
-    show "?th i" by auto
-  qed
-  from choice[OF this] obtain a where
-    a: "\<forall>u. a u \<in> \<rat> \<and> a u < x \<bullet> u \<and> x \<bullet> u - a u < e'" ..
-  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i")
-  proof
-    fix i
-    from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e
-    show "?th i" by auto
-  qed
-  from choice[OF this] obtain b where
-    b: "\<forall>u. b u \<in> \<rat> \<and> x \<bullet> u < b u \<and> b u - x \<bullet> u < e'" ..
+  have "\<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" for i
+    using Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e by force
+  then obtain a where
+    a: "\<forall>u. a u \<in> \<rat> \<and> a u < x \<bullet> u \<and> x \<bullet> u - a u < e'" by metis
+  have "\<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" for i
+    using Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e by force
+  then obtain b where
+    b: "\<forall>u. b u \<in> \<rat> \<and> x \<bullet> u < b u \<and> b u - x \<bullet> u < e'" by metis
   let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i"
   show ?thesis
   proof (rule exI[of _ ?a], rule exI[of _ ?b], safe)
@@ -556,10 +484,8 @@
       assume i: "i \<in> Basis"
       have "a i \<le> y\<bullet>i \<and> y\<bullet>i \<le> b i"
         using * i by (auto simp: cbox_def)
-      moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'"
-        using a by auto
-      moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'"
-        using b by auto
+      moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" "x\<bullet>i < b i" "b i - x\<bullet>i < e'"
+        using a b by auto
       ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'"
         by auto
       then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
@@ -626,54 +552,28 @@
   shows "(box a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i))" (is ?th1)
     and "(cbox a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i < a\<bullet>i))" (is ?th2)
 proof -
-  {
-    fix i x
-    assume i: "i\<in>Basis" and as:"b\<bullet>i \<le> a\<bullet>i" and x:"x\<in>box a b"
-    then have "a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i"
-      unfolding mem_box by (auto simp: box_def)
-    then have "a\<bullet>i < b\<bullet>i" by auto
-    then have False using as by auto
-  }
+  have False if "i \<in> Basis" and "b\<bullet>i \<le> a\<bullet>i" and "x \<in> box a b" for i x
+    by (smt (verit, ccfv_SIG) mem_box(1) that)
   moreover
-  {
-    assume as: "\<forall>i\<in>Basis. \<not> (b\<bullet>i \<le> a\<bullet>i)"
+  { assume as: "\<forall>i\<in>Basis. \<not> (b\<bullet>i \<le> a\<bullet>i)"
     let ?x = "(1/2) *\<^sub>R (a + b)"
-    {
-      fix i :: 'a
+    { fix i :: 'a
       assume i: "i \<in> Basis"
       have "a\<bullet>i < b\<bullet>i"
-        using as[THEN bspec[where x=i]] i by auto
+        using as i by fastforce
       then have "a\<bullet>i < ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i < b\<bullet>i"
         by (auto simp: inner_add_left)
     }
     then have "box a b \<noteq> {}"
-      using mem_box(1)[of "?x" a b] by auto
+      by (metis (no_types, opaque_lifting) emptyE mem_box(1))
   }
   ultimately show ?th1 by blast
 
-  {
-    fix i x
-    assume i: "i \<in> Basis" and as:"b\<bullet>i < a\<bullet>i" and x:"x\<in>cbox a b"
-    then have "a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i"
-      unfolding mem_box by auto
-    then have "a\<bullet>i \<le> b\<bullet>i" by auto
-    then have False using as by auto
-  }
+  have False if "i\<in>Basis" and "b\<bullet>i < a\<bullet>i" and "x \<in> cbox a b" for i x
+    using mem_box(2) that by force
   moreover
-  {
-    assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)"
-    let ?x = "(1/2) *\<^sub>R (a + b)"
-    {
-      fix i :: 'a
-      assume i:"i \<in> Basis"
-      have "a\<bullet>i \<le> b\<bullet>i"
-        using as[THEN bspec[where x=i]] i by auto
-      then have "a\<bullet>i \<le> ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i \<le> b\<bullet>i"
-        by (auto simp: inner_add_left)
-    }
-    then have "cbox a b \<noteq> {}"
-      using mem_box(2)[of "?x" a b] by auto
-  }
+  have "cbox a b \<noteq> {}" if "\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)"
+    by (metis emptyE linorder_linear mem_box(2) order.strict_iff_not that)
   ultimately show ?th2 by blast
 qed
 
@@ -685,11 +585,9 @@
 
 lemma
   fixes a :: "'a::euclidean_space"
-  shows cbox_sing [simp]: "cbox a a = {a}"
-    and box_sing [simp]: "box a a = {}"
-  unfolding set_eq_iff mem_box eq_iff [symmetric]
-  by (auto intro!: euclidean_eqI[where 'a='a])
-     (metis all_not_in_conv nonempty_Basis)
+  shows cbox_idem [simp]: "cbox a a = {a}"
+    and box_idem [simp]: "box a a = {}"
+  unfolding set_eq_iff mem_box eq_iff [symmetric] using euclidean_eq_iff by fastforce+
 
 lemma subset_box_imp:
   fixes a :: "'a::euclidean_space"
@@ -765,11 +663,7 @@
     by auto
   then show ?rhs
     by (force simp: subset_box box_eq_empty intro: antisym euclidean_eqI)
-next
-  assume ?rhs
-  then show ?lhs
-    by force
-qed
+qed auto
 
 lemma eq_cbox_box [simp]: "cbox a b = box c d \<longleftrightarrow> cbox a b = {} \<and> box c d = {}"
   (is "?lhs \<longleftrightarrow> ?rhs")
@@ -777,10 +671,8 @@
   assume L: ?lhs
   then have "cbox a b \<subseteq> box c d" "box c d \<subseteq> cbox a b"
     by auto
-  then show ?rhs
-    apply (simp add: subset_box)
-    using L box_ne_empty box_sing apply (fastforce simp add:)
-    done
+  with L subset_box show ?rhs
+    by (smt (verit) SOME_Basis box_ne_empty(1))
 qed force
 
 lemma eq_box_cbox [simp]: "box a b = cbox c d \<longleftrightarrow> box a b = {} \<and> cbox c d = {}"
@@ -793,11 +685,7 @@
   then have "box a b \<subseteq> box c d" "box c d \<subseteq> box a b"
     by auto
   then show ?rhs
-    apply (simp add: subset_box)
-    using box_ne_empty(2) L
-    apply auto
-     apply (meson euclidean_eqI less_eq_real_def not_less)+
-    done
+    unfolding subset_box by (smt (verit) box_ne_empty(2) euclidean_eq_iff)+
 qed force
 
 lemma subset_box_complex:
@@ -895,27 +783,30 @@
       by (simp add: dual_order.antisym euclidean_eqI)
   }
   moreover have "c \<in> cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)"
-    unfolding True by (auto)
+    unfolding True by auto
   ultimately show ?thesis using True by (auto simp: cbox_def)
 next
   case False
   {
     fix y
     assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m > 0"
-    then have "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i"
+    then have "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" 
+          and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i"
       by (auto simp: inner_distrib)
   }
   moreover
   {
     fix y
     assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m < 0"
-    then have "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i"
+    then have "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i"
+         and  "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i"
       by (auto simp: mult_left_mono_neg inner_distrib)
   }
   moreover
   {
     fix y
-    assume "m > 0" and "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> y \<bullet> i" and "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i"
+    assume "m > 0" and "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> y \<bullet> i"
+      and  "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i"
     then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b"
       unfolding image_iff Bex_def mem_box
       apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
@@ -946,11 +837,7 @@
   have "(\<lambda>(x, y). f y x) = (\<lambda>(x, y). f x y) \<circ> prod.swap"
     by auto
   then show ?thesis
-    apply (rule ssubst)
-    apply (rule continuous_on_compose)
-    apply (simp add: split_def)
-    apply (rule continuous_intros | simp add: assms)+
-    done
+    by (metis assms continuous_on_compose continuous_on_swap swap_cbox_Pair)
 qed
 
 
@@ -979,11 +866,11 @@
   unfolding is_interval_def  by simp
 
 lemma mem_is_intervalI:
-  assumes "is_interval s"
-    and "a \<in> s" "b \<in> s"
+  assumes "is_interval S"
+    and "a \<in> S" "b \<in> S"
     and "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i \<or> b \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> a \<bullet> i"
-  shows "x \<in> s"
-  by (rule assms(1)[simplified is_interval_def, rule_format, OF assms(2,3,4)])
+  shows "x \<in> S"
+  using assms is_interval_def by force
 
 lemma interval_subst:
   fixes S::"'a::euclidean_space set"
@@ -1016,11 +903,15 @@
     by (induct bs) (auto simp: y_def euclidean_representation intro!: euclidean_eqI[where 'a='a])
   also have "y bs \<in> S"
     using bs(1)[THEN equalityD1]
-    apply (induct bs)
-     apply (auto simp: y_def j)
-    apply (rule interval_subst[OF assms(1)])
-      apply (auto simp: s)
-    done
+  proof (induction bs)
+    case Nil
+    then show ?case
+      by (simp add: j y_def)
+  next
+    case (Cons a bs)
+    then show ?case
+      using interval_subst[OF assms(1)] s by (simp add: y_def)
+  qed
   finally show ?thesis .
 qed
 
@@ -1042,9 +933,8 @@
 next
   assume ?rhs
   have "cbox a b \<subseteq> S" if "a \<in> S" "b \<in> S"
-    using assms unfolding is_interval_def
-    apply (clarsimp simp add: mem_box)
-    using that by blast
+    using assms that 
+    by (force simp: mem_box intro: mem_is_intervalI)
   with \<open>?rhs\<close> show ?lhs
     by blast
 qed
@@ -1091,8 +981,7 @@
     "\<forall>i\<in>Basis. (- b) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (- d) \<bullet> i \<or>
        (- d) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (- b) \<bullet> i"
   hence "- e \<in> X"
-    by (intro mem_is_intervalI[OF assms \<open>b \<in> X\<close> \<open>d \<in> X\<close>, of "- e"])
-      (auto simp: algebra_simps)
+    by (smt (verit, ccfv_threshold) assms inner_minus_left mem_is_intervalI)
   thus "e \<in> uminus ` X" by force
 qed
 
@@ -1266,7 +1155,7 @@
     shows "closure {x. a \<bullet> x < b} = {x. a \<bullet> x \<le> b}"
 proof -
   have [simp]: "-{x. a \<bullet> x < b} = {x. a \<bullet> x \<ge> b}"
-    by (force simp:)
+    by force
   then show ?thesis
     using interior_halfspace_ge [of a b] assms
     by (force simp: closure_interior)
@@ -1281,7 +1170,7 @@
     shows "interior {x. a \<bullet> x = b} = {}"
 proof -
   have [simp]: "{x. a \<bullet> x = b} = {x. a \<bullet> x \<le> b} \<inter> {x. a \<bullet> x \<ge> b}"
-    by (force simp:)
+    by force
   then show ?thesis
     by (auto simp: assms)
 qed
@@ -1340,9 +1229,8 @@
   moreover have "?A \<inter> ?B = {}" by auto
   moreover have "S \<subseteq> ?A \<union> ?B" using as by auto
   ultimately show False
-    using \<open>connected S\<close>[unfolded connected_def not_ex,
-      THEN spec[where x="?A"], THEN spec[where x="?B"]]
-    using xy b by auto
+    using \<open>connected S\<close> unfolding connected_def
+    by (smt (verit, del_insts) as b disjoint_iff empty_iff mem_Collect_eq xy)
 qed
 
 lemma connected_ivt_component:
@@ -1419,18 +1307,11 @@
       assume i: "i \<in> Basis"
       have "dist (x - (e / 2) *\<^sub>R i) x < e"
         and "dist (x + (e / 2) *\<^sub>R i) x < e"
-        unfolding dist_norm
-        apply auto
-        unfolding norm_minus_cancel
-        using norm_Basis[OF i] \<open>e>0\<close>
-        apply auto
-        done
+         using norm_Basis[OF i] \<open>e>0\<close> by (auto simp: dist_norm)
       then have "a \<bullet> i \<le> (x - (e / 2) *\<^sub>R i) \<bullet> i" and "(x + (e / 2) *\<^sub>R i) \<bullet> i \<le> b \<bullet> i"
         using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]]
           and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]]
-        unfolding mem_box
-        using i
-        by blast+
+        unfolding mem_box using i by blast+
       then have "a \<bullet> i < x \<bullet> i" and "x \<bullet> i < b \<bullet> i"
         using \<open>e>0\<close> i
         by (auto simp: inner_diff_left inner_Basis inner_add_left)
@@ -1461,8 +1342,7 @@
 lemma bounded_box [simp]:
   fixes a :: "'a::euclidean_space"
   shows "bounded (box a b)"
-  using bounded_cbox[of a b] box_subset_cbox[of a b] bounded_subset[of "cbox a b" "box a b"]
-  by simp
+  by (metis bounded_cbox bounded_interior interior_cbox)
 
 lemma not_interval_UNIV [simp]:
   fixes a :: "'a::euclidean_space"
@@ -1497,12 +1377,7 @@
     have "a \<bullet> i = e * (a \<bullet> i) + (1 - e) * (a \<bullet> i)"
       unfolding left_diff_distrib by simp
     also have "\<dots> < e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)"
-    proof (rule add_less_le_mono)
-      show "e * (a \<bullet> i) < e * (x \<bullet> i)"
-        using \<open>0 < e\<close> i mem_box(1) x by auto
-      show "(1 - e) * (a \<bullet> i) \<le> (1 - e) * (y \<bullet> i)"
-        by (meson diff_ge_0_iff_ge \<open>e \<le> 1\<close> i mem_box(2) mult_left_mono y)
-    qed
+      by (smt (verit, best) e i mem_box mult_le_cancel_left_pos mult_left_mono x y)
     finally have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i"
       unfolding inner_simps by auto
     moreover
@@ -1510,12 +1385,7 @@
       have "b \<bullet> i = e * (b\<bullet>i) + (1 - e) * (b\<bullet>i)"
         unfolding left_diff_distrib by simp
       also have "\<dots> > e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)"
-      proof (rule add_less_le_mono)
-        show "e * (x \<bullet> i) < e * (b \<bullet> i)"
-          using \<open>0 < e\<close> i mem_box(1) x by auto
-        show "(1 - e) * (y \<bullet> i) \<le> (1 - e) * (b \<bullet> i)"
-          by (meson diff_ge_0_iff_ge \<open>e \<le> 1\<close> i mem_box(2) mult_left_mono y)
-      qed
+        by (smt (verit, best) e i mem_box mult_le_cancel_left_pos mult_left_mono x y)
       finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i"
         unfolding inner_simps by auto
     }
@@ -1539,7 +1409,7 @@
   let ?c = "(1 / 2) *\<^sub>R (a + b)"
   {
     fix x
-    assume as:"x \<in> cbox a b"
+    assume as: "x \<in> cbox a b"
     define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n
     {
       fix n
@@ -1557,22 +1427,16 @@
     }
     moreover
     {
-      assume "\<not> (f \<longlongrightarrow> x) sequentially"
-      {
-        fix e :: real
-        assume "e > 0"
-        then obtain N :: nat where N: "inverse (real (N + 1)) < e"
-          using reals_Archimedean by auto
-        have "inverse (real n + 1) < e" if "N \<le> n" for n
-          by (auto intro!: that le_less_trans [OF _ N])
-        then have "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto
-      }
-      then have "((\<lambda>n. inverse (real n + 1)) \<longlongrightarrow> 0) sequentially"
+      have "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < \<epsilon>" if "\<epsilon> > 0" for \<epsilon>
+          using reals_Archimedean [of \<epsilon>] that
+          by (metis inverse_inverse_eq inverse_less_imp_less nat_le_real_less order_less_trans 
+                  reals_Archimedean2)
+      then have "(\<lambda>n. inverse (real n + 1)) \<longlonglongrightarrow> 0"
         unfolding lim_sequentially by(auto simp: dist_norm)
-      then have "(f \<longlongrightarrow> x) sequentially"
+      then have "f \<longlonglongrightarrow> x"
         unfolding f_def
-        using tendsto_add[OF tendsto_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x]
-        using tendsto_scaleR [OF _ tendsto_const, of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"]
+        using tendsto_add[OF tendsto_const, of "\<lambda>n. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x]
+        using tendsto_scaleR [OF _ tendsto_const, of "\<lambda>n. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"]
         by auto
     }
     ultimately have "x \<in> closure (box a b)"
@@ -1603,12 +1467,7 @@
   fixes S :: "('a::euclidean_space) set"
   assumes "bounded S"
   obtains a where "S \<subseteq> cbox (-a) a"
-proof -
-  obtain a where "S \<subseteq> box (-a) a"
-    using bounded_subset_box_symmetric[OF assms] by auto
-  then show ?thesis
-    by (meson box_subset_cbox dual_order.trans that)
-qed
+  by (meson assms bounded_subset_box_symmetric box_subset_cbox order.trans)
 
 lemma frontier_cbox:
   fixes a b :: "'a::euclidean_space"
@@ -1618,16 +1477,7 @@
 lemma frontier_box:
   fixes a b :: "'a::euclidean_space"
   shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)"
-proof (cases "box a b = {}")
-  case True
-  then show ?thesis
-    using frontier_empty by auto
-next
-  case False
-  then show ?thesis
-    unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box]
-    by auto
-qed
+  by (simp add: frontier_def interior_open open_box)
 
 lemma Int_interval_mixed_eq_empty:
   fixes a :: "'a::euclidean_space"
@@ -1661,28 +1511,21 @@
     with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))) sequentially"
       by simp
     moreover
-    {
-      fix n
+    { fix n
       assume n: "\<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))"
       have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))"
-        apply (subst euclidean_dist_l2)
-        using zero_le_dist
-        apply (rule L2_set_le_sum)
-        done
+        using L2_set_le_sum [OF zero_le_dist] by (subst euclidean_dist_l2)
       also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))"
-        apply (rule sum_strict_mono)
-        using n
-        apply auto
-        done
+        by (meson eucl.finite_Basis n nonempty_Basis sum_strict_mono)
       finally have "dist (f (r n)) l < e"
         by auto
     }
-    ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
+    ultimately have "\<forall>\<^sub>F n in sequentially. dist (f (r n)) l < e"
       by (rule eventually_mono)
   }
-  then have *: "((f \<circ> r) \<longlongrightarrow> l) sequentially"
+  then have *: "(f \<circ> r) \<longlonglongrightarrow> l"
     unfolding o_def tendsto_iff by simp
-  with r show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
+  with r show "\<exists>l r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
     by auto
 qed
 
@@ -1704,10 +1547,8 @@
     fix A::"'a set"
     assume "open A"
     show "\<exists>B'\<subseteq>B. \<Union>B' = A"
-      apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"])
-      apply (subst (3) open_UNION_box[OF \<open>open A\<close>])
-      apply (auto simp: a b B_def)
-      done
+      using open_UNION_box[OF \<open>open A\<close>]
+      by (smt (verit, ccfv_threshold) B_def a b image_iff mem_Collect_eq subsetI)
   qed
   ultimately
   have "topological_basis B"
@@ -1770,10 +1611,11 @@
         by (simp add: False b cSUP_least)
       finally have bi: "x \<bullet> i \<le> b \<bullet> i" .
       show "x \<bullet> i \<in> (\<lambda>x. x \<bullet> i) ` S"
-        apply (rule_tac x="\<Sum>j\<in>Basis. (if j = i then x \<bullet> i else a \<bullet> j) *\<^sub>R j" in image_eqI)
+        apply (rule_tac x="\<Sum>j\<in>Basis. (((\<bullet>)a)(i := x \<bullet> j))j *\<^sub>R j" in image_eqI)
         apply (simp add: i)
         apply (rule mem_is_intervalI [OF \<open>is_interval S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close>])
-        using i ai bi apply force
+        using i ai bi 
+        apply force
         done
     qed
     have "S = cbox a b"
@@ -1826,10 +1668,7 @@
   assume ?lhs
   then show ?rhs
     unfolding tendsto_def
-    apply clarify
-    apply (drule_tac x="{s. s \<bullet> i \<in> S}" in spec)
-    apply (auto simp: open_preimage_inner)
-    done
+    by (smt (verit) eventually_elim2 mem_Collect_eq open_preimage_inner)
 next
   assume R: ?rhs
   then have "\<And>e. e > 0 \<Longrightarrow> \<forall>i\<in>Basis. \<forall>\<^sub>F x in F. dist (f x \<bullet> i) (l \<bullet> i) < e"
@@ -1847,19 +1686,15 @@
       have "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis \<le> sum (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis"
         by (simp add: L2_set_le_sum)
       also have "... < DIM('b) * (e / real DIM('b))"
-        apply (rule sum_bounded_above_strict)
-        using that by auto
+        by (meson DIM_positive sum_bounded_above_strict that)
       also have "... = e"
         by (simp add: field_simps)
       finally show "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis < e" .
     qed
     have "\<forall>\<^sub>F x in F. \<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e / DIM('b)"
-      apply (rule R')
-      using \<open>0 < e\<close> by simp
+      by (simp add: R' \<open>0 < e\<close>)
     then show "\<forall>\<^sub>F x in F. dist (f x) l < e"
-      apply (rule eventually_mono)
-      apply (subst euclidean_dist_l2)
-      using * by blast
+      by eventually_elim (metis (full_types) "*" euclidean_dist_l2)
   qed
 qed
 
@@ -1871,22 +1706,21 @@
 corollary continuous_on_componentwise:
   fixes S :: "'a :: t2_space set"
   shows "continuous_on S f \<longleftrightarrow> (\<forall>i \<in> Basis. continuous_on S (\<lambda>x. (f x \<bullet> i)))"
-  apply (simp add: continuous_on_eq_continuous_within)
-  using continuous_componentwise by blast
+  by (metis continuous_componentwise continuous_on_eq_continuous_within)
 
 lemma linear_componentwise_iff:
-     "(linear f') \<longleftrightarrow> (\<forall>i\<in>Basis. linear (\<lambda>x. f' x \<bullet> i))"
-  apply (auto simp: linear_iff inner_left_distrib)
-   apply (metis inner_left_distrib euclidean_eq_iff)
-  by (metis euclidean_eqI inner_scaleR_left)
+     "linear f' \<longleftrightarrow> (\<forall>i\<in>Basis. linear (\<lambda>x. f' x \<bullet> i))" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (simp add: Real_Vector_Spaces.linear_iff inner_left_distrib)
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (simp add: linear_iff) (metis euclidean_eqI inner_left_distrib inner_scaleR_left)
+qed
 
 lemma bounded_linear_componentwise_iff:
      "(bounded_linear f') \<longleftrightarrow> (\<forall>i\<in>Basis. bounded_linear (\<lambda>x. f' x \<bullet> i))"
      (is "?lhs = ?rhs")
 proof
-  assume ?lhs then show ?rhs
-    by (simp add: bounded_linear_inner_left_comp)
-next
   assume ?rhs
   then have "(\<forall>i\<in>Basis. \<exists>K. \<forall>x. \<bar>f' x \<bullet> i\<bar> \<le> norm x * K)" "linear f'"
     by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib)
@@ -1904,7 +1738,7 @@
   qed
   then show ?lhs
     by (force simp: bounded_linear_def bounded_linear_axioms_def \<open>linear f'\<close>)
-qed
+qed (simp add: bounded_linear_inner_left_comp)
 
 subsection\<^marker>\<open>tag unimportant\<close> \<open>Continuous Extension\<close>
 
@@ -1965,8 +1799,7 @@
     obtain d where d: "0 < d"
       "\<And>x'. x' \<in> cbox a b \<Longrightarrow> dist x' (clamp a b x) < d \<Longrightarrow> dist (f x') (f (clamp a b x)) < e"
       by force
-    show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow>
-      dist (f (clamp a b x')) (f (clamp a b x)) < e"
+    show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow> dist (f (clamp a b x')) (f (clamp a b x)) < e"
       using le
       by (auto intro!: d clamp_in_interval dist_clamps_le_dist_args[THEN le_less_trans])
   qed
@@ -1988,8 +1821,7 @@
   from bounded obtain c where f_bound: "\<forall>x\<in>f ` cbox a b. dist undefined x \<le> c"
     by (auto simp: bounded_any_center[where a=undefined])
   then show ?thesis
-    by (auto intro!: exI[where x=c] clamp_in_interval[OF le[rule_format]]
-        simp: bounded_any_center[where a=undefined])
+    by (metis bounded bounded_subset clamp_in_interval image_mono image_subsetI le range_composition)
 qed (auto simp: clamp_empty_interval image_def)
 
 
@@ -2000,9 +1832,7 @@
   fixes x a b :: "'a::euclidean_space"
   assumes x: "x \<in> cbox a b"
   shows "ext_cont f a b x = f x"
-  using assms
-  unfolding ext_cont_def
-  by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a] arg_cong[where f=f])
+  using assms by (simp add: ext_cont_def)
 
 lemma continuous_on_ext_cont[continuous_intros]:
   "continuous_on (cbox a b) f \<Longrightarrow> continuous_on S (ext_cont f a b)"
@@ -2021,18 +1851,16 @@
     and Un: "\<And>S. open S \<Longrightarrow> \<exists>U. U \<subseteq> \<B> \<and> S = \<Union>U"
     using univ_second_countable by blast
   have *: "infinite (range (\<lambda>n. ball (0::'a) (inverse(Suc n))))"
-    apply (rule Infinite_Set.range_inj_infinite)
-    apply (simp add: inj_on_def ball_eq_ball_iff)
-    done
+    by (simp add: inj_on_def ball_eq_ball_iff Infinite_Set.range_inj_infinite)
   have "infinite \<B>"
   proof
     assume "finite \<B>"
     then have "finite (Union ` (Pow \<B>))"
       by simp
-    then have "finite (range (\<lambda>n. ball (0::'a) (inverse(Suc n))))"
-      apply (rule rev_finite_subset)
+    moreover have "range (\<lambda>n. ball 0 (inverse (real (Suc n)))) \<subseteq> \<Union> ` Pow \<B>"
       by (metis (no_types, lifting) PowI image_eqI image_subset_iff Un [OF open_ball])
-    with * show False by simp
+    ultimately show False
+      by (metis finite_subset *)
   qed
   obtain f :: "nat \<Rightarrow> 'a set" where "\<B> = range f" "inj f"
     by (blast intro: countable_as_injective_image [OF \<open>countable \<B>\<close> \<open>infinite \<B>\<close>])
@@ -2043,9 +1871,7 @@
     using \<open>inj f\<close> \<open>\<B> = range f\<close> apply force
     done
   show ?thesis
-    apply (rule that [OF \<open>inj f\<close> _ *])
-    apply (auto simp: \<open>\<B> = range f\<close> opn)
-    done
+    using "*" \<open>\<B> = range f\<close> \<open>inj f\<close> opn that by force
 qed
 
 proposition separable:
@@ -2081,14 +1907,8 @@
           using \<open>0 < e\<close>  \<U> \<open>x \<in> S\<close> by auto
       next
         case False
-        then obtain C where "C \<in> \<U>" by blast
-        show ?thesis
-        proof
-          show "dist (f C) x < e"
-            by (metis Int_iff Union_iff \<U> \<open>C \<in> \<U>\<close> dist_commute f mem_ball subsetCE)
-          show "C \<in> \<B>"
-            using \<open>\<U> \<subseteq> \<B>\<close> \<open>C \<in> \<U>\<close> by blast
-        qed
+        then show ?thesis
+          by (metis IntI Union_iff \<U> \<open>0 < e\<close> \<open>x \<in> S\<close> dist_commute dist_self f inf_le2 mem_ball subset_eq)
       qed
     qed
   qed
@@ -2113,8 +1933,7 @@
       then show "norm (x - y) \<le> 2*r" by simp
     qed (simp add: that)
     have "2*r = dist (a + r *\<^sub>R (SOME i. i \<in> Basis)) (a - r *\<^sub>R (SOME i. i \<in> Basis))"
-      apply (simp add: dist_norm)
-      by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that)
+      using \<open>0 \<le> r\<close> that by (simp add: dist_norm flip: scaleR_2)
     also have "... \<le> diameter (cball a r)"
       apply (rule diameter_bounded_bound)
       using that by (auto simp: dist_norm)
@@ -2135,8 +1954,8 @@
 
 lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)"
 proof -
-  have "{a .. b} = cball ((a+b)/2) ((b-a)/2)"
-    by (auto simp: dist_norm abs_if field_split_simps split: if_split_asm)
+  have "{a..b} = cball ((a+b)/2) ((b-a)/2)"
+    using atLeastAtMost_eq_cball by blast
   then show ?thesis
     by simp
 qed
@@ -2144,7 +1963,7 @@
 lemma diameter_open_interval [simp]: "diameter {a<..<b} = (if b < a then 0 else b-a)"
 proof -
   have "{a <..< b} = ball ((a+b)/2) ((b-a)/2)"
-    by (auto simp: dist_norm abs_if field_split_simps split: if_split_asm)
+    using greaterThanLessThan_eq_ball by blast
   then show ?thesis
     by simp
 qed
@@ -2213,10 +2032,8 @@
     shows "open(f ` A) \<longleftrightarrow> open A"
 proof
   assume "open(f ` A)"
-  then have "open(f -` (f ` A))"
-    using assms by (force simp: linear_continuous_at linear_conv_bounded_linear continuous_open_vimage)
   then show "open A"
-    by (simp add: assms bij_is_inj inj_vimage_image_eq)
+    by (metis assms bij_is_inj continuous_open_vimage inj_vimage_image_eq linear_continuous_at linear_linear)
 next
   assume "open A"
   then show "open(f ` A)"
@@ -2226,20 +2043,9 @@
 corollary interior_bijective_linear_image:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "linear f" "bij f"
-  shows "interior (f ` S) = f ` interior S"  (is "?lhs = ?rhs")
-proof safe
-  fix x
-  assume x: "x \<in> ?lhs"
-  then obtain T where "open T" and "x \<in> T" and "T \<subseteq> f ` S"
-    by (metis interiorE)
-  then show "x \<in> ?rhs"
-    by (metis (no_types, opaque_lifting) assms subsetD interior_maximal open_bijective_linear_image_eq subset_image_iff)
-next
-  fix x
-  assume x: "x \<in> interior S"
-  then show "f x \<in> interior (f ` S)"
-    by (meson assms imageI image_mono interiorI interior_subset open_bijective_linear_image_eq open_interior)
-qed
+  shows "interior (f ` S) = f ` interior S" 
+  by (smt (verit) assms bij_is_inj inj_image_subset_iff interior_maximal interior_subset 
+      open_bijective_linear_image_eq open_interior subset_antisym subset_imageE)
 
 lemma interior_injective_linear_image:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
@@ -2346,10 +2152,9 @@
   shows "closed(f ` s)"
 proof -
   obtain e where "e > 0" and e: "\<forall>x\<in>s. e * norm x \<le> norm (f x)"
-    using injective_imp_isometric[OF assms(4,1,2,3)] by auto
-  show ?thesis
-    using complete_isometric_image[OF \<open>e>0\<close> assms(1,2) e] and assms(4)
-    unfolding complete_eq_closed[symmetric] by auto
+    using assms injective_imp_isometric by blast
+  with assms show ?thesis
+    by (meson complete_eq_closed complete_isometric_image)
 qed
                                
 
@@ -2386,9 +2191,8 @@
     show "closedin (top_of_set (range f)) (f ` S)"
       using continuous_closedin_preimage [OF confg cgf] by simp
     show "closed (range f)"
-      apply (rule closed_injective_image_subspace)
-      using f apply (auto simp: linear_linear linear_injective_0)
-      done
+      using closed_injective_image_subspace f linear_conv_bounded_linear 
+          linear_injective_0 subspace_UNIV by blast
   qed
 qed
 
@@ -2401,25 +2205,28 @@
 lemma closure_injective_linear_image:
     fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
     shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)"
-  apply (rule subset_antisym)
-  apply (simp add: closure_linear_image_subset)
-  by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono)
+  by (simp add: closed_injective_linear_image closure_linear_image_subset 
+        closure_minimal closure_subset image_mono subset_antisym)
 
 lemma closure_bounded_linear_image:
-    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-    shows "\<lbrakk>linear f; bounded S\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)"
-  apply (rule subset_antisym, simp add: closure_linear_image_subset)
-  apply (rule closure_minimal, simp add: closure_subset image_mono)
-  by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear)
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "linear f" "bounded S"
+    shows "f ` (closure S) = closure (f ` S)"  (is "?lhs = ?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    using assms closure_linear_image_subset by blast
+  show "?rhs \<subseteq> ?lhs"
+    using assms by (meson closure_minimal closure_subset compact_closure compact_eq_bounded_closed
+                      compact_continuous_image image_mono linear_continuous_on linear_linear)
+qed
 
 lemma closure_scaleR:
   fixes S :: "'a::real_normed_vector set"
-  shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)"
+  shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)"  (is "?lhs = ?rhs")
 proof
-  show "((*\<^sub>R) c) ` (closure S) \<subseteq> closure (((*\<^sub>R) c) ` S)"
-    using bounded_linear_scaleR_right
-    by (rule closure_bounded_linear_image_subset)
-  show "closure (((*\<^sub>R) c) ` S) \<subseteq> ((*\<^sub>R) c) ` (closure S)"
+  show "?lhs \<subseteq> ?rhs"
+    using bounded_linear_scaleR_right by (rule closure_bounded_linear_image_subset)
+  show "?rhs \<subseteq> ?lhs"
     by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure)
 qed
 
@@ -2438,73 +2245,55 @@
 qed
 
 lemma closed_subspace:
-  fixes s :: "'a::euclidean_space set"
-  assumes "subspace s"
-  shows "closed s"
+  fixes S :: "'a::euclidean_space set"
+  assumes "subspace S"
+  shows "closed S"
 proof -
-  have "dim s \<le> card (Basis :: 'a set)"
+  have "dim S \<le> card (Basis :: 'a set)"
     using dim_subset_UNIV by auto
-  with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \<subseteq> Basis"
-    by auto
+  with obtain_subset_with_card_n 
+  obtain d :: "'a set" where cd: "card d = dim S" and d: "d \<subseteq> Basis"
+    by metis
   let ?t = "{x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
-  have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s \<and>
+  have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = S \<and>
       inj_on f {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
-    using dim_substandard[of d] t d assms
+    using dim_substandard[of d] cd d assms
     by (intro subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]]) (auto simp: inner_Basis)
   then obtain f where f:
       "linear f"
-      "f ` {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s"
+      "f ` {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = S"
       "inj_on f {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
     by blast
   interpret f: bounded_linear f
     using f by (simp add: linear_conv_bounded_linear)
   have "x \<in> ?t \<Longrightarrow> f x = 0 \<Longrightarrow> x = 0" for x
     using f.zero d f(3)[THEN inj_onD, of x 0] by auto
-  moreover have "closed ?t" by (rule closed_substandard)
-  moreover have "subspace ?t" by (rule subspace_substandard)
-  ultimately show ?thesis
-    using closed_injective_image_subspace[of ?t f]
-    unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto
+  then show ?thesis
+    using closed_injective_image_subspace[of ?t f] closed_substandard subspace_substandard
+    using f(2) f.bounded_linear_axioms by force
 qed
 
-lemma complete_subspace: "subspace s \<Longrightarrow> complete s"
-  for s :: "'a::euclidean_space set"
+lemma complete_subspace: "subspace S \<Longrightarrow> complete S"
+  for S :: "'a::euclidean_space set"
   using complete_eq_closed closed_subspace by auto
 
-lemma closed_span [iff]: "closed (span s)"
-  for s :: "'a::euclidean_space set"
+lemma closed_span [iff]: "closed (span S)"
+  for S :: "'a::euclidean_space set"
   by (simp add: closed_subspace)
 
-lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d")
-  for s :: "'a::euclidean_space set"
-proof -
-  have "?dc \<le> ?d"
-    using closure_minimal[OF span_superset, of s]
-    using closed_subspace[OF subspace_span, of s]
-    using dim_subset[of "closure s" "span s"]
-    by simp
-  then show ?thesis
-    using dim_subset[OF closure_subset, of s]
-    by simp
-qed
+lemma dim_closure [simp]: "dim (closure S) = dim S" (is "?dc = ?d")
+  for S :: "'a::euclidean_space set"
+  by (metis closed_span closure_minimal closure_subset dim_eq_span span_eq_dim span_superset subset_le_dim)
 
 
 subsection \<open>Set Distance\<close>
 
 lemma setdist_compact_closed:
   fixes A :: "'a::heine_borel set"
-  assumes A: "compact A" and B: "closed B"
+  assumes "compact A" "closed B"
     and "A \<noteq> {}" "B \<noteq> {}"
   shows "\<exists>x \<in> A. \<exists>y \<in> B. dist x y = setdist A B"
-proof -
-  obtain x where "x \<in> A" "setdist A B = infdist x B"
-    by (metis A assms(3) setdist_attains_inf setdist_sym)
-  moreover
-  obtain y where"y \<in> B" "infdist x B = dist x y"
-    using B \<open>B \<noteq> {}\<close> infdist_attains_inf by blast
-  ultimately show ?thesis
-    using \<open>x \<in> A\<close> \<open>y \<in> B\<close> by auto
-qed
+  by (metis assms infdist_attains_inf setdist_attains_inf setdist_sym)
 
 lemma setdist_closed_compact:
   fixes S :: "'a::heine_borel set"
@@ -2518,14 +2307,10 @@
   assumes S: "compact S" and T: "closed T"
     shows "setdist S T = 0 \<longleftrightarrow> S = {} \<or> T = {} \<or> S \<inter> T \<noteq> {}"
 proof (cases "S = {} \<or> T = {}")
-  case True
-  then show ?thesis
-    by force
-next
   case False
   then show ?thesis
-    by (metis S T disjoint_iff_not_equal in_closed_iff_infdist_zero setdist_attains_inf setdist_eq_0I setdist_sym)
-qed
+    by (metis S T disjoint_iff in_closed_iff_infdist_zero setdist_attains_inf setdist_eq_0I setdist_sym)
+qed auto
 
 corollary setdist_gt_0_compact_closed:
   assumes S: "compact S" and T: "closed T"
--- a/src/HOL/Complex_Analysis/Contour_Integration.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Complex_Analysis/Contour_Integration.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -1413,7 +1413,7 @@
 proof (cases "r = 0 \<or> s = t")
   case True
   then show ?thesis
-    unfolding part_circlepath_def simple_path_def
+    unfolding part_circlepath_def simple_path_def loop_free_def
     by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+
 next
   case False then have "r \<noteq> 0" "s \<noteq> t" by auto
@@ -1445,7 +1445,7 @@
   have abs_away: "\<And>P. (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. P \<bar>x - y\<bar>) \<longleftrightarrow> (\<forall>x::real. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> P x)"
     by force
   show ?thesis using False
-    apply (simp add: simple_path_def)
+    apply (simp add: simple_path_def loop_free_def)
     apply (simp add: part_circlepath_def linepath_def exp_eq  * ** abs01 del: Set.insert_iff)
     apply (subst abs_away)
     apply (auto simp: 1)
--- a/src/HOL/Complex_Analysis/Winding_Numbers.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Complex_Analysis/Winding_Numbers.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -1932,7 +1932,7 @@
     proof (rule simple_path_join_loop, simp_all add: qt q01)
       have "inj_on q (closed_segment t 0)"
         using \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close> \<open>t \<noteq> 0\<close> \<open>t \<noteq> 1\<close>
-        by (fastforce simp: simple_path_def inj_on_def closed_segment_eq_real_ivl)
+        by (fastforce simp: simple_path_def loop_free_def inj_on_def closed_segment_eq_real_ivl)
       then show "arc (subpath t 0 q)"
         using \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close> \<open>t \<noteq> 0\<close>
         by (simp add: arc_subpath_eq simple_path_imp_path)
@@ -1972,7 +1972,7 @@
     proof
       show "?lhs \<subseteq> ?rhs"
         using \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close> \<open>t \<noteq> 1\<close> q_ends qt q01
-        by (force simp: pathfinish_def qt simple_path_def path_image_subpath)
+        by (force simp: pathfinish_def qt simple_path_def loop_free_def path_image_subpath)
       show "?rhs \<subseteq> ?lhs"
         using \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> q01 qt by (force simp: path_image_subpath)
     qed
--- a/src/HOL/Library/BigO.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Library/BigO.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -42,49 +42,22 @@
 lemma bigo_pos_const:
   "(\<exists>c::'a::linordered_idom. \<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>) \<longleftrightarrow>
     (\<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>))"
-  apply auto
-  apply (case_tac "c = 0")
-   apply simp
-   apply (rule_tac x = "1" in exI)
-   apply simp
-  apply (rule_tac x = "\<bar>c\<bar>" in exI)
-  apply auto
-  apply (subgoal_tac "c * \<bar>f x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>")
-   apply (erule_tac x = x in allE)
-   apply force
-  apply (rule mult_right_mono)
-   apply (rule abs_ge_self)
-  apply (rule abs_ge_zero)
-  done
+  by (metis (no_types, opaque_lifting) abs_ge_zero abs_not_less_zero abs_of_nonneg dual_order.trans 
+        mult_1 zero_less_abs_iff zero_less_mult_iff zero_less_one)
 
 lemma bigo_alt_def: "O(f) = {h. \<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>)}"
   by (auto simp add: bigo_def bigo_pos_const)
 
 lemma bigo_elt_subset [intro]: "f \<in> O(g) \<Longrightarrow> O(f) \<le> O(g)"
   apply (auto simp add: bigo_alt_def)
-  apply (rule_tac x = "ca * c" in exI)
-  apply (rule conjI)
-   apply simp
-  apply (rule allI)
-  apply (drule_tac x = "xa" in spec)+
-  apply (subgoal_tac "ca * \<bar>f xa\<bar> \<le> ca * (c * \<bar>g xa\<bar>)")
-   apply (erule order_trans)
-   apply (simp add: ac_simps)
-  apply (rule mult_left_mono, assumption)
-  apply (rule order_less_imp_le, assumption)
-  done
+  by (metis (no_types, opaque_lifting) mult.assoc mult_le_cancel_iff2 order.trans 
+      zero_less_mult_iff)
 
 lemma bigo_refl [intro]: "f \<in> O(f)"
-  apply (auto simp add: bigo_def)
-  apply (rule_tac x = 1 in exI)
-  apply simp
-  done
+  using bigo_def comm_monoid_mult_class.mult_1 dual_order.eq_iff by blast
 
 lemma bigo_zero: "0 \<in> O(g)"
-  apply (auto simp add: bigo_def func_zero)
-  apply (rule_tac x = 0 in exI)
-  apply auto
-  done
+  using bigo_def mult_le_cancel_left1 by fastforce
 
 lemma bigo_zero2: "O(\<lambda>x. 0) = {\<lambda>x. 0}"
   by (auto simp add: bigo_def)
@@ -92,21 +65,10 @@
 lemma bigo_plus_self_subset [intro]: "O(f) + O(f) \<subseteq> O(f)"
   apply (auto simp add: bigo_alt_def set_plus_def)
   apply (rule_tac x = "c + ca" in exI)
-  apply auto
-  apply (simp add: ring_distribs func_plus)
-  apply (rule order_trans)
-   apply (rule abs_triangle_ineq)
-  apply (rule add_mono)
-   apply force
-  apply force
-  done
+  by (smt (verit, best) abs_triangle_ineq add_mono add_pos_pos comm_semiring_class.distrib dual_order.trans)
 
 lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)"
-  apply (rule equalityI)
-   apply (rule bigo_plus_self_subset)
-  apply (rule set_zero_plus2)
-  apply (rule bigo_zero)
-  done
+  by (simp add: antisym bigo_plus_self_subset bigo_zero set_zero_plus2)
 
 lemma bigo_plus_subset [intro]: "O(f + g) \<subseteq> O(f) + O(g)"
   apply (rule subsetI)
@@ -117,37 +79,22 @@
    apply (rule_tac x = "c + c" in exI)
    apply (clarsimp)
    apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> (c + c) * \<bar>f xa\<bar>")
-    apply (erule_tac x = xa in allE)
-    apply (erule order_trans)
-    apply (simp)
+    apply (metis mult_2 order_trans)
    apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> c * (\<bar>f xa\<bar> + \<bar>g xa\<bar>)")
-    apply (erule order_trans)
-    apply (simp add: ring_distribs)
-   apply (rule mult_left_mono)
-    apply (simp add: abs_triangle_ineq)
-   apply (simp add: order_less_le)
+    apply auto[1]
+  using abs_triangle_ineq mult_le_cancel_iff2 apply blast
+  apply (simp add: order_less_le)
   apply (rule_tac x = "\<lambda>n. if \<bar>f n\<bar> < \<bar>g n\<bar> then x n else 0" in exI)
   apply (rule conjI)
    apply (rule_tac x = "c + c" in exI)
    apply auto
   apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> (c + c) * \<bar>g xa\<bar>")
-   apply (erule_tac x = xa in allE)
-   apply (erule order_trans)
-   apply simp
-  apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> c * (\<bar>f xa\<bar> + \<bar>g xa\<bar>)")
-   apply (erule order_trans)
-   apply (simp add: ring_distribs)
-  apply (rule mult_left_mono)
-   apply (rule abs_triangle_ineq)
-  apply (simp add: order_less_le)
+   apply (metis mult_2 order.trans)
+  apply simp
   done
 
 lemma bigo_plus_subset2 [intro]: "A \<subseteq> O(f) \<Longrightarrow> B \<subseteq> O(f) \<Longrightarrow> A + B \<subseteq> O(f)"
-  apply (subgoal_tac "A + B \<subseteq> O(f) + O(f)")
-   apply (erule order_trans)
-   apply simp
-  apply (auto del: subsetI simp del: bigo_plus_idemp)
-  done
+  using bigo_plus_idemp set_plus_mono2 by blast
 
 lemma bigo_plus_eq: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. 0 \<le> g x \<Longrightarrow> O(f + g) = O(f) + O(g)"
   apply (rule equalityI)
@@ -157,8 +104,7 @@
   apply (rule_tac x = "max c ca" in exI)
   apply (rule conjI)
    apply (subgoal_tac "c \<le> max c ca")
-    apply (erule order_less_le_trans)
-    apply assumption
+    apply linarith
    apply (rule max.cobounded1)
   apply clarify
   apply (drule_tac x = "xa" in spec)+
@@ -167,21 +113,9 @@
    apply (subgoal_tac "\<bar>a xa + b xa\<bar> \<le> \<bar>a xa\<bar> + \<bar>b xa\<bar>")
     apply (subgoal_tac "\<bar>a xa\<bar> + \<bar>b xa\<bar> \<le> max c ca * f xa + max c ca * g xa")
      apply force
-    apply (rule add_mono)
-     apply (subgoal_tac "c * f xa \<le> max c ca * f xa")
-      apply force
-     apply (rule mult_right_mono)
-      apply (rule max.cobounded1)
-     apply assumption
-    apply (subgoal_tac "ca * g xa \<le> max c ca * g xa")
-     apply force
-    apply (rule mult_right_mono)
-     apply (rule max.cobounded2)
-    apply assumption
-   apply (rule abs_triangle_ineq)
-  apply (rule add_nonneg_nonneg)
-   apply assumption+
-  done
+    apply (metis add_mono le_max_iff_disj max_mult_distrib_right)
+  using abs_triangle_ineq apply blast
+  using add_nonneg_nonneg by blast
 
 lemma bigo_bounded_alt: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. f x \<le> c * g x \<Longrightarrow> f \<in> O(g)"
   apply (auto simp add: bigo_def)
--- a/src/HOL/Library/Float.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Library/Float.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -100,22 +100,11 @@
 qed
 
 lemma uminus_float[simp]: "x \<in> float \<Longrightarrow> -x \<in> float"
-  apply (auto simp: float_def)
-  apply hypsubst_thin
-  apply (rename_tac m e)
-  apply (rule_tac x="-m" in exI)
-  apply (rule_tac x="e" in exI)
-  apply (simp add: field_simps)
-  done
+  by (simp add: float_def) (metis mult_minus_left of_int_minus)
 
 lemma times_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x * y \<in> float"
-  apply (auto simp: float_def)
-  apply hypsubst_thin
-  apply (rename_tac mx my ex ey)
-  apply (rule_tac x="mx * my" in exI)
-  apply (rule_tac x="ex + ey" in exI)
-  apply (simp add: powr_add)
-  done
+  apply (clarsimp simp: float_def)
+  by (metis (no_types, opaque_lifting) of_int_add powr_add mult.assoc mult.left_commute of_int_mult)
 
 lemma minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x - y \<in> float"
   using plus_float [of x "- y"] by simp
@@ -126,23 +115,11 @@
 lemma sgn_of_float[simp]: "x \<in> float \<Longrightarrow> sgn x \<in> float"
   by (cases x rule: linorder_cases[of 0]) (auto intro!: uminus_float)
 
-lemma div_power_2_float[simp]: "x \<in> float \<Longrightarrow> x / 2^d \<in> float"
-  apply (auto simp add: float_def)
-  apply hypsubst_thin
-  apply (rename_tac m e)
-  apply (rule_tac x="m" in exI)
-  apply (rule_tac x="e - d" in exI)
-  apply (simp flip: powr_realpow powr_add add: field_simps)
-  done
+lemma div_power_2_float[simp]: "x \<in> float \<Longrightarrow> x / 2^d \<in> float" 
+  by (simp add: float_def) (metis of_int_diff of_int_of_nat_eq powr_diff powr_realpow zero_less_numeral times_divide_eq_right)
 
 lemma div_power_2_int_float[simp]: "x \<in> float \<Longrightarrow> x / (2::int)^d \<in> float"
-  apply (auto simp add: float_def)
-  apply hypsubst_thin
-  apply (rename_tac m e)
-  apply (rule_tac x="m" in exI)
-  apply (rule_tac x="e - d" in exI)
-  apply (simp flip: powr_realpow powr_add add: field_simps)
-  done
+  by simp
 
 lemma div_numeral_Bit0_float[simp]:
   assumes "x / numeral n \<in> float"
@@ -158,13 +135,7 @@
 lemma div_neg_numeral_Bit0_float[simp]:
   assumes "x / numeral n \<in> float"
   shows "x / (- numeral (Num.Bit0 n)) \<in> float"
-proof -
-  have "- (x / numeral (Num.Bit0 n)) \<in> float"
-    using assms by simp
-  also have "- (x / numeral (Num.Bit0 n)) = x / - numeral (Num.Bit0 n)"
-    by simp
-  finally show ?thesis .
-qed
+  using assms by force
 
 lemma power_float[simp]:
   assumes "a \<in> float"
@@ -251,15 +222,9 @@
 proof
   fix a b :: float
   show "\<exists>c. a < c"
-    apply (intro exI[of _ "a + 1"])
-    apply transfer
-    apply simp
-    done
+    by (metis Float.real_of_float less_float.rep_eq reals_Archimedean2)
   show "\<exists>c. c < a"
-    apply (intro exI[of _ "a - 1"])
-    apply transfer
-    apply simp
-    done
+    by (metis add_0 add_strict_right_mono neg_less_0_iff_less zero_less_one)
   show "\<exists>c. a < c \<and> c < b" if "a < b"
     apply (rule exI[of _ "(a + b) * Float 1 (- 1)"])
     using that
@@ -283,11 +248,10 @@
 end
 
 lemma float_numeral[simp]: "real_of_float (numeral x :: float) = numeral x"
-  apply (induct x)
-  apply simp
-  apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq float_of_inverse
-          plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float)
-  done
+proof (induct x)
+  case One
+  then show ?case by simp
+qed (metis of_int_numeral real_of_float_of_int_eq)+
 
 lemma transfer_numeral [transfer_rule]:
   "rel_fun (=) pcr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)"
@@ -837,11 +801,11 @@
   finally show ?thesis
     using \<open>p + e < 0\<close>
     apply transfer
-    apply (simp add: round_down_def field_simps flip: floor_divide_of_int_eq)
+    apply (simp add: round_down_def field_simps flip: floor_divide_of_int_eq powr_add)
     apply (metis (no_types, opaque_lifting) Float.rep_eq
       add.inverse_inverse compute_real_of_float diff_minus_eq_add
       floor_divide_of_int_eq int_of_reals(1) linorder_not_le
-      minus_add_distrib of_int_eq_numeral_power_cancel_iff powr_add)
+      minus_add_distrib of_int_eq_numeral_power_cancel_iff )
     done
 next
   case False
@@ -885,10 +849,7 @@
     have "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> \<le> real_of_int a / real_of_int b"
       by simp
     moreover have "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> \<noteq> real_of_int a / real_of_int b"
-      apply (subst (2) real_of_int_div_aux)
-      unfolding floor_divide_of_int_eq
-      using ne \<open>b \<noteq> 0\<close> apply auto
-      done
+      by (smt (verit) floor_divide_of_int_eq ne real_of_int_div_aux)
     ultimately show "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> < real_of_int a / real_of_int b" by arith
   qed
   then show ?thesis
@@ -1254,12 +1215,7 @@
   by (auto simp: truncate_up_uminus_eq truncate_down_uminus_eq)
 
 lemma truncate_down_mono: "x \<le> y \<Longrightarrow> truncate_down p x \<le> truncate_down p y"
-  apply (cases "0 \<le> x")
-  apply (rule truncate_down_nonneg_mono, assumption+)
-  apply (simp add: truncate_down_eq_truncate_up)
-  apply (cases "0 \<le> y")
-  apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono)
-  done
+  by (smt (verit) truncate_down_nonneg_mono truncate_up_nonneg_mono truncate_up_uminus_eq)
 
 lemma truncate_up_mono: "x \<le> y \<Longrightarrow> truncate_up p x \<le> truncate_up p y"
   by (simp add: truncate_up_eq_truncate_down truncate_down_mono)
@@ -1271,22 +1227,7 @@
   by (meson less_le_trans that truncate_up)
 
 lemma truncate_up_less_zero_iff[simp]: "truncate_up p x < 0 \<longleftrightarrow> x < 0"
-proof -
-  have f1: "\<forall>n r. truncate_up n r + truncate_down n (- 1 * r) = 0"
-    by (simp add: truncate_down_uminus_eq)
-  have f2: "(\<forall>v0 v1. truncate_up v0 v1 + truncate_down v0 (- 1 * v1) = 0) = (\<forall>v0 v1. truncate_up v0 v1 = - 1 * truncate_down v0 (- 1 * v1))"
-    by (auto simp: truncate_up_eq_truncate_down)
-  have f3: "\<forall>x1. ((0::real) < x1) = (\<not> x1 \<le> 0)"
-    by fastforce
-  have "(- 1 * x \<le> 0) = (0 \<le> x)"
-    by force
-  then have "0 \<le> x \<or> \<not> truncate_down p (- 1 * x) \<le> 0"
-    using f3 by (meson truncate_down_pos)
-  then have "(0 \<le> truncate_up p x) \<noteq> (\<not> 0 \<le> x)"
-    using f2 f1 truncate_up_nonneg by force
-  then show ?thesis
-    by linarith
-qed
+  by (smt (verit) truncate_down_pos truncate_down_uminus_eq truncate_up_nonneg)
 
 lemma truncate_up_nonneg_iff[simp]: "truncate_up p x \<ge> 0 \<longleftrightarrow> x \<ge> 0"
   using truncate_up_less_zero_iff[of p x] truncate_up_nonneg[of x]
@@ -1913,7 +1854,6 @@
   by (meson dual_order.trans mult_left_mono_neg mult_right_mono_neg that)
 
 lemma mult_float_mono1:
-  notes mono_rules = plus_down_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono
   shows "a \<le> b \<Longrightarrow> ab \<le> bb \<Longrightarrow>
        aa \<le> a \<Longrightarrow>
        b \<le> ba \<Longrightarrow>
@@ -1927,20 +1867,10 @@
            (plus_down prec (nprt b * nprt bb)
              (plus_down prec (pprt a * pprt ab)
                (pprt b * nprt ab)))"
-  apply (rule order_trans)
-   apply (rule mono_rules | assumption)+
-    apply (rule mult_mono_nonpos_nonneg)
-       apply (rule mono_rules | assumption)+
-    apply (rule mult_mono_nonpos_nonpos)
-       apply (rule mono_rules | assumption)+
-    apply (rule mult_mono)
-       apply (rule mono_rules | assumption)+
-   apply (rule mult_mono_nonneg_nonpos)
-      apply (rule mono_rules | assumption)+
-  by (rule order_refl)+
+  by (smt (verit, del_insts) mult_mono plus_down_mono add_mono nprt_mono nprt_le_zero zero_le_pprt 
+pprt_mono mult_mono_nonpos_nonneg mult_mono_nonpos_nonpos mult_mono_nonneg_nonpos)
 
 lemma mult_float_mono2:
-  notes mono_rules = plus_up_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono
   shows "a \<le> b \<Longrightarrow>
        ab \<le> bb \<Longrightarrow>
        aa \<le> a \<Longrightarrow>
@@ -1955,17 +1885,8 @@
            (plus_up prec (pprt aa * nprt bc)
              (plus_up prec (nprt ba * pprt ac)
                (nprt aa * nprt ac)))"
-  apply (rule order_trans)
-   apply (rule mono_rules | assumption)+
-    apply (rule mult_mono)
-       apply (rule mono_rules | assumption)+
-    apply (rule mult_mono_nonneg_nonpos)
-       apply (rule mono_rules | assumption)+
-    apply (rule mult_mono_nonpos_nonneg)
-       apply (rule mono_rules | assumption)+
-   apply (rule mult_mono_nonpos_nonpos)
-      apply (rule mono_rules | assumption)+
-  by (rule order_refl)+
+  by (smt (verit, del_insts) plus_up_mono add_mono mult_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono 
+      mult_mono_nonpos_nonneg mult_mono_nonpos_nonpos mult_mono_nonneg_nonpos)
 
 
 subsection \<open>Approximate Power\<close>
@@ -2132,24 +2053,21 @@
       assume [simp]: "odd j"
       have "power_down prec 0 (Suc (j div 2)) \<le> - power_down prec b (Suc (j div 2))"
         if "b < 0" "even (j div 2)"
-        apply (rule order_trans[where y=0])
-        using IH that by (auto simp: div2_less_self)
+        by (metis even_Suc le_minus_iff Suc_neq_Zero neg_equal_zero power_down_eq_zero_iff
+              power_down_nonpos_iff that)
       then have "truncate_down (Suc prec) ((power_down prec a (Suc (j div 2)))\<^sup>2)
         \<le> truncate_down (Suc prec) ((power_down prec b (Suc (j div 2)))\<^sup>2)"
-        using IH
-        by (auto intro!: truncate_down_mono intro: order_trans[where y=0]
-            simp: abs_le_square_iff[symmetric] abs_real_def
-            div2_less_self)
+        by (smt (verit) IH Suc_less_eq \<open>odd j\<close> div2_less_self mult_mono_nonpos_nonpos 
+            Suc_neq_Zero power2_eq_square power_down_neg_iff power_down_nonpos_iff power_mono truncate_down_mono)
       then show ?thesis
-        unfolding j
-        by (simp add: power_down_simp)
+        unfolding j by (simp add: power_down_simp)
     qed
   qed simp
 qed
 
 lemma power_up_even_nonneg: "even n \<Longrightarrow> 0 \<le> power_up p x n"
   by (induct p x n rule: power_up.induct)
-    (auto simp: power_up.simps simp del: odd_Suc_div_two intro!: )
+    (auto simp: power_up.simps simp del: odd_Suc_div_two)
 
 lemma power_up_eq_zero_iff[simp]: "power_up prec b n = 0 \<longleftrightarrow> b = 0 \<and> n \<noteq> 0"
 proof (induction n arbitrary: b rule: less_induct)
@@ -2253,18 +2171,7 @@
   fixes a b :: int
   assumes "b > 0"
   shows "a \<ge> b * (a div b)"
-proof -
-  from minus_div_mult_eq_mod [symmetric, of a b] have "a = b * (a div b) + a mod b"
-    by simp
-  also have "\<dots> \<ge> b * (a div b) + 0"
-    apply (rule add_left_mono)
-    apply (rule pos_mod_sign)
-    using assms
-    apply simp
-    done
-  finally show ?thesis
-    by simp
-qed
+  by (smt (verit, ccfv_threshold) assms minus_div_mult_eq_mod mod_int_pos_iff mult.commute)
 
 lemma lapprox_rat_nonneg:
   assumes "0 \<le> x" and "0 \<le> y"
@@ -2393,9 +2300,8 @@
 qualified lemma compute_int_floor_fl[code]:
   "int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))"
   apply transfer
-  apply (simp add: powr_int floor_divide_of_int_eq)
-  apply (metis (no_types, opaque_lifting)floor_divide_of_int_eq of_int_numeral of_int_power floor_of_int of_int_mult)
-  done
+  by (smt (verit, ccfv_threshold) Float.rep_eq compute_real_of_float floor_divide_of_int_eq 
+      floor_of_int of_int_1 of_int_add of_int_mult of_int_power)
 
 lift_definition floor_fl :: "float \<Rightarrow> float" is "\<lambda>x. real_of_int \<lfloor>x\<rfloor>"
   by simp
@@ -2404,8 +2310,7 @@
   "floor_fl (Float m e) = (if 0 \<le> e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)"
   apply transfer
   apply (simp add: powr_int floor_divide_of_int_eq)
-  apply (metis (no_types, opaque_lifting)floor_divide_of_int_eq of_int_numeral of_int_power of_int_mult)
-  done
+  by (smt (z3) floor_divide_of_int_eq of_int_1 of_int_add of_int_power)
 
 end
 
--- a/src/HOL/SPARK/Tools/spark.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/SPARK/Tools/spark.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -11,10 +11,12 @@
 
 object SPARK {
   class Load_Command1 extends Command_Span.Load_Command("spark_vcg", Scala_Project.here) {
+    override def toString: String = print_extensions
     override val extensions: List[String] = List("vcg", "fdl", "rls")
   }
 
   class Load_Command2 extends Command_Span.Load_Command("spark_siv", Scala_Project.here) {
+    override def toString: String = print_extensions
     override val extensions: List[String] = List("siv", "fdl", "rls")
   }
 }
--- a/src/HOL/Transcendental.thy	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/HOL/Transcendental.thy	Tue Jan 03 18:23:52 2023 +0100
@@ -5477,6 +5477,64 @@
     using that by metis
 qed
 
+lemma arccos_arctan:
+  assumes "-1 < x" "x < 1"
+  shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))"
+proof -
+  have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0"
+  proof (rule sin_eq_0_pi)
+    show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)"
+      using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"]  arccos_bounded [of x] assms
+      by (simp add: algebra_simps)
+  next
+    show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x) < pi"
+      using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"]  arccos_bounded [of x] assms
+      by (simp add: algebra_simps)
+  next
+    show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)) = 0"
+      using assms
+      by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan
+                    power2_eq_square square_eq_1_iff)
+  qed
+  then show ?thesis
+    by simp
+qed
+
+lemma arcsin_plus_arccos:
+  assumes "-1 \<le> x" "x \<le> 1"
+    shows "arcsin x + arccos x = pi/2"
+proof -
+  have "arcsin x = pi/2 - arccos x"
+    apply (rule sin_inj_pi)
+    using assms arcsin [OF assms] arccos [OF assms]
+    by (auto simp: algebra_simps sin_diff)
+  then show ?thesis
+    by (simp add: algebra_simps)
+qed
+
+lemma arcsin_arccos_eq: "-1 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arcsin x = pi/2 - arccos x"
+  using arcsin_plus_arccos by force
+
+lemma arccos_arcsin_eq: "-1 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arccos x = pi/2 - arcsin x"
+  using arcsin_plus_arccos by force
+
+lemma arcsin_arctan: "-1 < x \<Longrightarrow> x < 1 \<Longrightarrow> arcsin x = arctan(x / sqrt(1 - x\<^sup>2))"
+  by (simp add: arccos_arctan arcsin_arccos_eq)
+
+lemma arcsin_arccos_sqrt_pos: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arcsin x = arccos(sqrt(1 - x\<^sup>2))"
+  by (smt (verit, del_insts) arccos_cos arcsin_0 arcsin_le_arcsin arcsin_pi cos_arcsin)
+
+lemma arcsin_arccos_sqrt_neg: "-1 \<le> x \<Longrightarrow> x \<le> 0 \<Longrightarrow> arcsin x = -arccos(sqrt(1 - x\<^sup>2))"
+  using arcsin_arccos_sqrt_pos [of "-x"]
+  by (simp add: arcsin_minus)
+
+lemma arccos_arcsin_sqrt_pos: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arccos x = arcsin(sqrt(1 - x\<^sup>2))"
+  by (smt (verit, del_insts) arccos_lbound arccos_le_pi2 arcsin_sin sin_arccos)
+
+lemma arccos_arcsin_sqrt_neg: "-1 \<le> x \<Longrightarrow> x \<le> 0 \<Longrightarrow> arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))"
+  using arccos_arcsin_sqrt_pos [of "-x"]
+  by (simp add: arccos_minus)
+
 lemma cos_limit_1:
   assumes "(\<lambda>j. cos (\<theta> j)) \<longlonglongrightarrow> 1"
   shows "\<exists>k. (\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> 0"
--- a/src/Pure/Admin/build_log.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Admin/build_log.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -772,7 +772,7 @@
       SQL.select(columns, distinct = true) +
         table1.query_named + SQL.join_outer + aux_table.query_named +
         " ON " + version(table1) + " = " + version(aux_table) +
-        " ORDER BY " + pull_date(afp)(table1) + " DESC"
+        SQL.order_by(List(pull_date(afp)(table1)), descending = true)
     }
 
 
--- a/src/Pure/Admin/check_sources.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Admin/check_sources.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -10,7 +10,7 @@
 object Check_Sources {
   def check_file(path: Path): Unit = {
     val file_name = path.implode
-    val file_pos = path.position
+    val file_pos = Position.File(path.implode_symbolic)
     def line_pos(i: Int) = Position.Line_File(i + 1, file_name)
 
     if (space_explode('/', Word.lowercase(path.expand.drop_ext.implode)).contains("aux"))
--- a/src/Pure/GUI/gui.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/GUI/gui.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -153,9 +153,7 @@
     }
 
     private def make_entries[A](batches: List[List[Entry[A]]]): List[Entry[A]] = {
-      val item_batches =
-        batches.map(_.flatMap(
-          { case item: Item[_] => Some(item.asInstanceOf[Item[A]]) case _ => None }))
+      val item_batches = batches.map(_.flatMap(Library.as_subclass(classOf[Item[A]])))
       val sep_entries: List[Entry[A]] =
         item_batches.filter(_.nonEmpty).zipWithIndex.flatMap({ case (batch, i) =>
           Separator[A](i) :: batch.map(_.copy(batch = i))
--- a/src/Pure/General/graph.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/General/graph.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -27,7 +27,7 @@
 
   def make[Key, A](
     entries: List[((Key, A), List[Key])],
-    symmetric: Boolean = false)(
+    converse: Boolean = false)(
     implicit ord: Ordering[Key]
   ): Graph[Key, A] = {
     val graph1 =
@@ -38,7 +38,7 @@
       entries.foldLeft(graph1) {
         case (graph, ((x, _), ys)) =>
           ys.foldLeft(graph) {
-            case (g, y) => if (symmetric) g.add_edge(y, x) else g.add_edge(x, y)
+            case (g, y) => if (converse) g.add_edge(y, x) else g.add_edge(x, y)
           }
       }
     graph2
--- a/src/Pure/General/path.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/General/path.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -312,8 +312,6 @@
       } catch { case ERROR(_) => None }).headOption.getOrElse(implode)
   }
 
-  def position: Position.T = Position.File(implode_symbolic)
-
 
   /* platform files */
 
--- a/src/Pure/General/sql.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/General/sql.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -127,6 +127,9 @@
     override def toString: Source = ident
   }
 
+  def order_by(columns: List[Column], descending: Boolean = false): Source =
+    " ORDER BY " + columns.mkString(", ") + (if (descending) " DESC" else "")
+
 
   /* tables */
 
--- a/src/Pure/General/url.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/General/url.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -102,13 +102,37 @@
   def canonical_file_name(uri: String): String = canonical_file(uri).getPath
 
 
-  /* generic path notation: local, ssh, rsync, ftp, http */
+  /* generic path notation: standard, platform, ssh, rsync, ftp, http */
+
+  private val separators1 = "/\\"
+  private val separators2 = ":/\\"
+
+  def is_base_name(s: String, suffix: String = ""): Boolean =
+    s.nonEmpty && !s.exists(separators2.contains) && s.endsWith(suffix)
+
+  def get_base_name(s: String, suffix: String = ""): Option[String] = {
+    val i = s.lastIndexWhere(separators2.contains)
+    if (i + 1 >= s.length) None else Library.try_unsuffix(suffix, s.substring(i + 1))
+  }
+
+  def strip_base_name(s: String, suffix: String = ""): Option[String] = {
+    val i = s.lastIndexWhere(separators2.contains)
+    val j = s.lastIndexWhere(c => !separators1.contains(c), end = i)
+    if (i + 1 >= s.length || !s.endsWith(suffix)) None
+    else if (j < 0) Some(s.substring(0, i + 1))
+    else Some(s.substring(0, j + 1))
+  }
 
   def append_path(prefix: String, suffix: String): String =
-    if (prefix.endsWith(":") || prefix.endsWith("/") || prefix.isEmpty) prefix + suffix
-    else if (prefix.endsWith(":.") || prefix.endsWith("/.") || prefix == ".") {
+    if (prefix.endsWith(":") || prefix.endsWith("/") || prefix.endsWith("\\") || prefix.isEmpty) {
+      prefix + suffix
+    }
+    else if (prefix.endsWith(":.") || prefix.endsWith("/.") || prefix.endsWith("\\.") || prefix == ".") {
       prefix.substring(0, prefix.length - 1) + suffix
     }
+    else if (prefix.contains('\\') || suffix.contains('\\')) {
+      prefix + "\\" + suffix
+    }
     else prefix + "/" + suffix
 
   def direct_path(prefix: String): String = append_path(prefix, ".")
--- a/src/Pure/Isar/method.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Isar/method.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -667,7 +667,7 @@
       (keyword_positions text);
 
 fun report text_range =
-  if Context_Position.pide_reports ()
+  if Context_Position.reports_enabled0 ()
   then Position.reports (reports_of text_range) else ();
 
 end;
--- a/src/Pure/Isar/outer_syntax.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Isar/outer_syntax.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -163,8 +163,7 @@
       Parse.opt_target -- parse >> (fn (target, f) => trans restricted target f)));
 
 val local_theory' =
-  local_theory_command
-    (fn a => fn b => fn c => Toplevel.local_theory' a b c Toplevel.no_presentation);
+  local_theory_command (fn a => fn b => fn c => Toplevel.local_theory' a b c NONE);
 val local_theory = local_theory_command Toplevel.local_theory;
 val local_theory_to_proof' = local_theory_command Toplevel.local_theory_to_proof';
 val local_theory_to_proof = local_theory_command Toplevel.local_theory_to_proof;
--- a/src/Pure/Isar/toplevel.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Isar/toplevel.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -29,9 +29,7 @@
   val pretty_state: state -> Pretty.T list
   val string_of_state: state -> string
   val pretty_abstract: state -> Pretty.T
-  type presentation = state -> Latex.text option
-  val presentation: (state -> Latex.text) -> presentation
-  val no_presentation: presentation
+  type presentation = state -> Latex.text
   type transition
   val empty: transition
   val name_of: transition -> string
@@ -46,7 +44,7 @@
   val is_init: transition -> bool
   val modify_init: (unit -> theory) -> transition -> transition
   val exit: transition -> transition
-  val present: (state -> Latex.text) -> transition -> transition
+  val present: presentation -> transition -> transition
   val keep: (state -> unit) -> transition -> transition
   val keep': (bool -> state -> unit) -> transition -> transition
   val keep_proof: (state -> unit) -> transition -> transition
@@ -55,17 +53,17 @@
   val ignored: Position.T -> transition
   val malformed: Position.T -> string -> transition
   val generic_theory: (generic_theory -> generic_theory) -> transition -> transition
-  val theory': (bool -> theory -> theory) -> presentation -> transition -> transition
+  val theory': (bool -> theory -> theory) -> presentation option -> transition -> transition
   val theory: (theory -> theory) -> transition -> transition
   val begin_main_target: bool -> (theory -> local_theory) -> transition -> transition
   val end_main_target: transition -> transition
   val begin_nested_target: (Context.generic -> local_theory) -> transition -> transition
   val end_nested_target: transition -> transition
   val local_theory': (bool * Position.T) option -> (xstring * Position.T) option ->
-    (bool -> local_theory -> local_theory) -> presentation -> transition -> transition
+    (bool -> local_theory -> local_theory) -> presentation option -> transition -> transition
   val local_theory: (bool * Position.T) option -> (xstring * Position.T) option ->
     (local_theory -> local_theory) -> transition -> transition
-  val present_local_theory: (xstring * Position.T) option -> (state -> Latex.text) ->
+  val present_local_theory: (xstring * Position.T) option -> presentation ->
     transition -> transition
   val local_theory_to_proof': (bool * Position.T) option -> (xstring * Position.T) option ->
     (bool -> local_theory -> Proof.state) -> transition -> transition
@@ -141,12 +139,12 @@
   (node, cases_node init_presentation Context.proof_of Proof.context_of node);
 
 datatype state =
-  State of node_presentation * (theory option * Latex.text option);
+  State of node_presentation * (theory option * Latex.text future option);
     (*current node with presentation context, previous theory, document output*)
 
 fun node_of (State ((node, _), _)) = node;
 fun previous_theory_of (State (_, (prev_thy, _))) = prev_thy;
-fun output_of (State (_, (_, output))) = output;
+fun output_of (State (_, (_, output))) = Option.map Future.join output;
 
 fun make_state opt_thy =
   let val node = (case opt_thy of NONE => Toplevel | SOME thy => Theory (Context.Theory thy))
@@ -256,14 +254,9 @@
 
 (** toplevel transitions **)
 
-(* presentation *)
+(* primitive transitions *)
 
-type presentation = state -> Latex.text option;
-fun presentation g : presentation = SOME o g;
-val no_presentation: presentation = K NONE;
-
-
-(* primitive transitions *)
+type presentation = state -> Latex.text;
 
 datatype trans =
   (*init theory*)
@@ -271,78 +264,76 @@
   (*formal exit of theory*)
   Exit |
   (*keep state unchanged*)
-  Keep of bool -> presentation |
+  Keep of (bool -> state -> unit) * presentation option |
   (*node transaction and presentation*)
-  Transaction of (bool -> node -> node_presentation) * presentation;
+  Transaction of (bool -> node -> node_presentation) * presentation option;
 
 local
 
-exception FAILURE of state * exn;
-
-fun apply_presentation g (st as State (node, (prev_thy, _))) =
-  State (node, (prev_thy, g st));
-
-fun apply f g node =
+fun present_state fork g node_pr prev_thy =
   let
-    val node_pr = node_presentation node;
-    val context = cases_proper_node I (Context.Proof o Proof.context_of) node;
-    fun make_state node_pr' = State (node_pr', (get_theory node, NONE));
+    val state = State (node_pr, (prev_thy, NONE));
+    fun present pr =
+      if fork andalso Future.proofs_enabled 1 then
+        Execution.fork {name = "Toplevel.present_state", pos = Position.thread_data (), pri = ~1}
+          (fn () => pr state)
+      else Future.value (pr state);
+  in State (node_pr, (prev_thy, Option.map present g)) end;
 
-    val (st', err) =
-      (Runtime.controlled_execution (SOME context) (f #> make_state #> apply_presentation g) node,
-        NONE) handle exn => (make_state node_pr, SOME exn);
-  in
-    (case err of
-      NONE => st'
-    | SOME exn => raise FAILURE (st', exn))
-  end;
+fun no_update f g state =
+  Runtime.controlled_execution (try generic_theory_of state)
+    (fn () =>
+      let
+        val prev_thy = previous_theory_of state;
+        val () = f state;
+        val node_pr = node_presentation (node_of state);
+      in present_state false g node_pr prev_thy end) ()
+
+fun update f g state =
+  Runtime.controlled_execution (try generic_theory_of state)
+    (fn () =>
+      let
+        val prev_thy = previous_theory_of state;
+        val node_pr' = f (node_of state);
+      in present_state true g node_pr' prev_thy end) ();
 
 fun apply_tr int trans state =
   (case (trans, node_of state) of
     (Init f, Toplevel) =>
       Runtime.controlled_execution NONE (fn () =>
         State (node_presentation (Theory (Context.Theory (f ()))), (NONE, NONE))) ()
-  | (Exit, node as Theory (Context.Theory thy)) =>
+  | (Exit, Theory (Context.Theory thy)) =>
       let
         val State ((node', pr_ctxt), _) =
-          node |> apply
+          state |> update
             (fn _ =>
               node_presentation
                 (Theory (Context.Theory (tap Thm.expose_theory (Theory.end_theory thy)))))
-            no_presentation;
+            NONE;
       in State ((Toplevel, pr_ctxt), (get_theory node', NONE)) end
-  | (Keep f, node) =>
-      Runtime.controlled_execution (try generic_theory_of state)
-        (fn () =>
-          let
-            val prev_thy = previous_theory_of state;
-            val state' = State (node_presentation node, (prev_thy, NONE));
-          in apply_presentation (fn st => f int st) state' end) ()
+  | (Keep (f, g), _) => no_update (fn x => f int x) g state
   | (Transaction _, Toplevel) => raise UNDEF
-  | (Transaction (f, g), node) => apply (fn x => f int x) g node
+  | (Transaction (f, g), _) => update (fn x => f int x) g state
   | _ => raise UNDEF);
 
-fun apply_union _ [] state = raise FAILURE (state, UNDEF)
-  | apply_union int (tr :: trs) state =
-      apply_union int trs state
-        handle Runtime.UNDEF => apply_tr int tr state
-          | FAILURE (alt_state, UNDEF) => apply_tr int tr alt_state
-          | exn as FAILURE _ => raise exn
-          | exn => raise FAILURE (state, exn);
+fun apply_body _ [] _ = raise UNDEF
+  | apply_body int [tr] state = apply_tr int tr state
+  | apply_body int (tr :: trs) state =
+      apply_body int trs state
+        handle Runtime.UNDEF => apply_tr int tr state;
 
 fun apply_markers name markers (state as State ((node, pr_ctxt), prev_thy)) =
   let
     val state' =
       Runtime.controlled_execution (try generic_theory_of state)
         (fn () => State ((node, fold (Document_Marker.evaluate name) markers pr_ctxt), prev_thy)) ();
-  in (state', NONE) end
-  handle exn => (state, SOME exn);
+  in (state', NONE) end;
 
 in
 
-fun apply_trans int name markers trans state =
-  (apply_union int trans state |> apply_markers name markers)
-  handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);
+fun apply_capture int name markers trans state =
+  (apply_body int trans state |> apply_markers name markers)
+    handle exn => (state, SOME exn);
 
 end;
 
@@ -412,12 +403,12 @@
 val exit = add_trans Exit;
 
 fun present_transaction f g = add_trans (Transaction (f, g));
-fun transaction f = present_transaction f no_presentation;
-fun transaction0 f = present_transaction (node_presentation oo f) no_presentation;
+fun transaction f = present_transaction f NONE;
+fun transaction0 f = present_transaction (node_presentation oo f) NONE;
 
-fun present f = add_trans (Keep (K (presentation f)));
-fun keep f = add_trans (Keep (fn _ => fn st => let val () = f st in NONE end));
-fun keep' f = add_trans (Keep (fn int => fn st => let val () = f int st in NONE end));
+fun present g = add_trans (Keep (fn _ => fn _ => (), SOME g));
+fun keep f = add_trans (Keep (K f, NONE));
+fun keep' f = add_trans (Keep (f, NONE));
 
 fun keep_proof f =
   keep (fn st =>
@@ -453,7 +444,7 @@
       in node_presentation (Theory (Context.Theory thy')) end
     | _ => raise UNDEF));
 
-fun theory f = theory' (K f) no_presentation;
+fun theory f = theory' (K f) NONE;
 
 fun begin_main_target begin f = transaction (fn _ =>
   (fn Theory (Context.Theory thy) =>
@@ -505,14 +496,14 @@
     | _ => raise UNDEF));
 
 fun local_theory restricted target f =
-  local_theory' restricted target (K f) no_presentation;
+  local_theory' restricted target (K f) NONE;
 
 fun present_local_theory target g = present_transaction (fn _ =>
   (fn Theory gthy =>
         let val (finish, lthy) = Target_Context.switch_named_cmd target gthy;
         in (Theory (finish lthy), lthy) end
     | _ => raise UNDEF))
-  (presentation g);
+  (SOME g);
 
 
 (* proof transitions *)
@@ -641,8 +632,8 @@
 
 fun app int (tr as Transition {name, markers, trans, ...}) =
   setmp_thread_position tr
-    (Timing.protocol (name_of tr) (pos_of tr) (apply_trans int name markers trans)
-      ##> Option.map (fn UNDEF => ERROR (type_error tr) | exn => exn)
+    (Timing.protocol (name_of tr) (pos_of tr) (apply_capture int name markers trans)
+      ##> Option.map (fn Runtime.UNDEF => ERROR (type_error tr) | exn => exn)
       #> show_state);
 
 in
@@ -790,8 +781,7 @@
 
             val future_proof =
               Proof.future_proof (fn state =>
-                Execution.fork
-                  {name = "Toplevel.future_proof", pos = pos_of head_tr, pri = ~1}
+                Execution.fork {name = "Toplevel.future_proof", pos = pos_of head_tr, pri = ~1}
                   (fn () =>
                     let
                       val State ((Proof (prf, (_, orig_gthy)), _), prev_thy) = st';
--- a/src/Pure/ML/ml_compiler.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/ML/ml_compiler.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -53,7 +53,7 @@
     val reports_enabled =
       (case Context.get_generic_context () of
         SOME context => Context_Position.reports_enabled_generic context
-      | NONE => Context_Position.pide_reports ());
+      | NONE => Context_Position.reports_enabled0 ());
     fun is_reported pos = reports_enabled andalso Position.is_reported pos;
 
 
--- a/src/Pure/PIDE/command.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/command.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -7,7 +7,6 @@
 signature COMMAND =
 sig
   type blob = {file_node: string, src_path: Path.T, content: (SHA1.digest * string list) option}
-  val read_file: Path.T -> Position.T -> bool -> Path.T -> Token.file
   val read_thy: Toplevel.state -> theory
   val read: Keyword.keywords -> theory -> Path.T-> (unit -> theory) ->
     blob Exn.result list * int -> Token.T list -> Toplevel.transition
@@ -57,38 +56,6 @@
 
 type blob = {file_node: string, src_path: Path.T, content: (SHA1.digest * string list) option};
 
-fun read_file_node file_node master_dir pos delimited src_path =
-  let
-    val _ =
-      if Context_Position.pide_reports ()
-      then Position.report pos (Markup.language_path delimited) else ();
-
-    fun read_local () =
-      let
-        val path = File.check_file (File.full_path master_dir src_path);
-        val text = File.read path;
-        val file_pos = Path.position path;
-      in (text, file_pos) end;
-
-    fun read_remote () =
-      let
-        val text = Bytes.content (Isabelle_System.download file_node);
-        val file_pos = Position.file file_node;
-      in (text, file_pos) end;
-
-    val (text, file_pos) =
-      (case try Url.explode file_node of
-        NONE => read_local ()
-      | SOME (Url.File _) => read_local ()
-      | _ => read_remote ());
-
-    val lines = split_lines text;
-    val digest = SHA1.digest text;
-  in {src_path = src_path, lines = lines, digest = digest, pos = Position.copy_id pos file_pos} end
-  handle ERROR msg => error (msg ^ Position.here pos);
-
-val read_file = read_file_node "";
-
 local
 
 fun blob_file src_path lines digest file_node =
@@ -110,12 +77,14 @@
             val pos = Input.pos_of source;
             val delimited = Input.is_delimited source;
 
-            fun make_file (Exn.Res {file_node, src_path, content = NONE}) =
-                  Exn.interruptible_capture (fn () =>
-                    read_file_node file_node master_dir pos delimited src_path) ()
-              | make_file (Exn.Res {file_node, src_path, content = SOME (digest, lines)}) =
-                  (Position.report pos (Markup.language_path delimited);
-                    Exn.Res (blob_file src_path lines digest file_node))
+            fun make_file (Exn.Res {file_node, src_path, content}) =
+                  let val _ = Position.report pos (Markup.language_path delimited) in
+                    case content of
+                      NONE =>
+                        Exn.interruptible_capture (fn () =>
+                          Resources.read_file_node file_node master_dir (src_path, pos)) ()
+                    | SOME (digest, lines) => Exn.Res (blob_file src_path lines digest file_node)
+                  end
               | make_file (Exn.Exn e) = Exn.Exn e;
             val files = map make_file blobs;
           in
--- a/src/Pure/PIDE/command.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/command.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -450,7 +450,7 @@
           for { ((import_name, pos), s) <- imports_pos zip raw_imports if !can_import(import_name) }
           yield {
             val completion =
-              if (Thy_Header.is_base_name(s)) resources.complete_import_name(node_name, s) else Nil
+              if (Url.is_base_name(s)) resources.complete_import_name(node_name, s) else Nil
             "Bad theory import " +
               Markup.Path(import_name.node).markup(quote(import_name.toString)) +
               Position.here(pos) + Completion.report_theories(pos, completion)
@@ -464,7 +464,7 @@
           loaded_files.files.map(file =>
             (Exn.capture {
               val src_path = Path.explode(file)
-              val name = Document.Node.Name(resources.append(node_name.master_dir, src_path))
+              val name = Document.Node.Name(resources.append_path(node_name.master_dir, src_path))
               val content = get_blob(name).map(blob => (blob.bytes.sha1_digest, blob.chunk))
               Blob(name, src_path, content)
             }).user_error)
--- a/src/Pure/PIDE/command_span.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/command_span.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -21,7 +21,15 @@
 
   abstract class Load_Command(val name: String, val here: Scala_Project.Here)
   extends Isabelle_System.Service {
-    override def toString: String = name
+    override def toString: String = print("")
+
+    def print(body: String): String =
+      if (body.nonEmpty) "Load_Command(" + body + ")"
+      else if (name.nonEmpty) print("name = " + quote(name))
+      else "Load_Command"
+
+    def print_extensions: String =
+      print("name = " + quote(name) + ", extensions = " + commas_quote(extensions))
 
     def position: Position.T = here.position
 
--- a/src/Pure/PIDE/document.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/document.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -91,8 +91,9 @@
     def bad_header(msg: String): Header = Header(errors = List(msg))
 
     object Name {
-      def apply(node: String, master_dir: String = "", theory: String = ""): Name =
-        new Name(node, master_dir, theory)
+      def apply(node: String, theory: String = ""): Name = new Name(node, theory)
+
+      def loaded_theory(theory: String): Document.Node.Name = Name(theory, theory = theory)
 
       val empty: Name = Name("")
 
@@ -103,10 +104,10 @@
       type Graph[A] = isabelle.Graph[Node.Name, A]
 
       def make_graph[A](entries: List[((Name, A), List[Name])]): Graph[A] =
-        Graph.make(entries, symmetric = true)(Ordering)
+        Graph.make(entries, converse = true)(Ordering)
     }
 
-    final class Name private(val node: String, val master_dir: String, val theory: String) {
+    final class Name private(val node: String, val theory: String) {
       override def hashCode: Int = node.hashCode
       override def equals(that: Any): Boolean =
         that match {
@@ -114,7 +115,11 @@
           case _ => false
         }
 
+      def file_name: String = Url.get_base_name(node).getOrElse("")
+
       def path: Path = Path.explode(File.standard_path(node))
+
+      def master_dir: String = Url.strip_base_name(node).getOrElse("")
       def master_dir_path: Path = Path.explode(File.standard_path(master_dir))
 
       def is_theory: Boolean = theory.nonEmpty
@@ -123,16 +128,11 @@
 
       override def toString: String = if (is_theory) theory else node
 
-      def map(f: String => String): Name =
-        new Name(f(node), f(master_dir), theory)
-
       def json: JSON.Object.T =
         JSON.Object("node_name" -> node, "theory_name" -> theory)
     }
 
     sealed case class Entry(name: Node.Name, header: Node.Header) {
-      def map(f: String => String): Entry = copy(name = name.map(f))
-
       override def toString: String = name.toString
     }
 
@@ -796,7 +796,9 @@
     def node_required: Boolean
 
     def get_blob: Option[Blob]
-    def bibtex_entries: List[Text.Info[String]]
+
+    def untyped_data: AnyRef
+    def get_data[C](c: Class[C]): Option[C] = Library.as_subclass(c)(untyped_data)
 
     def node_edits(
       node_header: Node.Header,
--- a/src/Pure/PIDE/editor.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/editor.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -8,12 +8,24 @@
 
 
 abstract class Editor[Context] {
-  /* session */
+  /* PIDE session and document model */
 
   def session: Session
   def flush(): Unit
   def invoke(): Unit
 
+  def get_models(): Iterable[Document.Model]
+
+
+  /* bibtex */
+
+  def bibtex_entries_iterator(): Iterator[Text.Info[(String, Document.Model)]] =
+    Bibtex.Entries.iterator(get_models())
+
+  def bibtex_completion(history: Completion.History, rendering: Rendering, caret: Text.Offset)
+      : Option[Completion.Result] =
+    Bibtex.completion(history, rendering, caret, get_models())
+
 
   /* document editor */
 
--- a/src/Pure/PIDE/headless.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/headless.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -306,7 +306,7 @@
       val dep_theories_set = dep_theories.toSet
       val dep_files =
         for (path <- dependencies.loaded_files)
-          yield Document.Node.Name(resources.append("", path))
+          yield Document.Node.Name(resources.append_path("", path))
 
       val use_theories_state = {
         val dep_graph = dependencies.theory_graph
--- a/src/Pure/PIDE/line.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/line.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -34,7 +34,9 @@
   sealed case class Position(line: Int = 0, column: Int = 0) {
     def line1: Int = line + 1
     def column1: Int = column + 1
-    def print: String = line1.toString + ":" + column1.toString
+    def print: String =
+      if (column == 0) line1.toString
+      else line1.toString + ":" + column1.toString
 
     def compare(that: Position): Int =
       line compare that.line match {
@@ -48,7 +50,7 @@
         val lines = logical_lines(text)
         val l = line + lines.length - 1
         val c = (if (l == line) column else 0) + Library.trim_line(lines.last).length
-        Position(l, c)
+        Position(line = l, column = c)
       }
   }
 
@@ -61,8 +63,9 @@
   }
 
   sealed case class Range(start: Position, stop: Position) {
-    if (start.compare(stop) > 0)
+    if (start.compare(stop) > 0) {
       error("Bad line range: " + start.print + ".." + stop.print)
+    }
 
     def print: String =
       if (start == stop) start.print
@@ -73,7 +76,7 @@
   /* positions within document node */
 
   object Node_Position {
-    def offside(name: String): Node_Position = Node_Position(name, Position.offside)
+    def offside(name: String): Node_Position = Node_Position(name, pos = Position.offside)
   }
 
   sealed case class Node_Position(name: String, pos: Position = Position.zero) {
@@ -134,11 +137,14 @@
     def position(text_offset: Text.Offset): Position = {
       @tailrec def move(i: Text.Offset, lines_count: Int, lines_rest: List[Line]): Position = {
         lines_rest match {
-          case Nil => require(i == 0, "bad Line.position.move"); Position(lines_count)
+          case Nil =>
+            require(i == 0, "bad Line.position.move")
+            Position(line = lines_count)
           case line :: ls =>
             val n = line.text.length
-            if (ls.isEmpty || i <= n)
-              Position(lines_count).advance(line.text.substring(n - i))
+            if (ls.isEmpty || i <= n) {
+              Position(line = lines_count).advance(line.text.substring(n - i))
+            }
             else move(i - (n + 1), lines_count + 1, ls)
         }
       }
--- a/src/Pure/PIDE/protocol.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/protocol.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -25,10 +25,8 @@
   object Loading_Theory {
     def unapply(props: Properties.T): Option[(Document.Node.Name, Document_ID.Exec)] =
       (props, props, props) match {
-        case (Markup.Name(name), Position.File(file), Position.Id(id))
-        if Path.is_wellformed(file) =>
-          val master_dir = Path.explode(file).dir.implode
-          Some((Document.Node.Name(file, master_dir, name), id))
+        case (Markup.Name(theory), Position.File(file), Position.Id(id))
+        if Path.is_wellformed(file) => Some((Document.Node.Name(file, theory = theory), id))
         case _ => None
       }
   }
--- a/src/Pure/PIDE/rendering.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/rendering.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -348,7 +348,7 @@
           if (text == "" || text.endsWith("/")) (path, "")
           else (path.dir, path.file_name)
 
-        val directory = new JFile(session.resources.append(snapshot.node_name.master_dir, dir))
+        val directory = new JFile(session.resources.append_path(snapshot.node_name.master_dir, dir))
         val files = directory.listFiles
         if (files == null) Nil
         else {
@@ -616,7 +616,7 @@
   }
 
   def perhaps_append_file(node_name: Document.Node.Name, name: String): String =
-    if (Path.is_valid(name)) session.resources.append(node_name.master_dir, Path.explode(name))
+    if (Path.is_valid(name)) session.resources.append_path(node_name.master_dir, Path.explode(name))
     else name
 
   def tooltips(elements: Markup.Elements, range: Text.Range): Option[Text.Info[List[XML.Tree]]] = {
--- a/src/Pure/PIDE/resources.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/resources.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -37,10 +37,15 @@
   val check_thy: Path.T -> string ->
    {master: Path.T * SHA1.digest, text: string, theory_pos: Position.T,
     imports: (string * Position.T) list, keywords: Thy_Header.keywords}
+  val read_file_node: string -> Path.T -> Path.T * Position.T -> Token.file
+  val read_file: Path.T -> Path.T * Position.T -> Token.file
+  val parsed_files: (Path.T -> Path.T list) ->
+    Token.file Exn.result list * Input.source -> theory -> Token.file list
   val parse_files: (Path.T -> Path.T list) -> (theory -> Token.file list) parser
   val parse_file: (theory -> Token.file) parser
   val provide: Path.T * SHA1.digest -> theory -> theory
   val provide_file: Token.file -> theory -> theory
+  val provide_file': Token.file -> theory -> Token.file * theory
   val provide_parse_files: (Path.T -> Path.T list) -> (theory -> Token.file list * theory) parser
   val provide_parse_file: (theory -> Token.file * theory) parser
   val loaded_files_current: theory -> bool
@@ -326,24 +331,57 @@
   end;
 
 
+(* read_file *)
+
+fun read_file_node file_node master_dir (src_path, pos) =
+  let
+    fun read_local () =
+      let
+        val path = File.check_file (File.full_path master_dir src_path);
+        val text = File.read path;
+        val file_pos = Path.position path;
+      in (text, file_pos) end;
+
+    fun read_remote () =
+      let
+        val text = Bytes.content (Isabelle_System.download file_node);
+        val file_pos = Position.file file_node;
+      in (text, file_pos) end;
+
+    val (text, file_pos) =
+      (case try Url.explode file_node of
+        NONE => read_local ()
+      | SOME (Url.File _) => read_local ()
+      | _ => read_remote ());
+
+    val lines = split_lines text;
+    val digest = SHA1.digest text;
+  in {src_path = src_path, lines = lines, digest = digest, pos = Position.copy_id pos file_pos} end
+  handle ERROR msg => error (msg ^ Position.here pos);
+
+val read_file = read_file_node "";
+
+
 (* load files *)
 
+fun parsed_files make_paths (files, source) thy =
+  if null files then
+    let
+      val master_dir = master_directory thy;
+      val name = Input.string_of source;
+      val pos = Input.pos_of source;
+      val delimited = Input.is_delimited source;
+      val src_paths = make_paths (Path.explode name);
+      val reports =
+        src_paths |> maps (fn src_path =>
+          [(pos, Markup.path (Path.implode_symbolic (master_dir + src_path))),
+           (pos, Markup.language_path delimited)]);
+      val _ = Position.reports reports;
+    in map (read_file master_dir o rpair pos) src_paths end
+  else map Exn.release files;
+
 fun parse_files make_paths =
-  Scan.ahead Parse.not_eof -- Parse.path_input >> (fn (tok, source) => fn thy =>
-    (case Token.get_files tok of
-      [] =>
-        let
-          val master_dir = master_directory thy;
-          val name = Input.string_of source;
-          val pos = Input.pos_of source;
-          val delimited = Input.is_delimited source;
-          val src_paths = make_paths (Path.explode name);
-          val reports =
-            src_paths |> map (fn src_path =>
-              (pos, Markup.path (Path.implode_symbolic (master_dir + src_path))));
-          val _ = Position.reports reports;
-        in map (Command.read_file master_dir pos delimited) src_paths end
-    | files => map Exn.release files));
+  (Scan.ahead Parse.not_eof >> Token.get_files) -- Parse.path_input >> parsed_files make_paths;
 
 val parse_file = parse_files single >> (fn f => f #> the_single);
 
@@ -355,13 +393,10 @@
     else (master_dir, imports, (src_path, id) :: provided));
 
 fun provide_file (file: Token.file) = provide (#src_path file, #digest file);
+fun provide_file' file thy = (file, provide_file file thy);
 
 fun provide_parse_files make_paths =
-  parse_files make_paths >> (fn files => fn thy =>
-    let
-      val fs = files thy;
-      val thy' = fold (fn {src_path, digest, ...} => provide (src_path, digest)) fs thy;
-    in (fs, thy') end);
+  parse_files make_paths >> (fn files => fn thy => fold_map provide_file' (files thy) thy);
 
 val provide_parse_file = provide_parse_files single >> (fn f => f #>> the_single);
 
@@ -383,7 +418,7 @@
 
 (* formal check *)
 
-fun formal_check check_file ctxt opt_dir source =
+fun formal_check (check: Path.T -> Path.T) ctxt opt_dir source =
   let
     val name = Input.string_of source;
     val pos = Input.pos_of source;
@@ -399,7 +434,7 @@
     val path = dir + Path.explode name handle ERROR msg => err msg;
     val _ = Path.expand path handle ERROR msg => err msg;
     val _ = Context_Position.report ctxt pos (Markup.path (Path.implode_symbolic path));
-    val _ : Path.T = check_file path handle ERROR msg => err msg;
+    val _ = check path handle ERROR msg => err msg;
   in path end;
 
 val check_path = formal_check I;
--- a/src/Pure/PIDE/resources.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/resources.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -72,17 +72,8 @@
 
   def migrate_name(name: Document.Node.Name): Document.Node.Name = name
 
-  def append(dir: String, source_path: Path): String =
-    (Path.explode(dir) + source_path).expand.implode
-
-  def file_node(file: Path, dir: String = "", theory: String = ""): Document.Node.Name = {
-    val node = append(dir, file)
-    val master_dir = append(dir, file.dir)
-    Document.Node.Name(node, master_dir, theory)
-  }
-
-  def loaded_theory_node(theory: String): Document.Node.Name =
-    Document.Node.Name(theory, "", theory)
+  def append_path(prefix: String, source_path: Path): String =
+    (Path.explode(prefix) + source_path).expand.implode
 
 
   /* source files of Isabelle/ML bootstrap */
@@ -141,7 +132,7 @@
     for {
       (name, theory) <- Thy_Header.ml_roots
       path = (pure_dir + Path.explode(name)).expand
-      node_name = Document.Node.Name(path.implode, path.dir.implode, theory)
+      node_name = Document.Node.Name(path.implode, theory = theory)
       file <- loaded_files(syntax, node_name, load_commands(syntax, node_name)())
     } yield file
   }
@@ -164,24 +155,24 @@
         case Some(info) => info.dirs
         case None => Nil
       }
-    dirs.collectFirst({
-      case dir if (dir + thy_file).is_file => file_node(dir + thy_file, theory = theory) })
+    dirs.collectFirst { case dir if (dir + thy_file).is_file =>
+      Document.Node.Name(append_path("", dir + thy_file), theory = theory) }
   }
 
-  def import_name(qualifier: String, dir: String, s: String): Document.Node.Name = {
+  def import_name(qualifier: String, prefix: String, s: String): Document.Node.Name = {
     val theory = theory_name(qualifier, Thy_Header.import_name(s))
     val literal_import =
       literal_theory(theory) && qualifier != sessions_structure.theory_qualifier(theory)
-    if (literal_import && !Thy_Header.is_base_name(s)) {
+    if (literal_import && !Url.is_base_name(s)) {
       error("Bad import of theory from other session via file-path: " + quote(s))
     }
-    if (session_base.loaded_theory(theory)) loaded_theory_node(theory)
+    if (session_base.loaded_theory(theory)) Document.Node.Name.loaded_theory(theory)
     else {
       find_theory_node(theory) match {
         case Some(node_name) => node_name
         case None =>
-          if (Thy_Header.is_base_name(s) && literal_theory(s)) loaded_theory_node(theory)
-          else file_node(Path.explode(s).thy, dir = dir, theory = theory)
+          if (Url.is_base_name(s) && literal_theory(s)) Document.Node.Name.loaded_theory(theory)
+          else Document.Node.Name(append_path(prefix, Path.explode(s).thy), theory = theory)
       }
     }
   }
@@ -210,7 +201,7 @@
     (for {
       (session, (info, _))  <- sessions_structure.imports_graph.iterator
       dir <- (if (session == context_session) context_dir.toList else info.dirs).iterator
-      theory <- Thy_Header.try_read_dir(dir).iterator
+      theory <- Thy_Header.list_thy_names(dir).iterator
       if Completion.completed(s)(theory)
     } yield {
       if (session == context_session || global_theory(theory)) theory
--- a/src/Pure/PIDE/session.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/PIDE/session.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -356,10 +356,7 @@
   private val protocol_handlers = Protocol_Handlers.init(session)
 
   def get_protocol_handler[C <: Session.Protocol_Handler](c: Class[C]): Option[C] =
-    protocol_handlers.get(c.getName) match {
-      case Some(h) if Library.is_subclass(h.getClass, c) => Some(h.asInstanceOf[C])
-      case _ => None
-    }
+    protocol_handlers.get(c.getName).flatMap(Library.as_subclass(c))
 
   def init_protocol_handler(handler: Session.Protocol_Handler): Unit =
     protocol_handlers.init(handler)
@@ -490,7 +487,7 @@
               case Protocol.Export(args)
               if args.id.isDefined && Value.Long.unapply(args.id.get).isDefined =>
                 val id = Value.Long.unapply(args.id.get).get
-                val entry = Export.make_entry(Sessions.DRAFT, args, msg.chunk, cache)
+                val entry = Export.Entry.make(Sessions.DRAFT, args, msg.chunk, cache)
                 change_command(_.add_export(id, (args.serial, entry)))
 
               case Protocol.Loading_Theory(node_name, id) =>
@@ -566,7 +563,7 @@
                 val snapshot = global_state.change_result(_.end_theory(id))
                 finished_theories.post(snapshot)
               }
-              file_formats.stop_session
+              file_formats.stop_session()
               phase = Session.Terminated(result)
               prover.reset()
 
--- a/src/Pure/ROOT.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/ROOT.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -311,9 +311,9 @@
 ML_file "Thy/document_antiquotations.ML";
 ML_file "General/graph_display.ML";
 ML_file "pure_syn.ML";
+ML_file "PIDE/resources.ML";
 ML_file "PIDE/command.ML";
 ML_file "PIDE/query_operation.ML";
-ML_file "PIDE/resources.ML";
 ML_file "Thy/thy_info.ML";
 ML_file "thm_deps.ML";
 ML_file "Thy/export_theory.ML";
@@ -366,4 +366,3 @@
 ML_file "Tools/jedit.ML";
 ML_file "Tools/ghc.ML";
 ML_file "Tools/generated_files.ML";
-
--- a/src/Pure/System/isabelle_process.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/System/isabelle_process.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -118,7 +118,7 @@
         in message name pos_props [XML.blob ss] end;
 
     fun report_message ss =
-      if Context_Position.pide_reports ()
+      if Context_Position.reports_enabled0 ()
       then standard_message [] Markup.reportN ss else ();
 
     val serial_props = Markup.serial_properties o serial;
--- a/src/Pure/Thy/bibtex.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/bibtex.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -21,15 +21,18 @@
     val format_name: String = "bibtex"
     val file_ext: String = "bib"
 
+    override def parse_data(name: String, text: String): Bibtex.Entries =
+      Bibtex.Entries.parse(text, file_pos = name)
+
     override def theory_suffix: String = "bibtex_file"
     override def theory_content(name: String): String =
-      """theory "bib" imports Pure begin bibtex_file """ +
-        Outer_Syntax.quote_string(name) + """ end"""
+      """theory "bib" imports Pure begin bibtex_file "." end"""
+    override def theory_excluded(name: String): Boolean = name == "bib"
 
     override def html_document(snapshot: Document.Snapshot): Option[Browser_Info.HTML_Document] = {
       val name = snapshot.node_name
       if (detect(name.node)) {
-        val title = "Bibliography " + quote(snapshot.node_name.path.file_name)
+        val title = "Bibliography " + quote(snapshot.node_name.file_name)
         val content =
           Isabelle_System.with_tmp_file("bib", "bib") { bib =>
             File.write(bib, snapshot.source)
@@ -154,35 +157,64 @@
 
   /* entries */
 
-  def entries(text: String): List[Text.Info[String]] = {
-    val result = new mutable.ListBuffer[Text.Info[String]]
-    var offset = 0
-    for (chunk <- Bibtex.parse(text)) {
-      val end_offset = offset + chunk.source.length
-      if (chunk.name != "" && !chunk.is_command)
-        result += Text.Info(Text.Range(offset, end_offset), chunk.name)
-      offset = end_offset
+  object Entries {
+    val empty: Entries = apply(Nil)
+
+    def apply(entries: List[Text.Info[String]], errors: List[String] = Nil): Entries =
+      new Entries(entries, errors)
+
+    def parse(text: String, file_pos: String = ""): Entries = {
+      val entries = new mutable.ListBuffer[Text.Info[String]]
+      var offset = 0
+      var line = 1
+      var err_line = 0
+
+      try {
+        for (chunk <- Bibtex.parse(text)) {
+          val end_offset = offset + chunk.source.length
+          if (chunk.name != "" && !chunk.is_command) {
+            entries += Text.Info(Text.Range(offset, end_offset), chunk.name)
+          }
+          if (chunk.is_malformed && err_line == 0) { err_line = line }
+          offset = end_offset
+          line += chunk.source.count(_ == '\n')
+        }
+
+        val err_pos =
+          if (err_line == 0 || file_pos.isEmpty) Position.none
+          else Position.Line_File(err_line, file_pos)
+        val errors =
+          if (err_line == 0) Nil
+          else List("Malformed bibtex file" + Position.here(err_pos))
+
+        apply(entries.toList, errors = errors)
+      }
+      catch { case ERROR(msg) => apply(Nil, errors = List(msg)) }
     }
-    result.toList
+
+    def iterator[A <: Document.Model](models: Iterable[A]): Iterator[Text.Info[(String, A)]] =
+      for {
+        model <- models.iterator
+        bibtex_entries <- model.get_data(classOf[Entries]).iterator
+        info <- bibtex_entries.entries.iterator
+      } yield info.map((_, model))
   }
 
-  def entries_iterator[A, B <: Document.Model](
-    models: Map[A, B]
-  ): Iterator[Text.Info[(String, B)]] = {
-    for {
-      (_, model) <- models.iterator
-      info <- model.bibtex_entries.iterator
-    } yield info.map((_, model))
+  final class Entries private(val entries: List[Text.Info[String]], val errors: List[String]) {
+    override def toString: String = "Bibtex.Entries(" + entries.length + ")"
+
+    def ::: (other: Entries): Entries =
+      new Entries(entries ::: other.entries, errors ::: other.errors)
   }
 
 
   /* completion */
 
-  def completion[A, B <: Document.Model](
+  def completion[A <: Document.Model](
     history: Completion.History,
     rendering: Rendering,
     caret: Text.Offset,
-    models: Map[A, B]
+    models: Iterable[A]
   ): Option[Completion.Result] = {
     for {
       Text.Info(r, name) <- rendering.citations(rendering.before_caret_range(caret)).headOption
@@ -193,7 +225,7 @@
 
       entries =
         (for {
-          Text.Info(_, (entry, _)) <- entries_iterator(models)
+          Text.Info(_, (entry, _)) <- Entries.iterator(models)
           if entry.toLowerCase.containsSlice(name1.toLowerCase) && entry != original1
         } yield entry).toList
       if entries.nonEmpty
@@ -382,7 +414,7 @@
       }
 
     def is_ignored: Boolean = kind == "" && tokens.forall(_.is_ignored)
-    def is_malformed: Boolean = kind == "" || tokens.exists(_.is_malformed)
+    def is_malformed: Boolean = tokens.exists(_.is_malformed)
     def is_command: Boolean = Bibtex.is_command(kind) && name != "" && content.isDefined
     def is_entry: Boolean = Bibtex.is_entry(kind) && name != "" && content.isDefined
   }
--- a/src/Pure/Thy/browser_info.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/browser_info.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -264,7 +264,7 @@
 
       val name = snapshot.node_name
       if (plain_text) {
-        val title = "File " + Symbol.cartouche_decoded(name.path.file_name)
+        val title = "File " + Symbol.cartouche_decoded(name.file_name)
         val body = HTML.text(snapshot.source)
         html_document(title, body, fonts_css)
       }
@@ -272,7 +272,7 @@
         Resources.html_document(snapshot) getOrElse {
           val title =
             if (name.is_theory) "Theory " + quote(name.theory_base_name)
-            else "File " + Symbol.cartouche_decoded(name.path.file_name)
+            else "File " + Symbol.cartouche_decoded(name.file_name)
           val xml = snapshot.xml_markup(elements = elements.html)
           val body = Node_Context.empty.make_html(elements, xml)
           html_document(title, body, fonts_css)
--- a/src/Pure/Thy/document_build.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/document_build.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -140,7 +140,7 @@
     def content: File.Content_XML = File.content(Path.basic(tex_name(name)), body)
     def file_pos: String = name.path.implode_symbolic
     def write(latex_output: Latex.Output, dir: Path): Unit =
-      content.output(latex_output(_, file_pos = file_pos)).write(dir)
+      content.output(latex_output.make(_, file_pos = file_pos)).write(dir)
   }
 
   def context(
--- a/src/Pure/Thy/export.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/export.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -88,7 +88,7 @@
   def read_theory_names(db: SQL.Database, session_name: String): List[String] = {
     val select =
       Data.table.select(List(Data.theory_name), Data.where_equal(session_name), distinct = true) +
-      " ORDER BY " + Data.theory_name
+      SQL.order_by(List(Data.theory_name))
     db.using_statement(select)(stmt =>
       stmt.execute_query().iterator(_.string(Data.theory_name)).toList)
   }
@@ -96,7 +96,7 @@
   def read_entry_names(db: SQL.Database, session_name: String): List[Entry_Name] = {
     val select =
       Data.table.select(List(Data.theory_name, Data.name), Data.where_equal(session_name)) +
-      " ORDER BY " + Data.theory_name + ", " + Data.name
+      SQL.order_by(List(Data.theory_name, Data.name))
     db.using_statement(select)(stmt =>
       stmt.execute_query().iterator(res =>
         Entry_Name(session = session_name,
@@ -107,13 +107,36 @@
   def message(msg: String, theory_name: String, name: String): String =
     msg + " " + quote(name) + " for theory " + quote(theory_name)
 
-  def empty_entry(theory_name: String, name: String): Entry =
-    Entry(Entry_Name(theory = theory_name, name = name),
-      false, Future.value(false, Bytes.empty), XML.Cache.none)
+  object Entry {
+    def apply(
+      entry_name: Entry_Name,
+      executable: Boolean,
+      body: Future[(Boolean, Bytes)],
+      cache: XML.Cache
+    ): Entry = new Entry(entry_name, executable, body, cache)
+
+    def empty(theory_name: String, name: String): Entry =
+      Entry(Entry_Name(theory = theory_name, name = name),
+        false, Future.value(false, Bytes.empty), XML.Cache.none)
 
-  sealed case class Entry(
-    entry_name: Entry_Name,
-    executable: Boolean,
+    def make(
+      session_name: String,
+      args: Protocol.Export.Args,
+      bytes: Bytes,
+      cache: XML.Cache
+    ): Entry = {
+      val body =
+        if (args.compress) Future.fork(bytes.maybe_compress(cache = cache.compress))
+        else Future.value((false, bytes))
+      val entry_name =
+        Entry_Name(session = session_name, theory = args.theory_name, name = args.name)
+      Entry(entry_name, args.executable, body, cache)
+    }
+  }
+
+  final class Entry private(
+    val entry_name: Entry_Name,
+    val executable: Boolean,
     body: Future[(Boolean, Bytes)],
     cache: XML.Cache
   ) {
@@ -122,6 +145,9 @@
     def name: String = entry_name.name
     override def toString: String = name
 
+    def is_finished: Boolean = body.is_finished
+    def cancel(): Unit = body.cancel()
+
     def compound_name: String = entry_name.compound_name
 
     def name_has_prefix(s: String): Boolean = name.startsWith(s)
@@ -130,25 +156,24 @@
     def name_extends(elems: List[String]): Boolean =
       name_elems.startsWith(elems) && name_elems != elems
 
-    def text: String = uncompressed.text
-
-    def uncompressed: Bytes = {
-      val (compressed, bytes) = body.join
-      if (compressed) bytes.uncompress(cache = cache.compress) else bytes
+    def bytes: Bytes = {
+      val (compressed, bs) = body.join
+      if (compressed) bs.uncompress(cache = cache.compress) else bs
     }
 
-    def uncompressed_yxml: XML.Body =
-      YXML.parse_body(UTF8.decode_permissive(uncompressed), cache = cache)
+    def text: String = bytes.text
+
+    def yxml: XML.Body = YXML.parse_body(UTF8.decode_permissive(bytes), cache = cache)
 
     def write(db: SQL.Database): Unit = {
-      val (compressed, bytes) = body.join
+      val (compressed, bs) = body.join
       db.using_statement(Data.table.insert()) { stmt =>
         stmt.string(1) = session_name
         stmt.string(2) = theory_name
         stmt.string(3) = name
         stmt.bool(4) = executable
         stmt.bool(5) = compressed
-        stmt.bytes(6) = bytes
+        stmt.bytes(6) = bs
         stmt.execute()
       }
     }
@@ -176,19 +201,6 @@
     (entry_name: Entry_Name) => regs.exists(_.matches(entry_name.compound_name))
   }
 
-  def make_entry(
-    session_name: String,
-    args: Protocol.Export.Args,
-    bytes: Bytes,
-    cache: XML.Cache
-  ): Entry = {
-    val body =
-      if (args.compress) Future.fork(bytes.maybe_compress(cache = cache.compress))
-      else Future.value((false, bytes))
-    val entry_name = Entry_Name(session = session_name, theory = args.theory_name, name = args.name)
-    Entry(entry_name, args.executable, body, cache)
-  }
-
 
   /* database consumer thread */
 
@@ -200,7 +212,7 @@
 
     private val consumer =
       Consumer_Thread.fork_bulk[(Entry, Boolean)](name = "export")(
-        bulk = { case (entry, _) => entry.body.is_finished },
+        bulk = { case (entry, _) => entry.is_finished },
         consume =
           { (args: List[(Entry, Boolean)]) =>
             val results =
@@ -208,7 +220,7 @@
                 for ((entry, strict) <- args)
                 yield {
                   if (progress.stopped) {
-                    entry.body.cancel()
+                    entry.cancel()
                     Exn.Res(())
                   }
                   else if (entry.entry_name.readable(db)) {
@@ -227,7 +239,7 @@
     def make_entry(session_name: String, args0: Protocol.Export.Args, body: Bytes): Unit = {
       if (!progress.stopped && !body.is_empty) {
         val args = if (db.is_server) args0.copy(compress = false) else args0
-        consumer.send(Export.make_entry(session_name, args, body, cache) -> args.strict)
+        consumer.send(Entry.make(session_name, args, body, cache) -> args.strict)
       }
     }
 
@@ -403,7 +415,7 @@
 
     def apply(theory: String, name: String, permissive: Boolean = false): Entry =
       get(theory, name) match {
-        case None if permissive => empty_entry(theory, name)
+        case None if permissive => Entry.empty(theory, name)
         case None => error("Missing export entry " + quote(compound_name(theory, name)))
         case Some(entry) => entry
       }
@@ -411,6 +423,24 @@
     def theory(theory: String, other_cache: Option[Term.Cache] = None): Theory_Context =
       new Theory_Context(session_context, theory, other_cache)
 
+    def node_source(name: Document.Node.Name): String = {
+      def snapshot_source: Option[String] =
+        for {
+          snapshot <- document_snapshot
+          node = snapshot.get_node(name)
+          text = node.source if text.nonEmpty
+        } yield text
+
+      val store = database_context.store
+      def db_source: Option[String] =
+        (for {
+          database <- db_hierarchy.iterator
+          file <- store.read_sources(database.db, database.session, name = name.node).iterator
+        } yield file.text).nextOption()
+
+      snapshot_source orElse db_source getOrElse ""
+    }
+
     def classpath(): List[File.Content] = {
       (for {
         session <- session_stack.iterator
@@ -420,7 +450,7 @@
         entry_name <- entry_names(session = session).iterator
         if matcher(entry_name)
         entry <- get(entry_name.theory, entry_name.name).iterator
-      } yield File.content(entry.entry_name.make_path(), entry.uncompressed)).toList
+      } yield File.content(entry.entry_name.make_path(), entry.bytes)).toList
     }
 
     override def toString: String =
@@ -438,9 +468,9 @@
     def apply(name: String, permissive: Boolean = false): Entry =
       session_context.apply(theory, name, permissive = permissive)
 
-    def uncompressed_yxml(name: String): XML.Body =
+    def yxml(name: String): XML.Body =
       get(name) match {
-        case Some(entry) => entry.uncompressed_yxml
+        case Some(entry) => entry.yxml
         case None => Nil
       }
 
@@ -492,7 +522,7 @@
           val path = export_dir + entry_name.make_path(prune = export_prune)
           progress.echo("export " + path + (if (entry.executable) " (executable)" else ""))
           Isabelle_System.make_directory(path.dir)
-          val bytes = entry.uncompressed
+          val bytes = entry.bytes
           if (!path.is_file || Bytes.read(path) != bytes) Bytes.write(path, bytes)
           File.set_executable(path, entry.executable)
         }
--- a/src/Pure/Thy/export_theory.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/export_theory.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -103,7 +103,7 @@
 
   def read_theory_parents(theory_context: Export.Theory_Context): Option[List[String]] =
     theory_context.get(Export.THEORY_PREFIX + "parents")
-      .map(entry => Library.trim_split_lines(entry.uncompressed.text))
+      .map(entry => Library.trim_split_lines(entry.text))
 
   def theory_names(
     session_context: Export.Session_Context,
@@ -233,7 +233,7 @@
         case _ => err()
       }
     }
-    theory_context.uncompressed_yxml(export_name).map(decode_entity)
+    theory_context.yxml(export_name).map(decode_entity)
   }
 
 
@@ -393,7 +393,7 @@
 
     for { entry <- theory_context.get(Export.PROOFS_PREFIX + id.serial) }
     yield {
-      val body = entry.uncompressed_yxml
+      val body = entry.yxml
       val (typargs, (args, (prop_body, proof_body))) = {
         import XML.Decode._
         import Term_XML.Decode._
@@ -539,7 +539,7 @@
   }
 
   def read_classrel(theory_context: Export.Theory_Context): List[Classrel] = {
-    val body = theory_context.uncompressed_yxml(Export.THEORY_PREFIX + "classrel")
+    val body = theory_context.yxml(Export.THEORY_PREFIX + "classrel")
     val classrel = {
       import XML.Decode._
       list(pair(decode_prop, pair(string, string)))(body)
@@ -559,7 +559,7 @@
   }
 
   def read_arities(theory_context: Export.Theory_Context): List[Arity] = {
-    val body = theory_context.uncompressed_yxml(Export.THEORY_PREFIX + "arities")
+    val body = theory_context.yxml(Export.THEORY_PREFIX + "arities")
     val arities = {
       import XML.Decode._
       import Term_XML.Decode._
@@ -577,7 +577,7 @@
   }
 
   def read_constdefs(theory_context: Export.Theory_Context): List[Constdef] = {
-    val body = theory_context.uncompressed_yxml(Export.THEORY_PREFIX + "constdefs")
+    val body = theory_context.yxml(Export.THEORY_PREFIX + "constdefs")
     val constdefs = {
       import XML.Decode._
       list(pair(string, string))(body)
@@ -606,7 +606,7 @@
   }
 
   def read_typedefs(theory_context: Export.Theory_Context): List[Typedef] = {
-    val body = theory_context.uncompressed_yxml(Export.THEORY_PREFIX + "typedefs")
+    val body = theory_context.yxml(Export.THEORY_PREFIX + "typedefs")
     val typedefs = {
       import XML.Decode._
       import Term_XML.Decode._
@@ -640,7 +640,7 @@
   }
 
   def read_datatypes(theory_context: Export.Theory_Context): List[Datatype] = {
-    val body = theory_context.uncompressed_yxml(Export.THEORY_PREFIX + "datatypes")
+    val body = theory_context.yxml(Export.THEORY_PREFIX + "datatypes")
     val datatypes = {
       import XML.Decode._
       import Term_XML.Decode._
@@ -730,7 +730,7 @@
   }
 
   def read_spec_rules(theory_context: Export.Theory_Context): List[Spec_Rule] = {
-    val body = theory_context.uncompressed_yxml(Export.THEORY_PREFIX + "spec_rules")
+    val body = theory_context.yxml(Export.THEORY_PREFIX + "spec_rules")
     val spec_rules = {
       import XML.Decode._
       import Term_XML.Decode._
@@ -753,7 +753,7 @@
   def read_others(theory_context: Export.Theory_Context): Map[String, List[Entity[Other]]] = {
     val kinds =
       theory_context.get(Export.THEORY_PREFIX + "other_kinds") match {
-        case Some(entry) => split_lines(entry.uncompressed.text)
+        case Some(entry) => split_lines(entry.text)
         case None => Nil
       }
     val other = Other()
--- a/src/Pure/Thy/file_format.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/file_format.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -8,6 +8,10 @@
 
 
 object File_Format {
+  object No_Data extends AnyRef {
+    override def toString: String = "File_Format.No_Data"
+  }
+
   sealed case class Theory_Context(name: Document.Node.Name, content: String)
 
 
@@ -24,6 +28,15 @@
     def get_theory(name: Document.Node.Name): Option[File_Format] = get(name.theory)
     def is_theory(name: Document.Node.Name): Boolean = get_theory(name).isDefined
 
+    def theory_excluded(name: String): Boolean = file_formats.exists(_.theory_excluded(name))
+
+    def parse_data(name: String, text: String): AnyRef =
+      get(name) match {
+        case Some(file_format) => file_format.parse_data(name, text)
+        case None => No_Data
+      }
+    def parse_data(name: Document.Node.Name, text: String): AnyRef = parse_data(name.node, text)
+
     def start_session(session: isabelle.Session): Session =
       new Session(file_formats.map(_.start(session)).filterNot(_ == No_Agent))
   }
@@ -38,7 +51,7 @@
     def prover_options(options: Options): Options =
       agents.foldLeft(options) { case (opts, agent) => agent.prover_options(opts) }
 
-    def stop_session: Unit = agents.foreach(_.stop())
+    def stop_session(): Unit = agents.foreach(_.stop())
   }
 
   trait Agent {
@@ -56,31 +69,35 @@
   def file_ext: String
   def detect(name: String): Boolean = name.endsWith("." + file_ext)
 
+  def parse_data(name: String, text: String): AnyRef = File_Format.No_Data
+
 
   /* implicit theory context: name and content */
 
   def theory_suffix: String = ""
   def theory_content(name: String): String = ""
+  def theory_excluded(name: String): Boolean = false
 
   def make_theory_name(
     resources: Resources,
     name: Document.Node.Name
   ): Option[Document.Node.Name] = {
     for {
-      (_, thy) <- Thy_Header.split_file_name(name.node)
+      theory <- Url.get_base_name(name.node)
       if detect(name.node) && theory_suffix.nonEmpty
     }
     yield {
-      val thy_node = resources.append(name.node, Path.explode(theory_suffix))
-      Document.Node.Name(thy_node, name.master_dir, thy)
+      val node = resources.append_path(name.node, Path.explode(theory_suffix))
+      Document.Node.Name(node, theory = theory)
     }
   }
 
   def make_theory_content(resources: Resources, thy_name: Document.Node.Name): Option[String] = {
     for {
-      (prefix, suffix) <- Thy_Header.split_file_name(thy_name.node)
+      prefix <- Url.strip_base_name(thy_name.node)
+      suffix <- Url.get_base_name(thy_name.node)
       if detect(prefix) && suffix == theory_suffix
-      (_, thy) <- Thy_Header.split_file_name(prefix)
+      thy <- Url.get_base_name(prefix)
       s <- proper_string(theory_content(thy))
     } yield s
   }
--- a/src/Pure/Thy/latex.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/latex.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -158,7 +158,7 @@
     else List("\\endinput\n", position(Markup.FILE, file_pos))
 
   class Output(val options: Options) {
-    def latex_output(latex_text: Text): String = apply(latex_text)
+    def latex_output(latex_text: Text): String = make(latex_text)
 
     def latex_macro0(name: String, optional_argument: String = ""): Text =
       XML.string("\\" + name + optional_argument)
@@ -214,7 +214,7 @@
       error("Unknown latex markup element " + quote(elem.name) + Position.here(pos) +
         ":\n" + XML.string_of_tree(elem))
 
-    def apply(latex_text: Text, file_pos: String = ""): String = {
+    def make(latex_text: Text, file_pos: String = ""): String = {
       var line = 1
       val result = new mutable.ListBuffer[String]
       val positions = new mutable.ListBuffer[String] ++= init_position(file_pos)
--- a/src/Pure/Thy/sessions.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/sessions.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -32,25 +32,83 @@
   def is_pure(name: String): Boolean = name == Thy_Header.PURE
 
   def illegal_session(name: String): Boolean = name == "" || name == DRAFT
-  def illegal_theory(name: String): Boolean = name == root_name || name == "bib"
+  def illegal_theory(name: String): Boolean =
+    name == root_name || File_Format.registry.theory_excluded(name)
 
 
   /* ROOTS file format */
 
-  class File_Format extends isabelle.File_Format {
+  class ROOTS_File_Format extends File_Format {
     val format_name: String = roots_name
     val file_ext = ""
 
     override def detect(name: String): Boolean =
-      Thy_Header.split_file_name(name) match {
-        case Some((_, file_name)) => file_name == roots_name
+      Url.get_base_name(name) match {
+        case Some(base_name) => base_name == roots_name
         case None => false
       }
 
     override def theory_suffix: String = "ROOTS_file"
     override def theory_content(name: String): String =
-      """theory "ROOTS" imports Pure begin ROOTS_file """ +
-        Outer_Syntax.quote_string(name) + """ end"""
+      """theory "ROOTS" imports Pure begin ROOTS_file "." end"""
+    override def theory_excluded(name: String): Boolean = name == "ROOTS"
+  }
+
+
+  /* source files */
+
+  sealed case class Source_File(
+    name: String,
+    digest: SHA1.Digest,
+    compressed: Boolean,
+    body: Bytes,
+    cache: Compress.Cache
+  ) {
+    override def toString: String = name
+
+    def bytes: Bytes = if (compressed) body.uncompress(cache = cache) else body
+    def text: String = bytes.text
+  }
+
+  object Sources {
+    val session_name = SQL.Column.string("session_name").make_primary_key
+    val name = SQL.Column.string("name").make_primary_key
+    val digest = SQL.Column.string("digest")
+    val compressed = SQL.Column.bool("compressed")
+    val body = SQL.Column.bytes("body")
+
+    val table =
+      SQL.Table("isabelle_sources", List(session_name, name, digest, compressed, body))
+
+    def where_equal(session_name: String, name: String = ""): SQL.Source =
+      "WHERE " + Sources.session_name.equal(session_name) +
+        (if (name == "") "" else " AND " + Sources.name.equal(name))
+
+    def load(session_base: Base, cache: Compress.Cache = Compress.Cache.none): Sources =
+      new Sources(
+        session_base.session_sources.foldLeft(Map.empty) {
+          case (sources, (path, digest)) =>
+            def err(): Nothing = error("Incoherent digest for source file: " + path)
+            val name = path.implode_symbolic
+            sources.get(name) match {
+              case Some(source_file) =>
+                if (source_file.digest == digest) sources else err()
+              case None =>
+                val bytes = Bytes.read(path)
+                if (bytes.sha1_digest == digest) {
+                  val (compressed, body) =
+                    bytes.maybe_compress(Compress.Options_Zstd(), cache = cache)
+                  val file = Source_File(name, digest, compressed, body, cache)
+                  sources + (name -> file)
+                }
+                else err()
+            }
+        })
+  }
+
+  class Sources private(rep: Map[String, Source_File]) extends Iterable[Source_File] {
+    override def toString: String = rep.values.toList.sortBy(_.name).mkString("Sources(", ", ", ")")
+    override def iterator: Iterator[Source_File] = rep.valuesIterator
   }
 
 
@@ -67,7 +125,7 @@
     document_theories: List[Document.Node.Name] = Nil,
     loaded_theories: Graph[String, Outer_Syntax] = Graph.string,  // cumulative imports
     used_theories: List[(Document.Node.Name, Options)] = Nil,  // new imports
-    load_commands: Map[Document.Node.Name, List[Command_Span.Span]] = Map.empty,
+    theory_load_commands: Map[String, List[Command_Span.Span]] = Map.empty,
     known_theories: Map[String, Document.Node.Entry] = Map.empty,
     known_loaded_files: Map[String, List[Path]] = Map.empty,
     overall_syntax: Outer_Syntax = Outer_Syntax.empty,
@@ -296,6 +354,9 @@
               try { if (inlined_files) (dependencies.load_commands, Nil) else (Nil, Nil) }
               catch { case ERROR(msg) => (Nil, List(msg)) }
 
+            val theory_load_commands =
+              (for ((name, span) <- load_commands.iterator) yield name.theory -> span).toMap
+
             val loaded_files =
               load_commands.map({ case (name, spans) => dependencies.loaded_files(name, spans) })
 
@@ -426,8 +487,10 @@
                   name <- proper_session_theories.iterator
                   name1 <- resources.find_theory_node(name.theory)
                   if name.node != name1.node
-                } yield "Incoherent theory file import:\n  " + name.path + " vs. \n  " + name1.path)
-                .toList
+                } yield {
+                  "Incoherent theory file import:\n  " + quote(name.node) +
+                    " vs. \n  " + quote(name1.node)
+                }).toList
 
               errs1 ::: errs2 ::: errs3 ::: errs4
             }
@@ -439,9 +502,7 @@
               try { Path.check_case_insensitive(session_files ::: imported_files); Nil }
               catch { case ERROR(msg) => List(msg) }
 
-            val bibtex_errors =
-              try { info.bibtex_entries; Nil }
-              catch { case ERROR(msg) => List(msg) }
+            val bibtex_errors = info.bibtex_entries.errors
 
             val base =
               Base(
@@ -451,7 +512,7 @@
                 document_theories = document_theories,
                 loaded_theories = dependencies.loaded_theories,
                 used_theories = dependencies.theories_adjunct,
-                load_commands = load_commands.toMap,
+                theory_load_commands = theory_load_commands,
                 known_theories = known_theories,
                 known_loaded_files = known_loaded_files,
                 overall_syntax = overall_syntax,
@@ -639,12 +700,14 @@
 
     def browser_info: Boolean = options.bool("browser_info")
 
-    lazy val bibtex_entries: List[Text.Info[String]] =
+    lazy val bibtex_entries: Bibtex.Entries =
       (for {
         (document_dir, file) <- document_files.iterator
         if File.is_bib(file.file_name)
-        info <- Bibtex.entries(File.read(dir + document_dir + file)).iterator
-      } yield info).toList
+      } yield {
+        val path = dir + document_dir + file
+        Bibtex.Entries.parse(File.read(path), file_pos = path.expand.implode)
+      }).foldRight(Bibtex.Entries.empty)(_ ::: _)
 
     def record_proofs: Boolean = options.int("record_proofs") >= 2
 
@@ -923,11 +986,12 @@
     def imports_topological_order: List[String] = imports_graph.topological_order
 
     def bibtex_entries: List[(String, List[String])] =
-      build_topological_order.flatMap(name =>
-        apply(name).bibtex_entries match {
+      build_topological_order.flatMap { name =>
+        apply(name).bibtex_entries.entries match {
           case Nil => None
           case entries => Some(name -> entries.map(_.info))
-        })
+        }
+      }
 
     override def toString: String =
       imports_graph.keys_iterator.mkString("Sessions.Structure(", ", ", ")")
@@ -1099,8 +1163,7 @@
   def parse_root(path: Path): List[Entry] = Parsers.parse_root(path)
 
   def parse_root_entries(path: Path): List[Session_Entry] =
-    for (entry <- Parsers.parse_root(path) if entry.isInstanceOf[Session_Entry])
-    yield entry.asInstanceOf[Session_Entry]
+    Parsers.parse_root(path).flatMap(Library.as_subclass(classOf[Session_Entry]))
 
   def parse_roots(roots: Path): List[String] = {
     for {
@@ -1491,6 +1554,9 @@
           db.using_statement(Session_Info.augment_table)(_.execute())
         }
 
+        db.create_table(Sources.table)
+        db.using_statement(Sources.table.delete(Sources.where_equal(name)))(_.execute())
+
         db.create_table(Export.Data.table)
         db.using_statement(
           Export.Data.table.delete(Export.Data.session_name.where_equal(name)))(_.execute())
@@ -1517,14 +1583,15 @@
 
     def write_session_info(
       db: SQL.Database,
-      name: String,
+      session_name: String,
+      sources: Sources,
       build_log: Build_Log.Session_Info,
       build: Build.Session_Info
     ): Unit = {
       db.transaction {
-        val table = Session_Info.table
-        db.using_statement(table.insert()) { stmt =>
-          stmt.string(1) = name
+        write_sources(db, session_name, sources)
+        db.using_statement(Session_Info.table.insert()) { stmt =>
+          stmt.string(1) = session_name
           stmt.bytes(2) = Properties.encode(build_log.session_timing)
           stmt.bytes(3) = Properties.compress(build_log.command_timings, cache = cache.compress)
           stmt.bytes(4) = Properties.compress(build_log.theory_timings, cache = cache.compress)
@@ -1584,5 +1651,40 @@
       }
       else None
     }
+
+
+    /* session sources */
+
+    def write_sources(db: SQL.Database, session_name: String, sources: Sources): Unit = {
+      for (source_file <- sources) {
+        db.using_statement(Sources.table.insert()) { stmt =>
+          stmt.string(1) = session_name
+          stmt.string(2) = source_file.name
+          stmt.string(3) = source_file.digest.toString
+          stmt.bool(4) = source_file.compressed
+          stmt.bytes(5) = source_file.body
+          stmt.execute()
+        }
+      }
+    }
+
+    def read_sources(
+      db: SQL.Database,
+      session_name: String,
+      name: String = ""
+    ): List[Source_File] = {
+      val select =
+        Sources.table.select(Nil,
+          Sources.where_equal(session_name, name = name) + SQL.order_by(List(Sources.name)))
+      db.using_statement(select) { stmt =>
+        (stmt.execute_query().iterator { res =>
+          val res_name = res.string(Sources.name)
+          val digest = SHA1.fake_digest(res.string(Sources.digest))
+          val compressed = res.bool(Sources.compressed)
+          val body = res.bytes(Sources.body)
+          Source_File(res_name, digest, compressed, body, cache.compress)
+        }).toList
+      }
+    }
   }
 }
--- a/src/Pure/Thy/thy_header.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Thy/thy_header.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -76,32 +76,32 @@
   val bootstrap_global_theories =
     (Sessions.root_name :: (ml_roots ::: bootstrap_thys).map(_._2)).map(_ -> PURE)
 
-  private val Thy_File_Name = new Regex(""".*?([^/\\:]+)\.thy""")
-  private val Split_File_Name = new Regex("""(.*?)[/\\]*([^/\\:]+)""")
-  private val File_Name = new Regex(""".*?([^/\\:]+)""")
-
-  def is_base_name(s: String): Boolean =
-    s != "" && !s.exists("/\\:".contains(_))
-
-  def split_file_name(s: String): Option[(String, String)] =
-    s match {
-      case Split_File_Name(s1, s2) => Some((s1, s2))
-      case _ => None
-    }
-
   def import_name(s: String): String =
-    s match {
-      case File_Name(name) if !File.is_thy(name) => name
+    Url.get_base_name(s) match {
+      case Some(name) if !File.is_thy(name) => name
       case _ => error("Malformed theory import: " + quote(s))
     }
 
+  def get_thy_name(s: String): Option[String] = Url.get_base_name(s, suffix = ".thy")
+
+  def list_thy_names(dir: Path): List[String] = {
+    val files =
+      try { File.read_dir(dir) }
+      catch { case ERROR(_) => Nil }
+    files.flatMap(get_thy_name)
+  }
+
   def theory_name(s: String): String =
-    s match {
-      case Thy_File_Name(name) => bootstrap_name(name)
-      case File_Name(name) =>
-        if (name == Sessions.root_name) name
-        else ml_roots.collectFirst({ case (a, b) if a == name => b }).getOrElse("")
-      case _ => ""
+    get_thy_name(s) match {
+      case Some(name) =>
+        bootstrap_thys.collectFirst({ case (a, b) if a == name => b }).getOrElse(name)
+      case None =>
+        Url.get_base_name(s) match {
+          case Some(name) =>
+            if (name == Sessions.root_name) name
+            else ml_roots.collectFirst({ case (a, b) if a == name => b }).getOrElse("")
+          case None => ""
+        }
     }
 
   def is_ml_root(theory: String): Boolean =
@@ -110,16 +110,6 @@
   def is_bootstrap(theory: String): Boolean =
     bootstrap_thys.exists({ case (_, b) => b == theory })
 
-  def bootstrap_name(theory: String): String =
-    bootstrap_thys.collectFirst({ case (a, b) if a == theory => b }).getOrElse(theory)
-
-  def try_read_dir(dir: Path): List[String] = {
-    val files =
-      try { File.read_dir(dir) }
-      catch { case ERROR(_) => Nil }
-    for { Thy_File_Name(name) <- files } yield name
-  }
-
 
   /* parser */
 
--- a/src/Pure/Tools/build.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Tools/build.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -348,7 +348,7 @@
 
             // write database
             using(store.open_database(session_name, output = true))(db =>
-              store.write_session_info(db, session_name,
+              store.write_session_info(db, session_name, job.session_sources,
                 build_log =
                   if (process_result.timeout) build_log.error("Timeout") else build_log,
                 build =
--- a/src/Pure/Tools/build_job.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Tools/build_job.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -22,7 +22,7 @@
 
     def read_xml(name: String): XML.Body =
       YXML.parse_body(
-        Symbol.output(unicode_symbols, UTF8.decode_permissive(read(name).uncompressed)),
+        Symbol.output(unicode_symbols, UTF8.decode_permissive(read(name).bytes)),
         cache = theory_context.cache)
 
     for {
@@ -31,12 +31,9 @@
     }
     yield {
       val master_dir =
-        Thy_Header.split_file_name(thy_file) match {
-          case Some((dir, _)) => dir
-          case None => error("Cannot determine theory master directory: " + quote(thy_file))
-        }
-      val node_name =
-        Document.Node.Name(thy_file, master_dir = master_dir, theory = theory_context.theory)
+        Url.strip_base_name(thy_file).getOrElse(
+          error("Cannot determine theory master directory: " + quote(thy_file)))
+      val node_name = Document.Node.Name(thy_file, theory = theory_context.theory)
 
       val results =
         Command.Results.make(
@@ -249,6 +246,9 @@
   val info: Sessions.Info = session_background.sessions_structure(session_name)
   val options: Options = NUMA.policy_options(info.options, numa_node)
 
+  val session_sources: Sessions.Sources =
+    Sessions.Sources.load(session_background.base, cache = store.cache.compress)
+
   private val future_result: Future[Process_Result] =
     Future.thread("build", uninterruptible = true) {
       val parent = info.parent.getOrElse("")
@@ -276,8 +276,7 @@
           override val cache: Term.Cache = store.cache
 
           override def build_blobs_info(name: Document.Node.Name): Command.Blobs_Info = {
-            val name1 = name.map(s => Path.explode(s).expand.implode)
-            session_background.base.load_commands.get(name1) match {
+            session_background.base.theory_load_commands.get(name.theory) match {
               case Some(spans) =>
                 val syntax = session_background.base.theory_syntax(name)
                 Command.build_blobs_info(syntax, name, spans)
--- a/src/Pure/Tools/dump.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Tools/dump.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -57,13 +57,13 @@
           for {
             entry <- args.snapshot.exports
             if entry.name_has_prefix(Export.DOCUMENT_PREFIX)
-          } args.write(Path.explode(entry.name), entry.uncompressed)),
+          } args.write(Path.explode(entry.name), entry.bytes)),
       Aspect("theory", "foundational theory content",
         args =>
           for {
             entry <- args.snapshot.exports
             if entry.name_has_prefix(Export.THEORY_PREFIX)
-          } args.write(Path.explode(entry.name), entry.uncompressed),
+          } args.write(Path.explode(entry.name), entry.bytes),
         options = List("export_theory"))
     ).sortBy(_.name)
 
--- a/src/Pure/Tools/generated_files.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Tools/generated_files.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -285,10 +285,9 @@
 fun check_external_files ctxt (raw_files, raw_base_dir) =
   let
     val base_dir = Resources.check_dir ctxt NONE raw_base_dir;
-    fun check source =
+    val files = raw_files |> map (fn source =>
      (Resources.check_file ctxt (SOME base_dir) source;
-      Path.explode (Input.string_of source));
-    val files = map check raw_files;
+      Path.explode (Input.string_of source)));
   in (files, base_dir) end;
 
 fun get_external_files dir (files, base_dir) =
--- a/src/Pure/Tools/server_commands.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/Tools/server_commands.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -266,7 +266,7 @@
                     val matcher = Export.make_matcher(List(args.export_pattern))
                     for { entry <- snapshot.exports if matcher(entry.entry_name) }
                     yield {
-                      val (base64, body) = entry.uncompressed.maybe_encode_base64
+                      val (base64, body) = entry.bytes.maybe_encode_base64
                       JSON.Object("name" -> entry.name, "base64" -> base64, "body" -> body)
                     }
                   }))
--- a/src/Pure/context_position.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/context_position.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -17,7 +17,7 @@
   val restore_visible_generic: Context.generic -> Context.generic -> Context.generic
   val restore_visible: Proof.context -> Proof.context -> Proof.context
   val restore_visible_global: theory -> theory -> theory
-  val pide_reports: unit -> bool
+  val reports_enabled0: unit -> bool
   val reports_enabled_generic: Context.generic -> bool
   val reports_enabled: Proof.context -> bool
   val reports_enabled_global: theory -> bool
@@ -64,9 +64,8 @@
 
 (* PIDE reports *)
 
-fun pide_reports () = Options.default_bool "pide_reports";
-
-fun reports_enabled_generic context = pide_reports () andalso is_visible_generic context;
+fun reports_enabled0 () = Options.default_bool "pide_reports";
+fun reports_enabled_generic context = reports_enabled0 () andalso is_visible_generic context;
 val reports_enabled = reports_enabled_generic o Context.Proof;
 val reports_enabled_global = reports_enabled_generic o Context.Theory;
 
--- a/src/Pure/library.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/library.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -291,4 +291,7 @@
     }
     subclass(a)
   }
+
+  def as_subclass[C](c: Class[C])(x: AnyRef): Option[C] =
+    if (x == null || is_subclass(x.getClass, c)) Some(x.asInstanceOf[C]) else None
 }
--- a/src/Pure/logic.ML	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Pure/logic.ML	Tue Jan 03 18:23:52 2023 +0100
@@ -44,6 +44,7 @@
   val mk_type: typ -> term
   val dest_type: term -> typ
   val type_map: (term -> term) -> typ -> typ
+  val class_suffix: string
   val const_of_class: class -> string
   val class_of_const: string -> class
   val mk_of_class: typ * class -> term
@@ -288,11 +289,11 @@
 
 (* const names *)
 
-val classN = "_class";
+val class_suffix = "_class";
 
-val const_of_class = suffix classN;
+val const_of_class = suffix class_suffix;
 
-fun class_of_const c = unsuffix classN c
+fun class_of_const c = unsuffix class_suffix c
   handle Fail _ => raise TERM ("class_of_const: bad name " ^ quote c, []);
 
 
--- a/src/Tools/VSCode/src/language_server.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/VSCode/src/language_server.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -480,12 +480,14 @@
   /* abstract editor operations */
 
   object editor extends Language_Server.Editor {
-    /* session */
+    /* PIDE session and document model */
 
     override def session: Session = server.session
     override def flush(): Unit = resources.flush_input(session, channel)
     override def invoke(): Unit = delay_input.invoke()
 
+    override def get_models(): Iterable[Document.Model] = resources.get_models()
+
 
     /* current situation */
 
--- a/src/Tools/VSCode/src/vscode_model.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/VSCode/src/vscode_model.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -26,11 +26,7 @@
 
   /* content */
 
-  object Content {
-    val empty: Content = Content(Line.Document.empty)
-  }
-
-  sealed case class Content(doc: Line.Document) {
+  sealed case class Content(node_name: Document.Node.Name, doc: Line.Document) {
     override def toString: String = doc.toString
     def text_length: Text.Offset = doc.text_length
     def text_range: Text.Range = doc.text_range
@@ -38,9 +34,7 @@
 
     lazy val bytes: Bytes = Bytes(Symbol.encode(text))
     lazy val chunk: Symbol.Text_Chunk = Symbol.Text_Chunk(text)
-    lazy val bibtex_entries: List[Text.Info[String]] =
-      try { Bibtex.entries(text) }
-      catch { case ERROR(_) => Nil }
+    lazy val data: AnyRef = File_Format.registry.parse_data(node_name, text)
 
     def recode_symbols: List[LSP.TextEdit] =
       (for {
@@ -58,15 +52,15 @@
     editor: Language_Server.Editor,
     node_name: Document.Node.Name
   ): VSCode_Model = {
-    VSCode_Model(session, editor, node_name, Content.empty,
-      node_required = File_Format.registry.is_theory(node_name))
+    val content = Content(node_name, Line.Document.empty)
+    val is_theory = File_Format.registry.is_theory(node_name)
+    VSCode_Model(session, editor, content, node_required = is_theory)
   }
 }
 
 sealed case class VSCode_Model(
   session: Session,
   editor: Language_Server.Editor,
-  node_name: Document.Node.Name,
   content: VSCode_Model.Content,
   version: Option[Long] = None,
   external_file: Boolean = false,
@@ -81,6 +75,8 @@
 
   /* content */
 
+  def node_name: Document.Node.Name = content.node_name
+
   def get_text(range: Text.Range): Option[String] = content.doc.get_text(range)
 
   def set_version(new_version: Long): VSCode_Model = copy(version = Some(new_version))
@@ -151,10 +147,9 @@
     else Some(Document.Blob(content.bytes, content.text, content.chunk, pending_edits.nonEmpty))
 
 
-  /* bibtex entries */
+  /* data */
 
-  def bibtex_entries: List[Text.Info[String]] =
-    model.content.bibtex_entries
+  def untyped_data: AnyRef = model.content.data
 
 
   /* edits */
@@ -166,7 +161,7 @@
         Text.Edit.replace(0, content.text, insert) match {
           case Nil => None
           case edits =>
-            val content1 = VSCode_Model.Content(Line.Document(insert))
+            val content1 = VSCode_Model.Content(node_name, Line.Document(insert))
             Some(copy(content = content1, pending_edits = pending_edits ::: edits))
         }
       case Some(remove) =>
@@ -174,7 +169,7 @@
           case None => error("Failed to apply document change: " + remove)
           case Some((Nil, _)) => None
           case Some((edits, doc1)) =>
-            val content1 = VSCode_Model.Content(doc1)
+            val content1 = VSCode_Model.Content(node_name, doc1)
             Some(copy(content = content1, pending_edits = pending_edits ::: edits))
         }
     }
--- a/src/Tools/VSCode/src/vscode_rendering.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/VSCode/src/vscode_rendering.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -75,15 +75,6 @@
   override def get_text(range: Text.Range): Option[String] = model.get_text(range)
 
 
-  /* bibtex */
-
-  def bibtex_entries_iterator(): Iterator[Text.Info[(String, VSCode_Model)]] =
-    Bibtex.entries_iterator(resources.get_models())
-
-  def bibtex_completion(history: Completion.History, caret: Text.Offset): Option[Completion.Result] =
-    Bibtex.completion(history, rendering, caret, resources.get_models())
-
-
   /* completion */
 
   def completion(node_pos: Line.Node_Position, caret: Text.Offset): List[LSP.CompletionItem] = {
@@ -114,7 +105,7 @@
               syntax_completion,
               VSCode_Spell_Checker.completion(rendering, caret),
               path_completion(caret),
-              bibtex_completion(history, caret))
+              model.editor.bibtex_completion(history, rendering, caret))
           val items =
             results match {
               case None => Nil
@@ -323,7 +314,8 @@
           case (links, Text.Info(info_range, XML.Elem(Markup.Citation(name), _))) =>
             val iterator =
               for {
-                Text.Info(entry_range, (entry, model)) <- bibtex_entries_iterator()
+                Text.Info(entry_range, (entry, model: VSCode_Model)) <-
+                  model.editor.bibtex_entries_iterator()
                 if entry == name
               } yield Line.Node_Range(model.node_name.node, model.content.doc.range(entry_range))
             if (iterator.isEmpty) None else Some(iterator.foldLeft(links)(_ :+ _))
--- a/src/Tools/VSCode/src/vscode_resources.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/VSCode/src/vscode_resources.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -93,28 +93,25 @@
     find_theory(file) getOrElse {
       val node = file.getPath
       val theory = theory_name(Sessions.DRAFT, Thy_Header.theory_name(node))
-      if (session_base.loaded_theory(theory)) loaded_theory_node(theory)
-      else {
-        val master_dir = file.getParent
-        Document.Node.Name(node, master_dir, theory)
-      }
+      if (session_base.loaded_theory(theory)) Document.Node.Name.loaded_theory(theory)
+      else Document.Node.Name(node, theory = theory)
     }
 
   override def migrate_name(standard_name: Document.Node.Name): Document.Node.Name =
     node_name(Path.explode(standard_name.node).canonical_file)
 
-  override def append(dir: String, source_path: Path): String = {
+  override def append_path(prefix: String, source_path: Path): String = {
     val path = source_path.expand
-    if (dir == "" || path.is_absolute) File.platform_path(path)
-    else if (path.is_current) dir
-    else if (path.is_basic && !dir.endsWith("/") && !dir.endsWith(JFile.separator))
-      dir + JFile.separator + File.platform_path(path)
-    else if (path.is_basic) dir + File.platform_path(path)
-    else File.absolute_name(new JFile(dir + JFile.separator + File.platform_path(path)))
+    if (prefix == "" || path.is_absolute) File.platform_path(path)
+    else if (path.is_current) prefix
+    else if (path.is_basic && !prefix.endsWith("/") && !prefix.endsWith(JFile.separator))
+      prefix + JFile.separator + File.platform_path(path)
+    else if (path.is_basic) prefix + File.platform_path(path)
+    else File.absolute_name(new JFile(prefix + JFile.separator + File.platform_path(path)))
   }
 
-  def get_models(): Map[JFile, VSCode_Model] = state.value.models
-  def get_model(file: JFile): Option[VSCode_Model] = get_models().get(file)
+  def get_models(): Iterable[VSCode_Model] = state.value.models.values
+  def get_model(file: JFile): Option[VSCode_Model] = state.value.models.get(file)
   def get_model(name: Document.Node.Name): Option[VSCode_Model] = get_model(node_file(name))
 
 
@@ -123,7 +120,7 @@
   def snapshot(model: VSCode_Model): Document.Snapshot =
     model.session.snapshot(
       node_name = model.node_name,
-      pending_edits = Document.Pending_Edits.make(get_models().values))
+      pending_edits = Document.Pending_Edits.make(get_models()))
 
   def get_snapshot(file: JFile): Option[Document.Snapshot] =
     get_model(file).map(snapshot)
--- a/src/Tools/jEdit/src/completion_popup.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/completion_popup.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -122,7 +122,7 @@
                   Completion.Result.merge(Completion.History.empty,
                     syntax_completion(Completion.History.empty, true, Some(rendering)),
                     rendering.path_completion(caret),
-                    Document_Model.bibtex_completion(Completion.History.empty, rendering, caret))
+                    PIDE.editor.bibtex_completion(Completion.History.empty, rendering, caret))
                   .map(_.range)
                 rendering.semantic_completion(range0, range) match {
                   case None => range0
@@ -304,7 +304,7 @@
                   result1,
                   JEdit_Spell_Checker.completion(rendering, explicit, caret),
                   rendering.path_completion(caret),
-                  Document_Model.bibtex_completion(history, rendering, caret))
+                  PIDE.editor.bibtex_completion(history, rendering, caret))
             }
           }
           result match {
--- a/src/Tools/jEdit/src/document_model.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/document_model.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -32,8 +32,8 @@
     def file_models_iterator: Iterator[(Document.Node.Name, File_Model)] =
       for {
         (node_name, model) <- models.iterator
-        if model.isInstanceOf[File_Model]
-      } yield (node_name, model.asInstanceOf[File_Model])
+        file_model <- Library.as_subclass(classOf[File_Model])(model)
+      } yield (node_name, file_model)
 
     def document_blobs: Document.Blobs =
       Document.Blobs(
@@ -53,7 +53,7 @@
           case Some(buffer_model: Buffer_Model) => Some(buffer_model.exit())
           case _ => None
         }
-      val buffer_model = Buffer_Model(session, node_name, buffer).init(old_model)
+      val buffer_model = Buffer_Model.init(old_model, session, node_name, buffer)
       (buffer_model,
         copy(models = models + (node_name -> buffer_model),
           buffer_models = buffer_models + (buffer -> buffer_model)))
@@ -73,7 +73,8 @@
       if (models.isDefinedAt(node_name)) this
       else {
         val edit = Text.Edit.insert(0, text)
-        val model = File_Model.init(session, node_name, text, pending_edits = List(edit))
+        val content = File_Content(node_name, text)
+        val model = File_Model.init(session, content = content, pending_edits = List(edit))
         copy(models = models + (node_name -> model))
       }
   }
@@ -84,30 +85,21 @@
 
   def document_blobs(): Document.Blobs = state.value.document_blobs
 
-  def get_models(): Map[Document.Node.Name, Document_Model] = state.value.models
-  def get_model(name: Document.Node.Name): Option[Document_Model] = get_models().get(name)
+  def get_models_map(): Map[Document.Node.Name, Document_Model] = state.value.models
+  def get_models(): Iterable[Document_Model] = get_models_map().values
+  def get_model(name: Document.Node.Name): Option[Document_Model] = get_models_map().get(name)
   def get_model(buffer: JEditBuffer): Option[Buffer_Model] =
     state.value.buffer_models.get(buffer)
 
   def snapshot(model: Document_Model): Document.Snapshot =
     PIDE.session.snapshot(
       node_name = model.node_name,
-      pending_edits = Document.Pending_Edits.make(get_models().values))
+      pending_edits = Document.Pending_Edits.make(get_models()))
 
   def get_snapshot(name: Document.Node.Name): Option[Document.Snapshot] = get_model(name).map(snapshot)
   def get_snapshot(buffer: JEditBuffer): Option[Document.Snapshot] = get_model(buffer).map(snapshot)
 
 
-  /* bibtex */
-
-  def bibtex_entries_iterator(): Iterator[Text.Info[(String, Document_Model)]] =
-    Bibtex.entries_iterator(state.value.models)
-
-  def bibtex_completion(history: Completion.History, rendering: Rendering, caret: Text.Offset)
-      : Option[Completion.Result] =
-    Bibtex.completion(history, rendering, caret, state.value.models)
-
-
   /* overlays */
 
   def node_overlays(name: Document.Node.Name): Document.Node.Overlays =
@@ -131,7 +123,7 @@
           text <- PIDE.resources.read_file_content(node_name)
           if model.content.text != text
         } yield {
-          val content = Document_Model.File_Content(text)
+          val content = Document_Model.File_Content(node_name, text)
           val edits = Text.Edit.replace(0, model.content.text, text)
           (node_name, model.copy(content = content, pending_edits = model.pending_edits ::: edits))
         }).toList
@@ -287,12 +279,18 @@
 
   /* file content */
 
-  sealed case class File_Content(text: String) {
+  object File_Content {
+    val empty: File_Content = apply(Document.Node.Name.empty, "")
+
+    def apply(node_name: Document.Node.Name, text: String): File_Content =
+      new File_Content(node_name, text)
+  }
+
+  final class File_Content private(val node_name: Document.Node.Name, val text: String) {
+    override def toString: String = "Document_Model.File_Content(" + node_name.node + ")"
     lazy val bytes: Bytes = Bytes(Symbol.encode(text))
     lazy val chunk: Symbol.Text_Chunk = Symbol.Text_Chunk(text)
-    lazy val bibtex_entries: List[Text.Info[String]] =
-      try { Bibtex.entries(text) }
-      catch { case ERROR(_) => Nil }
+    lazy val data: AnyRef = File_Format.registry.parse_data(node_name, text)
   }
 
 
@@ -316,7 +314,7 @@
       PIDE.plugin.http_server.url + "/" + service.name + "?" +
         (if (plain_text) plain_text_prefix else "") + Url.encode(node_name.node)
 
-    def apply(request: HTTP.Request): Option[HTTP.Response] =
+    def apply(request: HTTP.Request): Option[HTTP.Response] = GUI_Thread.now {
       for {
         query <- request.decode_query
         name = Library.perhaps_unprefix(plain_text_prefix, query)
@@ -333,6 +331,7 @@
             fonts_css = HTML.fonts_css_dir(HTTP.url_path(request.server_name)))
         HTTP.Response.html(document.content)
       }
+    }
   }
 }
 
@@ -384,46 +383,43 @@
 }
 
 object File_Model {
-  def empty(session: Session): File_Model =
-    File_Model(session, Document.Node.Name.empty, None, Document_Model.File_Content(""),
-      false, Document.Node.Perspective_Text.empty, Nil)
-
   def init(session: Session,
-    node_name: Document.Node.Name,
-    text: String,
+    content: Document_Model.File_Content = Document_Model.File_Content.empty,
     node_required: Boolean = false,
     last_perspective: Document.Node.Perspective_Text.T = Document.Node.Perspective_Text.empty,
     pending_edits: List[Text.Edit] = Nil
   ): File_Model = {
+    val node_name = content.node_name
+
     val file = JEdit_Lib.check_file(node_name.node)
     file.foreach(PIDE.plugin.file_watcher.register_parent(_))
 
-    val content = Document_Model.File_Content(text)
     val node_required1 = node_required || File_Format.registry.is_theory(node_name)
-    File_Model(session, node_name, file, content, node_required1, last_perspective, pending_edits)
+    File_Model(session, file, content, node_required1, last_perspective, pending_edits)
   }
 }
 
 case class File_Model(
   session: Session,
-  node_name: Document.Node.Name,
   file: Option[JFile],
   content: Document_Model.File_Content,
   node_required: Boolean,
   last_perspective: Document.Node.Perspective_Text.T,
   pending_edits: List[Text.Edit]
 ) extends Document_Model {
+  /* content */
+
+  def node_name: Document.Node.Name = content.node_name
+
+  def get_text(range: Text.Range): Option[String] =
+    range.try_substring(content.text)
+
+
   /* required */
 
   def set_node_required(b: Boolean): File_Model = copy(node_required = b)
 
 
-  /* text */
-
-  def get_text(range: Text.Range): Option[String] =
-    range.try_substring(content.text)
-
-
   /* header */
 
   def node_header: Document.Node.Header =
@@ -441,8 +437,7 @@
     if (is_theory) None
     else Some(Document.Blob(content.bytes, content.text, content.chunk, pending_edits.nonEmpty))
 
-  def bibtex_entries: List[Text.Info[String]] =
-    if (File.is_bib(node_name.node)) content.bibtex_entries else Nil
+  def untyped_data: AnyRef = content.data
 
 
   /* edits */
@@ -451,7 +446,7 @@
     Text.Edit.replace(0, content.text, text) match {
       case Nil => None
       case edits =>
-        val content1 = Document_Model.File_Content(text)
+        val content1 = Document_Model.File_Content(node_name, text)
         val pending_edits1 = pending_edits ::: edits
         Some(copy(content = content1, pending_edits = pending_edits1))
     }
@@ -480,8 +475,20 @@
   def is_stable: Boolean = pending_edits.isEmpty
 }
 
-case class Buffer_Model(session: Session, node_name: Document.Node.Name, buffer: Buffer)
-extends Document_Model {
+object Buffer_Model {
+  def init(
+    old_model: Option[File_Model],
+    session: Session,
+    node_name: Document.Node.Name,
+    buffer: Buffer
+  ): Buffer_Model = (new Buffer_Model(session, node_name, buffer)).init(old_model)
+}
+
+class Buffer_Model private(
+  val session: Session,
+  val node_name: Document.Node.Name,
+  val buffer: Buffer
+) extends Document_Model {
   /* text */
 
   def get_text(range: Text.Range): Option[String] =
@@ -502,11 +509,6 @@
 
   /* perspective */
 
-  // owned by GUI thread
-  private var _node_required = false
-  def node_required: Boolean = _node_required
-  def set_node_required(b: Boolean): Unit = GUI_Thread.require { _node_required = b }
-
   def document_view_iterator: Iterator[Document_View] =
     for {
       text_area <- JEdit_Lib.jedit_text_areas(buffer)
@@ -523,105 +525,100 @@
   }
 
 
-  /* blob */
-
-  // owned by GUI thread
-  private var _blob: Option[(Bytes, String, Symbol.Text_Chunk)] = None
-
-  private def reset_blob(): Unit = GUI_Thread.require { _blob = None }
+  /* mutable buffer state: owned by GUI thread */
 
-  def get_blob: Option[Document.Blob] =
-    GUI_Thread.require {
-      if (is_theory) None
-      else {
-        val (bytes, text, chunk) =
-          _blob match {
-            case Some(x) => x
-            case None =>
-              val bytes = PIDE.resources.make_file_content(buffer)
-              val text = buffer.getText(0, buffer.getLength)
-              val chunk = Symbol.Text_Chunk(text)
-              val x = (bytes, text, chunk)
-              _blob = Some(x)
-              x
-          }
-        val changed = !buffer_edits.is_empty
-        Some(Document.Blob(bytes, text, chunk, changed))
-      }
-    }
-
-
-  /* bibtex entries */
+  private object buffer_state {
+    // perspective and edits
 
-  // owned by GUI thread
-  private var _bibtex_entries: Option[List[Text.Info[String]]] = None
-
-  private def reset_bibtex_entries(): Unit = GUI_Thread.require { _bibtex_entries = None }
+    private var last_perspective = Document.Node.Perspective_Text.empty
+    def get_last_perspective: Document.Node.Perspective_Text.T =
+      GUI_Thread.require { last_perspective }
+    def set_last_perspective(perspective: Document.Node.Perspective_Text.T): Unit =
+      GUI_Thread.require { last_perspective = perspective }
 
-  def bibtex_entries: List[Text.Info[String]] =
-    GUI_Thread.require {
-      if (File.is_bib(node_name.node)) {
-        _bibtex_entries match {
-          case Some(entries) => entries
-          case None =>
-            val text = JEdit_Lib.buffer_text(buffer)
-            val entries =
-              try { Bibtex.entries(text) }
-              catch { case ERROR(msg) => Output.warning(msg); Nil }
-            _bibtex_entries = Some(entries)
-            entries
-        }
-      }
-      else Nil
-    }
+    private var node_required = false
+    def get_node_required: Boolean = GUI_Thread.require { node_required }
+    def set_node_required(b: Boolean): Unit = GUI_Thread.require { node_required = b }
 
-
-  /* pending buffer edits */
-
-  private object buffer_edits {
-    private val pending = new mutable.ListBuffer[Text.Edit]
-    private var last_perspective = Document.Node.Perspective_Text.empty
-
-    def is_empty: Boolean = synchronized { pending.isEmpty }
-    def get_edits: List[Text.Edit] = synchronized { pending.toList }
-    def get_last_perspective: Document.Node.Perspective_Text.T = synchronized { last_perspective }
-    def set_last_perspective(perspective: Document.Node.Perspective_Text.T): Unit =
-      synchronized { last_perspective = perspective }
+    private val pending_edits = new mutable.ListBuffer[Text.Edit]
+    def is_stable: Boolean = GUI_Thread.require { pending_edits.isEmpty }
+    def get_pending_edits: List[Text.Edit] = GUI_Thread.require { pending_edits.toList }
 
     def flush_edits(doc_blobs: Document.Blobs, hidden: Boolean): List[Document.Edit_Text] =
-      synchronized {
-        GUI_Thread.require {}
-
-        val edits = get_edits
+      GUI_Thread.require {
+        val edits = get_pending_edits
         val (reparse, perspective) = node_perspective(doc_blobs, hidden)
         if (reparse || edits.nonEmpty || last_perspective != perspective) {
-          pending.clear()
+          pending_edits.clear()
           last_perspective = perspective
           node_edits(node_header(), edits, perspective)
         }
         else Nil
       }
 
-    def edit(edits: List[Text.Edit]): Unit = synchronized {
-      GUI_Thread.require {}
-
+    def edit(edits: List[Text.Edit]): Unit = GUI_Thread.require {
       reset_blob()
-      reset_bibtex_entries()
+      reset_data()
 
       for (doc_view <- document_view_iterator) {
         doc_view.rich_text_area.active_reset()
       }
 
-      pending ++= edits
+      pending_edits ++= edits
       PIDE.editor.invoke()
     }
+
+
+    // blob
+
+    private var blob: Option[(Bytes, String, Symbol.Text_Chunk)] = None
+
+    def reset_blob(): Unit = GUI_Thread.require { blob = None }
+
+    def get_blob: Option[Document.Blob] = GUI_Thread.require {
+      if (is_theory) None
+      else {
+        val (bytes, text, chunk) =
+          blob getOrElse {
+            val bytes = PIDE.resources.make_file_content(buffer)
+            val text = buffer.getText(0, buffer.getLength)
+            val chunk = Symbol.Text_Chunk(text)
+            val x = (bytes, text, chunk)
+            blob = Some(x)
+            x
+          }
+        val changed = !is_stable
+        Some(Document.Blob(bytes, text, chunk, changed))
+      }
+    }
+
+
+    // parsed data
+
+    private var data: Option[AnyRef] = None
+
+    def reset_data(): Unit = GUI_Thread.require { data = None }
+
+    def untyped_data: AnyRef = GUI_Thread.require {
+      data getOrElse {
+        val text = JEdit_Lib.buffer_text(buffer)
+        val res = File_Format.registry.parse_data(node_name, text)
+        data = Some(res)
+        res
+      }
+    }
   }
 
-  def is_stable: Boolean = buffer_edits.is_empty
-  def pending_edits: List[Text.Edit] = buffer_edits.get_edits
+  def is_stable: Boolean = buffer_state.is_stable
+  def pending_edits: List[Text.Edit] = buffer_state.get_pending_edits
+  def flush_edits(doc_blobs: Document.Blobs, hidden: Boolean): List[Document.Edit_Text] =
+    buffer_state.flush_edits(doc_blobs, hidden)
 
-  def flush_edits(doc_blobs: Document.Blobs, hidden: Boolean): List[Document.Edit_Text] =
-    buffer_edits.flush_edits(doc_blobs, hidden)
+  def node_required: Boolean = buffer_state.get_node_required
+  def set_node_required(b: Boolean): Unit = buffer_state.set_node_required(b)
+
+  def get_blob: Option[Document.Blob] = buffer_state.get_blob
+  def untyped_data: AnyRef = buffer_state.untyped_data
 
 
   /* buffer listener */
@@ -634,7 +631,7 @@
       num_lines: Int,
       length: Int
     ): Unit = {
-      buffer_edits.edit(List(Text.Edit.insert(offset, buffer.getText(offset, length))))
+      buffer_state.edit(List(Text.Edit.insert(offset, buffer.getText(offset, length))))
     }
 
     override def preContentRemoved(
@@ -644,7 +641,7 @@
       num_lines: Int,
       removed_length: Int
     ): Unit = {
-      buffer_edits.edit(List(Text.Edit.remove(offset, buffer.getText(offset, removed_length))))
+      buffer_state.edit(List(Text.Edit.remove(offset, buffer.getText(offset, removed_length))))
     }
   }
 
@@ -672,16 +669,14 @@
 
   /* init */
 
-  def init(old_model: Option[File_Model]): Buffer_Model = {
-    GUI_Thread.require {}
-
+  def init(old_model: Option[File_Model]): Buffer_Model = GUI_Thread.require {
     old_model match {
       case None =>
-        buffer_edits.edit(List(Text.Edit.insert(0, JEdit_Lib.buffer_text(buffer))))
+        buffer_state.edit(List(Text.Edit.insert(0, JEdit_Lib.buffer_text(buffer))))
       case Some(file_model) =>
         set_node_required(file_model.node_required)
-        buffer_edits.set_last_perspective(file_model.last_perspective)
-        buffer_edits.edit(
+        buffer_state.set_last_perspective(file_model.last_perspective)
+        buffer_state.edit(
           file_model.pending_edits :::
             Text.Edit.replace(0, file_model.content.text, JEdit_Lib.buffer_text(buffer)))
     }
@@ -695,15 +690,14 @@
 
   /* exit */
 
-  def exit(): File_Model = {
-    GUI_Thread.require {}
-
+  def exit(): File_Model = GUI_Thread.require {
     buffer.removeBufferListener(buffer_listener)
     init_token_marker()
 
-    File_Model.init(session, node_name, JEdit_Lib.buffer_text(buffer),
-      node_required = _node_required,
-      last_perspective = buffer_edits.get_last_perspective,
+    File_Model.init(session,
+      content = Document_Model.File_Content(node_name, JEdit_Lib.buffer_text(buffer)),
+      node_required = node_required,
+      last_perspective = buffer_state.get_last_perspective,
       pending_edits = pending_edits)
   }
 }
--- a/src/Tools/jEdit/src/graphview_dockable.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/graphview_dockable.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -87,7 +87,7 @@
           override def make_tooltip(parent: JComponent, x: Int, y: Int, body: XML.Body): String = {
             Pretty_Tooltip.invoke(() =>
               {
-                val model = File_Model.empty(PIDE.session)
+                val model = File_Model.init(PIDE.session)
                 val rendering = JEdit_Rendering(snapshot, model, options)
                 val info = Text.Info(Text.Range.offside, body)
                 Pretty_Tooltip(view, parent, new Point(x, y), rendering, Command.Results.empty, info)
--- a/src/Tools/jEdit/src/isabelle_export.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/isabelle_export.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -56,7 +56,7 @@
                 } yield {
                   val is_dir = entry.name_elems.length > depth
                   val path = Export.implode_name(theory :: entry.name_elems.take(depth))
-                  val file_size = if (is_dir) None else Some(entry.uncompressed.length.toLong)
+                  val file_size = if (is_dir) None else Some(entry.bytes.length.toLong)
                   (path, file_size)
                 }).toSet.toList
               (for ((path, file_size) <- entries.iterator) yield {
@@ -86,7 +86,7 @@
               if node_name.theory == theory
               (_, entry) <- snapshot.state.node_exports(version, node_name).iterator
               if entry.name_elems == name_elems
-            } yield entry.uncompressed).find(_ => true).getOrElse(Bytes.empty)
+            } yield entry.bytes).find(_ => true).getOrElse(Bytes.empty)
 
           bytes.stream()
       }
--- a/src/Tools/jEdit/src/jedit_editor.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/jedit_editor.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -18,7 +18,7 @@
 
 
 class JEdit_Editor extends Editor[View] {
-  /* session */
+  /* PIDE session and document model */
 
   override def session: Session = PIDE.session
 
@@ -53,6 +53,9 @@
     } yield doc_view.model.node_name).contains(name)
 
 
+  override def get_models(): Iterable[Document.Model] = Document_Model.get_models()
+
+
   /* global changes */
 
   def state_changed(): Unit = {
--- a/src/Tools/jEdit/src/jedit_lib.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/jedit_lib.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -59,10 +59,10 @@
   }
 
 
-  /* files */
+  /* plain files */
 
   def is_file(name: String): Boolean =
-    VFSManager.getVFSForPath(name).isInstanceOf[FileVFS]
+    name != null && name.nonEmpty && VFSManager.getVFSForPath(name).isInstanceOf[FileVFS]
 
   def check_file(name: String): Option[JFile] =
     if (is_file(name)) Some(new JFile(name)) else None
--- a/src/Tools/jEdit/src/jedit_rendering.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/jedit_rendering.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -32,7 +32,7 @@
   ): (String, JEdit_Rendering) = {
     val command = Command.rich_text(Document_ID.make(), results, formatted_body)
     val snippet = snapshot.snippet(command)
-    val model = File_Model.empty(PIDE.session)
+    val model = File_Model.init(PIDE.session)
     val rendering = apply(snippet, model, PIDE.options.value)
     (command.source, rendering)
   }
@@ -272,8 +272,8 @@
 
           case (links, Text.Info(info_range, XML.Elem(Markup.Citation(name), _))) =>
             val opt_link =
-              Document_Model.bibtex_entries_iterator().collectFirst(
-                { case Text.Info(entry_range, (entry, model)) if entry == name =>
+              PIDE.editor.bibtex_entries_iterator().collectFirst(
+                { case Text.Info(entry_range, (entry, model: Document_Model)) if entry == name =>
                     PIDE.editor.hyperlink_model(true, model, entry_range.start) })
             opt_link.map(link => links :+ Text.Info(snapshot.convert(info_range), link))
 
--- a/src/Tools/jEdit/src/jedit_resources.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/jedit_resources.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -35,11 +35,8 @@
       val vfs = VFSManager.getVFSForPath(path)
       val node = if (vfs.isInstanceOf[FileVFS]) MiscUtilities.resolveSymlinks(path) else path
       val theory = theory_name(Sessions.DRAFT, Thy_Header.theory_name(node))
-      if (session_base.loaded_theory(theory)) loaded_theory_node(theory)
-      else {
-        val master_dir = vfs.getParentOfPath(path)
-        Document.Node.Name(node, master_dir, theory)
-      }
+      if (session_base.loaded_theory(theory)) Document.Node.Name.loaded_theory(theory)
+      else Document.Node.Name(node, theory = theory)
     }
 
   def node_name(buffer: Buffer): Document.Node.Name =
@@ -53,19 +50,19 @@
   override def migrate_name(standard_name: Document.Node.Name): Document.Node.Name =
     node_name(File.platform_path(Path.explode(standard_name.node).canonical))
 
-  override def append(dir: String, source_path: Path): String = {
+  override def append_path(prefix: String, source_path: Path): String = {
     val path = source_path.expand
-    if (dir == "" || path.is_absolute) {
+    if (prefix == "" || path.is_absolute) {
       MiscUtilities.resolveSymlinks(File.platform_path(path))
     }
-    else if (path.is_current) dir
+    else if (path.is_current) prefix
     else {
-      val vfs = VFSManager.getVFSForPath(dir)
+      val vfs = VFSManager.getVFSForPath(prefix)
       if (vfs.isInstanceOf[FileVFS]) {
         MiscUtilities.resolveSymlinks(
-          vfs.constructPath(dir, File.platform_path(path)))
+          vfs.constructPath(prefix, File.platform_path(path)))
       }
-      else vfs.constructPath(dir, File.standard_path(path))
+      else vfs.constructPath(prefix, File.standard_path(path))
     }
   }
 
@@ -116,7 +113,7 @@
     def content(): Bytes = Bytes(this.buf, 0, this.count)
   }
 
-  private class File_Content(buffer: Buffer)
+  private class File_Content_Request(buffer: Buffer)
   extends BufferIORequest(null, buffer, null, VFSManager.getVFSForPath(buffer.getPath), buffer.getPath) {
     def _run(): Unit = {}
     def content(): Bytes = {
@@ -126,7 +123,7 @@
     }
   }
 
-  def make_file_content(buffer: Buffer): Bytes = (new File_Content(buffer)).content()
+  def make_file_content(buffer: Buffer): Bytes = (new File_Content_Request(buffer)).content()
 
 
   /* theory text edits */
--- a/src/Tools/jEdit/src/main_plugin.scala	Mon Dec 26 14:34:32 2022 +0100
+++ b/src/Tools/jEdit/src/main_plugin.scala	Tue Jan 03 18:23:52 2023 +0100
@@ -110,7 +110,7 @@
 
   private def delay_load_body(): Unit = {
     val required_files = {
-      val models = Document_Model.get_models()
+      val models = Document_Model.get_models_map()
 
       val thy_files =
         resources.resolve_dependencies(models.values, PIDE.editor.document_required())
@@ -124,7 +124,7 @@
         }
         else Nil
 
-      (thy_files ::: aux_files).filterNot(models.isDefinedAt)
+      (thy_files ::: aux_files).filterNot(models.keySet)
     }
     if (required_files.nonEmpty) {
       try {