merged
authorwenzelm
Tue, 16 Apr 2024 17:28:58 +0200
changeset 80126 b73df63e0f52
parent 80107 247751d25102 (diff)
parent 80125 761bd2b35217 (current diff)
child 80127 39f9084a9668
merged
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -36,11 +36,19 @@
   by (metis path_connected_continuous_image retract_of_def retraction)
 
 lemma retract_of_simply_connected:
-  "\<lbrakk>simply_connected T; S retract_of T\<rbrakk> \<Longrightarrow> simply_connected S"
-  apply (simp add: retract_of_def retraction_def Pi_iff, clarify)
-  apply (rule simply_connected_retraction_gen)
-       apply (force elim!: continuous_on_subset)+
-  done
+  assumes T: "simply_connected T" and "S retract_of T"
+  shows "simply_connected S"
+proof -
+  obtain r where r: "retraction T S r"
+    using assms by (metis retract_of_def) 
+  have "S \<subseteq> T"
+    by (meson \<open>retraction T S r\<close> retraction)
+  then have "(\<lambda>a. a) \<in> S \<rightarrow> T"
+    by blast
+  then show ?thesis
+    using simply_connected_retraction_gen [OF T]
+    by (metis (no_types) r retraction retraction_refl)
+qed
 
 lemma retract_of_homotopically_trivial:
   assumes ts: "T retract_of S"
@@ -56,10 +64,7 @@
   then obtain k where "Retracts S r T k"
     unfolding Retracts_def using continuous_on_id by blast
   then show ?thesis
-    apply (rule Retracts.homotopically_trivial_retraction_gen)
-    using assms
-    apply (force simp: hom)+
-    done
+    by (rule Retracts.homotopically_trivial_retraction_gen) (use assms hom in force)+
 qed
 
 lemma retract_of_homotopically_trivial_null:
@@ -74,9 +79,11 @@
   then obtain k where "Retracts S r T k"
     unfolding Retracts_def by fastforce
   then show ?thesis
-    apply (rule Retracts.homotopically_trivial_retraction_null_gen)
-    apply (rule TrueI refl assms that | assumption)+
-    done
+  proof (rule Retracts.homotopically_trivial_retraction_null_gen)
+    show "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S\<rbrakk>
+         \<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>a. True) U S f (\<lambda>x. c)"
+      using hom by blast
+  qed (use assms that in auto)
 qed
 
 lemma retraction_openin_vimage_iff:
@@ -101,13 +108,13 @@
   assumes "locally connected T" "S retract_of T"
   shows "locally connected S"
   using assms
-  by (metis Abstract_Topology_2.retraction_openin_vimage_iff idempotent_imp_retraction locally_connected_quotient_image retract_ofE)
+  by (metis retraction_openin_vimage_iff idempotent_imp_retraction locally_connected_quotient_image retract_ofE)
 
 lemma retract_of_locally_path_connected:
   assumes "locally path_connected T" "S retract_of T"
   shows "locally path_connected S"
   using assms
-  by (metis Abstract_Topology_2.retraction_openin_vimage_iff idempotent_imp_retraction locally_path_connected_quotient_image retract_ofE)
+  by (metis retraction_openin_vimage_iff idempotent_imp_retraction locally_path_connected_quotient_image retract_ofE)
 
 text \<open>A few simple lemmas about deformation retracts\<close>
 
@@ -137,15 +144,16 @@
   then show ?rhs
     by (auto simp: retract_of_def retraction_def)
 next
-  assume ?rhs
-  then show ?lhs
-    apply (clarsimp simp add: retract_of_def retraction_def)
-    apply (rule_tac x=r in exI, simp)
-     apply (rule homotopic_with_trans, assumption)
-     apply (rule_tac f = "r \<circ> f" and g="r \<circ> id" in homotopic_with_eq)
-        apply (rule_tac Y=S in homotopic_with_compose_continuous_left)
+  assume R: ?rhs
+  have "\<And>r f. \<lbrakk>T \<subseteq> S; continuous_on S r; homotopic_with_canon (\<lambda>x. True) S S id f;
+            f \<in> S \<rightarrow> T; r \<in> S \<rightarrow> T; \<forall>x\<in>T. r x = x\<rbrakk>
+           \<Longrightarrow> homotopic_with_canon (\<lambda>x. True) S S f r"
+    apply (rule_tac f = "r \<circ> f" and g="r \<circ> id" in homotopic_with_eq)
+       apply (rule_tac Y=S in homotopic_with_compose_continuous_left)
          apply (auto simp: homotopic_with_sym Pi_iff)
     done
+  with R homotopic_with_trans show ?lhs
+    unfolding retract_of_def retraction_def by blast
 qed
 
 lemma deformation_retract_of_contractible_sing:
@@ -959,7 +967,7 @@
       qed }
     with ks_f' eq \<open>a \<noteq> f' n\<close> \<open>n \<noteq> 0\<close> show ?thesis
       apply (intro ex1I[of _ "f' ` {.. n}"])
-      apply auto []
+       apply auto []
       apply metis
       done
   next
@@ -1523,7 +1531,7 @@
              (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)"
   unfolding all_conj_distrib [symmetric] 
   apply (subst choice_iff[symmetric])+
-  by (metis assms bot_nat_0.extremum nle_le zero_neq_one)
+  by (metis assms choice_iff bot_nat_0.extremum nle_le zero_neq_one)
 
 subsection \<open>Brouwer's fixed point theorem\<close>
 
@@ -1543,10 +1551,8 @@
     by auto
   obtain d where
       d: "d > 0" "\<And>x. x \<in> cbox 0 One \<Longrightarrow> d \<le> norm (f x - x)"
-    apply (rule brouwer_compactness_lemma[OF compact_cbox _ *])
-    apply (rule continuous_intros assms)+
-    apply blast
-    done
+    using brouwer_compactness_lemma[OF compact_cbox _ *] assms
+    by (metis (no_types, lifting) continuous_on_cong continuous_on_diff continuous_on_id)
   have *: "\<forall>x. x \<in> cbox 0 One \<longrightarrow> f x \<in> cbox 0 One"
     "\<forall>x. x \<in> (cbox 0 One::'a set) \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
     using assms(2)[unfolded image_subset_iff Ball_def]
@@ -1581,8 +1587,7 @@
       case False
       then show ?thesis
         using label [OF \<open>i \<in> Basis\<close>] i(1) x y
-        apply (auto simp: inner_diff_left le_Suc_eq)
-        by (metis "*")
+        by (smt (verit, ccfv_threshold) inner_diff_left less_one order_le_less)
     qed
     also have "\<dots> \<le> norm (f y - f x) + norm (y - x)"
       by (simp add: add_mono i(2) norm_bound_Basis_le)
@@ -1980,9 +1985,8 @@
   then obtain a where "a \<in> S" "a \<notin> T" by blast
   define g where "g \<equiv> \<lambda>z. if z \<in> closure S then f z else z"
   have "continuous_on (closure S \<union> closure(-S)) g"
-    unfolding g_def
-    apply (rule continuous_on_cases)
-    using fros fid frontier_closures by (auto simp: contf)
+    unfolding g_def using fros fid frontier_closures
+    by (intro continuous_on_cases) (auto simp: contf)
   moreover have "closure S \<union> closure(- S) = UNIV"
     using closure_Un by fastforce
   ultimately have contg: "continuous_on UNIV g" by metis
@@ -2027,10 +2031,8 @@
 proof -
   have "\<exists>d. 0 < d \<and> (a + d *\<^sub>R l) \<in> rel_frontier S \<and>
             (\<forall>e. 0 \<le> e \<and> e < d \<longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S)"
-       if "(a + l) \<in> affine hull S" "l \<noteq> 0" for l
-    apply (rule ray_to_rel_frontier [OF \<open>bounded S\<close> arelS])
-    apply (rule that)+
-    by metis
+    if "(a + l) \<in> affine hull S" "l \<noteq> 0" for l
+    using ray_to_rel_frontier [OF \<open>bounded S\<close> arelS] that by metis
   then obtain dd
     where dd1: "\<And>l. \<lbrakk>(a + l) \<in> affine hull S; l \<noteq> 0\<rbrakk> \<Longrightarrow> 0 < dd l \<and> (a + dd l *\<^sub>R l) \<in> rel_frontier S"
       and dd2: "\<And>l e. \<lbrakk>(a + l) \<in> affine hull S; e < dd l; 0 \<le> e; l \<noteq> 0\<rbrakk>
@@ -2071,11 +2073,11 @@
         then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \<in> rel_frontier S"
           using dd1 by auto
         moreover have "a + dd x *\<^sub>R x \<in> open_segment a (a + k *\<^sub>R x)"
-          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: that vector_add_divide_simps algebra_simps)
-          done
+          unfolding in_segment
+        proof (intro conjI exI)
+          show "a + dd x *\<^sub>R x = (1 - dd x / k) *\<^sub>R a + (dd x / k) *\<^sub>R (a + k *\<^sub>R x)"
+            using k by (simp add: that algebra_simps)
+        qed (use \<open>x \<noteq> 0\<close> \<open>0 < dd x\<close> that in auto)
         ultimately show ?thesis
           using segsub by (auto simp: rel_frontier_def)
       qed
@@ -2094,47 +2096,45 @@
     by (blast intro: continuous_on_subset)
   show ?thesis
   proof
-    show "homotopic_with_canon (\<lambda>x. True) (T - {a}) (T - {a}) id (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
+    show "homotopic_with_canon (\<lambda>x. True) (T - {a}) (T - {a}) id (\<lambda>x. a + dd (x-a) *\<^sub>R (x-a))"
     proof (rule homotopic_with_linear)
       show "continuous_on (T - {a}) id"
         by (intro continuous_intros continuous_on_compose)
-      show "continuous_on (T - {a}) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
+      show "continuous_on (T - {a}) (\<lambda>x. a + dd (x-a) *\<^sub>R (x-a))"
         using contdd by (simp add: o_def)
-      show "closed_segment (id x) (a + dd (x - a) *\<^sub>R (x - a)) \<subseteq> T - {a}"
+      show "closed_segment (id x) (a + dd (x-a) *\<^sub>R (x-a)) \<subseteq> T - {a}"
            if "x \<in> T - {a}" for x
       proof (clarsimp simp: in_segment, intro conjI)
         fix u::real assume u: "0 \<le> u" "u \<le> 1"
-        have "a + dd (x - a) *\<^sub>R (x - a) \<in> T"
+        have "a + dd (x-a) *\<^sub>R (x-a) \<in> T"
           by (metis DiffD1 DiffD2 add.commute add.right_neutral affS dd1 diff_add_cancel relS singletonI subsetCE that)
-        then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<in> T"
+        then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x-a) *\<^sub>R (x-a)) \<in> T"
           using convexD [OF \<open>convex T\<close>] that u by simp
-        have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x - a)) = a \<longleftrightarrow>
-                  (1 - u + u * d) *\<^sub>R (x - a) = 0" for d
+        have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x-a)) = a \<longleftrightarrow>
+                  (1 - u + u * d) *\<^sub>R (x-a) = 0" for d
           by (auto simp: algebra_simps)
         have "x \<in> T" "x \<noteq> a" using that by auto
-        then have axa: "a + (x - a) \<in> affine hull S"
+        then have axa: "a + (x-a) \<in> affine hull S"
            by (metis (no_types) add.commute affS diff_add_cancel rev_subsetD)
-        then have "\<not> dd (x - a) \<le> 0 \<and> a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S"
+        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"
+        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"
           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))"
+    show "retraction (T - {a}) (rel_frontier S) (\<lambda>x. a + dd (x-a) *\<^sub>R (x-a))"
     proof (simp add: retraction_def, intro conjI ballI)
       show "rel_frontier S \<subseteq> T - {a}"
         using arelS relS rel_frontier_def by fastforce
-      show "continuous_on (T - {a}) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
+      show "continuous_on (T - {a}) (\<lambda>x. a + dd (x-a) *\<^sub>R (x-a))"
         using contdd by (simp add: o_def)
-      show "(\<lambda>x. a + dd (x - a) *\<^sub>R (x - a)) \<in> (T - {a}) \<rightarrow> rel_frontier S"
-        apply (auto simp: rel_frontier_def)
-        apply (metis Diff_subset add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def subset_iff)
-        by (metis DiffE add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def rev_subsetD)
-      show "a + dd (x - a) *\<^sub>R (x - a) = x" if x: "x \<in> rel_frontier S" for x
+      show "(\<lambda>x. a + dd (x-a) *\<^sub>R (x-a)) \<in> (T - {a}) \<rightarrow> rel_frontier S"
+        unfolding Pi_iff using affS dd1 subset_eq by force
+      show "a + dd (x-a) *\<^sub>R (x-a) = x" if x: "x \<in> rel_frontier S" for x
       proof -
         have "x \<noteq> a"
           using that arelS by (auto simp: rel_frontier_def)
-        have False if "dd (x - a) < 1"
+        have False if "dd (x-a) < 1"
         proof -
           have "x \<in> closure S"
             using x by (auto simp: rel_frontier_def)
@@ -2142,21 +2142,21 @@
             by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS])
           have  xaffS: "x \<in> affine hull S"
             using affS relS x by auto
-          then have "0 < dd (x - a)" and inS: "a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S"
+          then have "0 < dd (x-a)" and inS: "a + dd (x-a) *\<^sub>R (x-a) \<in> rel_frontier S"
             using dd1 by (auto simp: \<open>x \<noteq> a\<close>)
-          moreover have "a + dd (x - a) *\<^sub>R (x - a) \<in> open_segment a x"
-            using  \<open>x \<noteq> a\<close> \<open>0 < dd (x - a)\<close>
-            apply (simp add: in_segment)
-            apply (rule_tac x = "dd (x - a)" in exI)
-            apply (simp add: algebra_simps that)
-            done
+          moreover have "a + dd (x-a) *\<^sub>R (x-a) \<in> open_segment a x"
+            unfolding in_segment
+          proof (intro exI conjI)
+            show "a + dd (x-a) *\<^sub>R (x-a) = (1 - dd (x-a)) *\<^sub>R a + (dd (x-a)) *\<^sub>R x"
+              by (simp add: algebra_simps)
+          qed (use  \<open>x \<noteq> a\<close> \<open>0 < dd (x-a)\<close> that in auto)
           ultimately show ?thesis
             using segsub by (auto simp: rel_frontier_def)
         qed
-        moreover have False if "1 < dd (x - a)"
+        moreover have False if "1 < dd (x-a)"
           using x that dd2 [of "x - a" 1] \<open>x \<noteq> a\<close> closure_affine_hull
           by (auto simp: rel_frontier_def)
-        ultimately have "dd (x - a) = 1" \<comment> \<open>similar to another proof above\<close>
+        ultimately have "dd (x-a) = 1" \<comment> \<open>similar to another proof above\<close>
           by fastforce
         with that show ?thesis
           by (simp add: rel_frontier_def)
@@ -2211,13 +2211,13 @@
 subsubsection\<open>Borsuk-style characterization of separation\<close>
 
 lemma continuous_on_Borsuk_map:
-   "a \<notin> S \<Longrightarrow>  continuous_on S (\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a))"
+   "a \<notin> S \<Longrightarrow>  continuous_on S (\<lambda>x. inverse(norm (x-a)) *\<^sub>R (x-a))"
 by (rule continuous_intros | force)+
 
 lemma Borsuk_map_into_sphere:
-   "(\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a)) \<in> S \<rightarrow> sphere 0 1 \<longleftrightarrow> (a \<notin> S)"
+   "(\<lambda>x. inverse(norm (x-a)) *\<^sub>R (x-a)) \<in> S \<rightarrow> sphere 0 1 \<longleftrightarrow> (a \<notin> S)"
 proof -
-  have "\<And>x. \<lbrakk>a \<notin> S; x \<in> S\<rbrakk> \<Longrightarrow> inverse (norm (x - a)) * norm (x - a) = 1"
+  have "\<And>x. \<lbrakk>a \<notin> S; x \<in> S\<rbrakk> \<Longrightarrow> inverse (norm (x-a)) * norm (x-a) = 1"
     by (metis left_inverse norm_eq_zero right_minus_eq)
   then show ?thesis
     by force
@@ -2226,16 +2226,19 @@
 lemma Borsuk_maps_homotopic_in_path_component:
   assumes "path_component (- S) a b"
     shows "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1)
-                   (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a))
+                   (\<lambda>x. inverse(norm(x-a)) *\<^sub>R (x-a))
                    (\<lambda>x. inverse(norm(x - b)) *\<^sub>R (x - b))"
 proof -
-  obtain g where "path g" "path_image g \<subseteq> -S" "pathstart g = a" "pathfinish g = b"
+  obtain g where g: "path g" "path_image g \<subseteq> -S" "pathstart g = a" "pathfinish g = b"
     using assms by (auto simp: path_component_def)
-  then show ?thesis
-    apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def)
-    apply (rule_tac x = "\<lambda>z. inverse(norm(snd z - (g \<circ> fst)z)) *\<^sub>R (snd z - (g \<circ> fst)z)" in exI)
-    apply (rule conjI continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+
-    done
+  define h where "h \<equiv> \<lambda>z. (snd z - (g \<circ> fst) z) /\<^sub>R norm (snd z - (g \<circ> fst) z)"
+  have "continuous_on ({0..1} \<times> S) h"
+    unfolding h_def using g by (intro continuous_intros) (auto simp: path_defs)
+  moreover
+  have "h ` ({0..1} \<times> S) \<subseteq> sphere 0 1"
+    unfolding h_def using g  by (auto simp: divide_simps path_defs)
+  ultimately show ?thesis
+    using g by (auto simp: h_def path_defs homotopic_with_def)
 qed
 
 lemma non_extensible_Borsuk_map:
@@ -2243,7 +2246,7 @@
   assumes "compact S" and cin: "C \<in> components(- S)" and boc: "bounded C" and "a \<in> C"
     shows "\<not> (\<exists>g. continuous_on (S \<union> C) g \<and>
                   g \<in> (S \<union> C) \<rightarrow> sphere 0 1 \<and>
-                  (\<forall>x \<in> S. g x = inverse(norm(x - a)) *\<^sub>R (x - a)))"
+                  (\<forall>x \<in> S. g x = inverse(norm(x-a)) *\<^sub>R (x-a)))"
 proof -
   have "closed S" using assms by (simp add: compact_imp_closed)
   have "C \<subseteq> -S"
@@ -2258,7 +2261,7 @@
   { fix g
     assume "continuous_on (S \<union> C) g"
             "g \<in> (S \<union> C) \<rightarrow> sphere 0 1"
-       and [simp]: "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)"
+       and [simp]: "\<And>x. x \<in> S \<Longrightarrow> g x = (x-a) /\<^sub>R norm (x-a)"
     then have norm_g1[simp]: "\<And>x. x \<in> S \<union> C \<Longrightarrow> norm (g x) = 1"
       by force
     have cb_eq: "cball a r = (S \<union> connected_component_set (- S) a) \<union>
@@ -2269,12 +2272,12 @@
       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)))"
+            (\<lambda>x. a + r *\<^sub>R ((x-a) /\<^sub>R norm (x-a)))"
       by (rule continuous_intros | force simp: \<open>a \<notin> S\<close>)+
     have 1: "continuous_on (cball a r)
              (\<lambda>x. if connected_component (- S) a x
                   then a + r *\<^sub>R g x
-                  else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))"
+                  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 \<open>closed S\<close> ceq cin 
@@ -2285,7 +2288,7 @@
     have "retraction (cball a r) (sphere a r)
             (\<lambda>x. if x \<in> connected_component_set (- S) a
                  then a + r *\<^sub>R g x
-                 else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))"
+                 else a + r *\<^sub>R ((x-a) /\<^sub>R norm (x-a)))"
       using  \<open>0 < r\<close> \<open>a \<notin> S\<close> \<open>a \<in> C\<close> r
       by (auto simp: norm_minus_commute retraction_def Pi_iff ceq dist_norm abs_if 
           mult_less_0_iff divide_simps 1 2)
@@ -2294,7 +2297,7 @@
              [OF \<open>0 < r\<close>, of a, unfolded retract_of_def, simplified, rule_format,
               of "\<lambda>x. if x \<in> connected_component_set (- S) a
                       then a + r *\<^sub>R g x
-                      else a + r *\<^sub>R inverse(norm(x - a)) *\<^sub>R (x - a)"]
+                      else a + r *\<^sub>R inverse(norm(x-a)) *\<^sub>R (x-a)"]
       by blast
   }
   then show ?thesis
@@ -2305,10 +2308,8 @@
 
 lemma brouwer_surjective:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
-  assumes "compact T"
-    and "convex T"
-    and "T \<noteq> {}"
-    and "continuous_on T f"
+  assumes T: "compact T" "convex T" "T \<noteq> {}"
+    and f: "continuous_on T f"
     and "\<And>x y. \<lbrakk>x\<in>S; y\<in>T\<rbrakk> \<Longrightarrow> x + (y - f y) \<in> T"
     and "x \<in> S"
   shows "\<exists>y\<in>T. f y = x"
@@ -2317,11 +2318,10 @@
     by (auto simp add: algebra_simps)
   show ?thesis
     unfolding *
-    apply (rule brouwer[OF assms(1-3), of "\<lambda>y. x + (y - f y)"])
-    apply (intro continuous_intros)
-    using assms
-    apply auto
-    done
+  proof (rule brouwer[OF T])
+    show "continuous_on T (\<lambda>y. x + (y - f y))"
+      by (intro continuous_intros f)
+  qed (use assms in auto)
 qed
 
 lemma brouwer_surjective_cball:
--- a/src/HOL/Data_Structures/Selection.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Data_Structures/Selection.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -310,6 +310,12 @@
         if y < x then (y # ls, es, gs) else if x = y then (ls, y # es, gs) else (ls, es, y # gs))"
   by (auto simp: partition3_def)
 
+lemma length_partition3:
+  assumes "partition3 x xs = (ls, es, gs)"
+  shows   "length xs = length ls + length es + length gs"
+  using assms by (induction xs arbitrary: ls es gs)
+                 (auto simp: partition3_code split: if_splits prod.splits)
+
 lemma sort_append:
   assumes "\<forall>x\<in>set xs. \<forall>y\<in>set ys. x \<le> y"
   shows   "sort (xs @ ys) = sort xs @ sort ys"
@@ -634,30 +640,51 @@
 lemma T_partition3_eq: "T_partition3 x xs = length xs + 1"
   by (induction x xs rule: T_partition3.induct) auto
 
-definition T_slow_select :: "nat \<Rightarrow> 'a :: linorder list \<Rightarrow> nat" where
-  "T_slow_select k xs = T_insort xs + T_nth (insort xs) k + 1"
+
+time_definition slow_select
+
+lemmas T_slow_select_def [simp del] = T_slow_select.simps
+
 
 definition T_slow_median :: "'a :: linorder list \<Rightarrow> nat" where
-  "T_slow_median xs = T_slow_select ((length xs - 1) div 2) xs + 1"
+  "T_slow_median xs = T_length xs + T_slow_select ((length xs - 1) div 2) xs"
 
-lemma T_slow_select_le: "T_slow_select k xs \<le> length xs ^ 2 + 3 * length xs + 3"
+lemma T_slow_select_le:
+  assumes "k < length xs"
+  shows   "T_slow_select k xs \<le> length xs ^ 2 + 3 * length xs + 1"
 proof -
-  have "T_slow_select k xs \<le> (length xs + 1)\<^sup>2 + (length (insort xs) + 1) + 1"
-    unfolding T_slow_select_def
-    by (intro add_mono T_insort_length) (auto simp: T_nth_eq)
-  also have "\<dots> = length xs ^ 2 + 3 * length xs + 3"
-    by (simp add: insort_correct algebra_simps power2_eq_square)
-  finally show ?thesis .
+  have "T_slow_select k xs = T_insort xs + T_nth (Sorting.insort xs) k"
+    unfolding T_slow_select_def ..
+  also have "T_insort xs \<le> (length xs + 1) ^ 2"
+    by (rule T_insort_length)
+  also have "T_nth (Sorting.insort xs) k = k + 1"
+    using assms by (subst T_nth_eq) (auto simp: length_insort)
+  also have "k + 1 \<le> length xs"
+    using assms by linarith
+  also have "(length xs + 1) ^ 2 + length xs = length xs ^ 2 + 3 * length xs + 1"
+    by (simp add: algebra_simps power2_eq_square)
+  finally show ?thesis by - simp_all
 qed
 
-lemma T_slow_median_le: "T_slow_median xs \<le> length xs ^ 2 + 3 * length xs + 4"
-  unfolding T_slow_median_def using T_slow_select_le[of "(length xs - 1) div 2" xs] by simp
+lemma T_slow_median_le:
+  assumes "xs \<noteq> []"
+  shows   "T_slow_median xs \<le> length xs ^ 2 + 4 * length xs + 2"
+proof -
+  have "T_slow_median xs = length xs + T_slow_select ((length xs - 1) div 2) xs + 1"
+    by (simp add: T_slow_median_def T_length_eq)
+  also from assms have "length xs > 0"
+    by simp
+  hence "(length xs - 1) div 2 < length xs"
+    by linarith
+  hence "T_slow_select ((length xs - 1) div 2) xs \<le> length xs ^ 2 + 3 * length xs + 1"
+    by (intro T_slow_select_le) auto
+  also have "length xs + \<dots> + 1 = length xs ^ 2 + 4 * length xs + 2"
+    by (simp add: algebra_simps)
+  finally show ?thesis by - simp_all
+qed
 
 
-fun T_chop :: "nat \<Rightarrow> 'a list \<Rightarrow> nat" where
-  "T_chop 0 _  = 1"
-| "T_chop _ [] = 1"
-| "T_chop n xs = T_take n xs + T_drop n xs + T_chop n (drop n xs)"
+time_fun chop
 
 lemmas [simp del] = T_chop.simps
 
@@ -668,19 +695,20 @@
   by (auto simp: T_chop.simps)
 
 lemma T_chop_reduce:
-  "n > 0 \<Longrightarrow> xs \<noteq> [] \<Longrightarrow> T_chop n xs = T_take n xs + T_drop n xs + T_chop n (drop n xs)"
+  "n > 0 \<Longrightarrow> xs \<noteq> [] \<Longrightarrow> T_chop n xs = T_take n xs + T_drop n xs + T_chop n (drop n xs) + 1"
   by (cases n; cases xs) (auto simp: T_chop.simps)
 
 lemma T_chop_le: "T_chop d xs \<le> 5 * length xs + 1"
   by (induction d xs rule: T_chop.induct) (auto simp: T_chop_reduce T_take_eq T_drop_eq)
 
+
 text \<open>
   The option \<open>domintros\<close> here allows us to explicitly reason about where the function does and
   does not terminate. With this, we can skip the termination proof this time because we can
   reuse the one for \<^const>\<open>mom_select\<close>.
 \<close>
 function (domintros) T_mom_select :: "nat \<Rightarrow> 'a :: linorder list \<Rightarrow> nat" where
-  "T_mom_select k xs = (
+  "T_mom_select k xs = T_length xs + (
      if length xs \<le> 20 then
        T_slow_select k xs
      else
@@ -693,10 +721,9 @@
            ne = length es
        in
          (if k < nl then T_mom_select k ls 
-          else if k < nl + ne then 0
-          else T_mom_select (k - nl - ne) gs) +
+          else T_length es + (if k < nl + ne then 0 else T_mom_select (k - nl - ne) gs)) +
          T_mom_select idx ms + T_chop 5 xs + T_map T_slow_median xss +
-         T_partition3 x xs + T_length ls + T_length es + 1
+         T_partition3 x xs + T_length ls + 1
       )"
   by auto
 
@@ -716,16 +743,16 @@
 function T'_mom_select :: "nat \<Rightarrow> nat" where
   "T'_mom_select n =
      (if n \<le> 20 then
-        463
+        482
       else
-        T'_mom_select (nat \<lceil>0.2*n\<rceil>) + T'_mom_select (nat \<lceil>0.7*n+3\<rceil>) + 17 * n + 50)"
+        T'_mom_select (nat \<lceil>0.2*n\<rceil>) + T'_mom_select (nat \<lceil>0.7*n+3\<rceil>) + 19 * n + 54)"
   by force+
 termination by (relation "measure id"; simp; linarith)
 
 lemmas [simp del] = T'_mom_select.simps
 
 
-lemma T'_mom_select_ge: "T'_mom_select n \<ge> 463"
+lemma T'_mom_select_ge: "T'_mom_select n \<ge> 482"
   by (induction n rule: T'_mom_select.induct; subst T'_mom_select.simps) auto
 
 lemma T'_mom_select_mono:
@@ -735,7 +762,7 @@
   show ?case
   proof (cases "m \<le> 20")
     case True
-    hence "T'_mom_select m = 463"
+    hence "T'_mom_select m = 482"
       by (subst T'_mom_select.simps) auto
     also have "\<dots> \<le> T'_mom_select n"
       by (rule T'_mom_select_ge)
@@ -743,9 +770,9 @@
   next
     case False
     hence "T'_mom_select m =
-             T'_mom_select (nat \<lceil>0.2*m\<rceil>) + T'_mom_select (nat \<lceil>0.7*m + 3\<rceil>) + 17 * m + 50"
+             T'_mom_select (nat \<lceil>0.2*m\<rceil>) + T'_mom_select (nat \<lceil>0.7*m + 3\<rceil>) + 19 * m + 54"
       by (subst T'_mom_select.simps) auto
-    also have "\<dots> \<le> T'_mom_select (nat \<lceil>0.2*n\<rceil>) + T'_mom_select (nat \<lceil>0.7*n + 3\<rceil>) + 17 * n + 50"
+    also have "\<dots> \<le> T'_mom_select (nat \<lceil>0.2*n\<rceil>) + T'_mom_select (nat \<lceil>0.7*n + 3\<rceil>) + 19 * n + 54"
       using \<open>m \<le> n\<close> and False by (intro add_mono less.IH; linarith)
     also have "\<dots> = T'_mom_select n"
       using \<open>m \<le> n\<close> and False by (subst T'_mom_select.simps) auto
@@ -753,7 +780,10 @@
   qed
 qed
 
-lemma T_mom_select_le_aux: "T_mom_select k xs \<le> T'_mom_select (length xs)"
+lemma T_mom_select_le_aux:
+  assumes "k < length xs"
+  shows   "T_mom_select k xs \<le> T'_mom_select (length xs)"
+  using assms
 proof (induction k xs rule: T_mom_select.induct)
   case (1 k xs)
   define n where [simp]: "n = length xs"
@@ -770,11 +800,12 @@
   show ?case
   proof (cases "length xs \<le> 20")
     case True \<comment> \<open>base case\<close>
-    hence "T_mom_select k xs \<le> (length xs)\<^sup>2 + 3 * length xs + 3"
-      using T_slow_select_le[of k xs] by (subst T_mom_select.simps) auto
-    also have "\<dots> \<le> 20\<^sup>2 + 3 * 20 + 3"
+    hence "T_mom_select k xs \<le> (length xs)\<^sup>2 + 4 * length xs + 2"
+      using T_slow_select_le[of k xs] \<open>k < length xs\<close>
+      by (subst T_mom_select.simps) (auto simp: T_length_eq)
+    also have "\<dots> \<le> 20\<^sup>2 + 4 * 20 + 2"
       using True by (intro add_mono power_mono) auto
-    also have "\<dots> \<le> 463"
+    also have "\<dots> = 482"
       by simp
     also have "\<dots> = T'_mom_select (length xs)"
       using True by (simp add: T'_mom_select.simps)
@@ -793,34 +824,36 @@
 
     text \<open>The cost of computing the medians of all the subgroups:\<close>
     define T_ms where "T_ms = T_map T_slow_median (chop 5 xs)"
-    have "T_ms \<le> 9 * n + 45"
+    have "T_ms \<le> 10 * n + 48"
     proof -
       have "T_ms = (\<Sum>ys\<leftarrow>chop 5 xs. T_slow_median ys) + length (chop 5 xs) + 1"
         by (simp add: T_ms_def T_map_eq)
-      also have "(\<Sum>ys\<leftarrow>chop 5 xs. T_slow_median ys) \<le> (\<Sum>ys\<leftarrow>chop 5 xs. 44)"
+      also have "(\<Sum>ys\<leftarrow>chop 5 xs. T_slow_median ys) \<le> (\<Sum>ys\<leftarrow>chop 5 xs. 47)"
       proof (intro sum_list_mono)
         fix ys assume "ys \<in> set (chop 5 xs)"
-        hence "length ys \<le> 5"
-          using length_chop_part_le by blast
-        have "T_slow_median ys \<le> (length ys) ^ 2 + 3 * length ys + 4"
+        hence "length ys \<le> 5" "ys \<noteq> []"
+          using length_chop_part_le[of ys 5 xs] by auto
+        from \<open>ys \<noteq> []\<close> have "T_slow_median ys \<le> (length ys) ^ 2 + 4 * length ys + 2"
           by (rule T_slow_median_le)
-        also have "\<dots> \<le> 5 ^ 2 + 3 * 5 + 4"
+        also have "\<dots> \<le> 5 ^ 2 + 4 * 5 + 2"
           using \<open>length ys \<le> 5\<close> by (intro add_mono power_mono) auto
-        finally show "T_slow_median ys \<le> 44" by simp
+        finally show "T_slow_median ys \<le> 47" by simp
       qed
-      also have "(\<Sum>ys\<leftarrow>chop 5 xs. 44) + length (chop 5 xs) + 1 =
-                   45 * nat \<lceil>real n / 5\<rceil> + 1"
+      also have "(\<Sum>ys\<leftarrow>chop 5 xs. 47) + length (chop 5 xs) + 1 =
+                   48 * nat \<lceil>real n / 5\<rceil> + 1"
         by (simp add: map_replicate_const length_chop)
-      also have "\<dots> \<le> 9 * n + 45"
+      also have "\<dots> \<le> 10 * n + 48"
         by linarith
-      finally show "T_ms \<le> 9 * n + 45" by simp
+      finally show "T_ms \<le> 10 * n + 48" by simp
     qed
 
     text \<open>The cost of the first recursive call (to compute the median of medians):\<close>
     define T_rec1 where
       "T_rec1 = T_mom_select (((length xs + 4) div 5 - 1) div 2) (map slow_median (chop 5 xs))"
-    have "T_rec1 \<le> T'_mom_select (length (map slow_median (chop 5 xs)))"
-      using False unfolding T_rec1_def by (intro IH(3)) auto
+    from False have "((length xs + 4) div 5 - Suc 0) div 2 < nat \<lceil>real (length xs) / 5\<rceil>"
+      by linarith
+    hence "T_rec1 \<le> T'_mom_select (length (map slow_median (chop 5 xs)))"
+      using False unfolding T_rec1_def by (intro IH(3)) (auto simp: length_chop)
     hence "T_rec1 \<le> T'_mom_select (nat \<lceil>0.2 * n\<rceil>)"
       by (simp add: length_chop)
 
@@ -854,7 +887,17 @@
       hence "T_rec2 = T_mom_select (k - nl - ne) gs"
         by (simp add: T_rec2_def)
       also have "\<dots> \<le> T'_mom_select (length gs)"
-        unfolding nl_def ne_def by (rule IH(2)) (use \<open>k \<ge> nl + ne\<close> False in \<open>auto simp: defs\<close>)
+        unfolding nl_def ne_def 
+      proof (rule IH(2))
+        show "\<not> length xs \<le> 20"
+          using False by auto
+        show "\<not> k < length ls" "\<not>k < length ls + length es"
+          using \<open>k \<ge> nl + ne\<close> by (auto simp: nl_def ne_def)
+        have "length xs = nl + ne + length gs"
+          unfolding defs by (rule length_partition3) (simp_all add: partition3_def)
+        thus "k - length ls - length es < length gs"
+          using \<open>k \<ge> nl + ne\<close> \<open>k < length xs\<close> by (auto simp: nl_def ne_def)
+      qed
       also have "length gs \<le> nat \<lceil>0.7 * n + 3\<rceil>"
         unfolding gs_def using size_greater_than_median_of_medians[of xs]
         by (auto simp: length_filter_conv_size_filter_mset slow_median_correct[abs_def] x_eq)
@@ -864,19 +907,19 @@
     qed
 
     text \<open>Now for the final inequality chain:\<close>
-    have "T_mom_select k xs = T_rec2 + T_rec1 + T_ms + n + nl + ne + T_chop 5 xs + 4" using False
+    have "T_mom_select k xs \<le> T_rec2 + T_rec1 + T_ms + 2 * n + nl + ne + T_chop 5 xs + 5" using False
       by (subst T_mom_select.simps, unfold Let_def tw [symmetric] defs [symmetric])
          (simp_all add: nl_def ne_def T_rec1_def T_rec2_def T_partition3_eq
                         T_length_eq T_ms_def)
     also have "nl \<le> n" by (simp add: nl_def ls_def)
     also have "ne \<le> n" by (simp add: ne_def es_def)
-    also note \<open>T_ms \<le> 9 * n + 45\<close>
+    also note \<open>T_ms \<le> 10 * n + 48\<close>
     also have "T_chop 5 xs \<le> 5 * n + 1"
       using T_chop_le[of 5 xs] by simp 
     also note \<open>T_rec1 \<le> T'_mom_select (nat \<lceil>0.2*n\<rceil>)\<close>
     also note \<open>T_rec2 \<le> T'_mom_select (nat \<lceil>0.7*n + 3\<rceil>)\<close>
     finally have "T_mom_select k xs \<le>
-                    T'_mom_select (nat \<lceil>0.7*n + 3\<rceil>) + T'_mom_select (nat \<lceil>0.2*n\<rceil>) + 17 * n + 50"
+                    T'_mom_select (nat \<lceil>0.7*n + 3\<rceil>) + T'_mom_select (nat \<lceil>0.2*n\<rceil>) + 19 * n + 54"
       by simp
     also have "\<dots> = T'_mom_select n"
       using False by (subst T'_mom_select.simps) auto
@@ -1033,7 +1076,7 @@
 lemma T'_mom_select_le': "\<exists>C\<^sub>1 C\<^sub>2. \<forall>n. T'_mom_select n \<le> C\<^sub>1 * n + C\<^sub>2"
 proof (rule akra_bazzi_light_nat)
   show "\<forall>n>20. T'_mom_select n = T'_mom_select (nat \<lceil>0.2 * n + 0\<rceil>) +
-                 T'_mom_select (nat \<lceil>0.7 * n + 3\<rceil>) + 17 * n + 50"
+                 T'_mom_select (nat \<lceil>0.7 * n + 3\<rceil>) + 19 * n + 54"
     using T'_mom_select.simps by auto
 qed auto
 
--- a/src/HOL/Data_Structures/Time_Funs.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Data_Structures/Time_Funs.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -43,12 +43,9 @@
 
 lemmas [simp del] = T_filter.simps
 
+time_fun nth
 
-fun T_nth :: "'a list \<Rightarrow> nat \<Rightarrow> nat" where
-  "T_nth [] n = 1"
-| "T_nth (x # xs) n = (case n of 0 \<Rightarrow> 1 | Suc n' \<Rightarrow> T_nth xs n' + 1)"
-
-lemma T_nth_eq: "T_nth xs n = min n (length xs) + 1"
+lemma T_nth_eq: "n < length xs \<Longrightarrow> T_nth xs n = n + 1"
   by (induction xs n rule: T_nth.induct) (auto split: nat.splits)
 
 lemmas [simp del] = T_nth.simps
--- a/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -34,43 +34,49 @@
   by (cases i) auto
 
 lemma norm_Pinj: "isnorm (Pinj i Q) \<Longrightarrow> isnorm Q"
-  by (cases i) (simp add: norm_Pinj_0_False norm_PX_0_False, cases Q, auto)
+  using isnorm.elims(2) by fastforce
 
 lemma norm_PX2: "isnorm (PX P i Q) \<Longrightarrow> isnorm Q"
-  apply (cases i)
-   apply auto
-  apply (cases P)
-    apply auto
-  subgoal for \<dots> pol2 by (cases pol2) auto
-  done
+  using isnorm.elims(1) by auto
+
+lemma norm_PX1: assumes "isnorm (PX P i Q)" shows "isnorm P"
+proof (cases P)
+  case (Pc x1)
+  then show ?thesis
+    by auto 
+qed (use assms isnorm.elims(1) in auto)
 
-lemma norm_PX1: "isnorm (PX P i Q) \<Longrightarrow> isnorm P"
-  apply (cases i)
-   apply auto
-  apply (cases P)
-    apply auto
-  subgoal for \<dots> pol2 by (cases pol2) auto
-  done
-
-lemma mkPinj_cn: "y \<noteq> 0 \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (mkPinj y Q)"
-  apply (auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split)
-   apply (rename_tac nat a, case_tac nat, auto simp add: norm_Pinj_0_False)
-   apply (rename_tac pol a, case_tac pol, auto)
-  apply (case_tac y, auto)
-  done
+lemma mkPinj_cn:
+  assumes "y \<noteq> 0" and "isnorm Q" 
+  shows "isnorm (mkPinj y Q)"
+proof (cases Q)
+  case Pc
+  with assms show ?thesis
+    by (simp add: mkPinj_def)
+next
+  case Pinj
+  with assms show ?thesis
+    using isnorm.elims(2) isnorm.simps(5) mkPinj_def by fastforce
+next
+  case PX
+  with assms show ?thesis
+  using isnorm.elims(1) mkPinj_def by auto
+qed
 
 lemma norm_PXtrans:
-  assumes A: "isnorm (PX P x Q)"
+  assumes "isnorm (PX P x Q)" and "isnorm Q2"
+  shows "isnorm (PX P x Q2)"
+  using assms isnorm.elims(3) by fastforce
+
+
+lemma norm_PXtrans2:
+  assumes "isnorm (PX P x Q)"
     and "isnorm Q2"
-  shows "isnorm (PX P x Q2)"
+  shows "isnorm (PX P (Suc (n + x)) Q2)"
 proof (cases P)
   case (PX p1 y p2)
   with assms show ?thesis
-    apply (cases x)
-     apply auto
-    apply (cases p2)
-      apply auto
-    done
+  using isnorm.elims(2) by fastforce
 next
   case Pc
   with assms show ?thesis
@@ -81,29 +87,7 @@
     by (cases x) auto
 qed
 
-lemma norm_PXtrans2:
-  assumes "isnorm (PX P x Q)"
-    and "isnorm Q2"
-  shows "isnorm (PX P (Suc (n + x)) Q2)"
-proof (cases P)
-  case (PX p1 y p2)
-  with assms show ?thesis
-    apply (cases x)
-     apply auto
-    apply (cases p2)
-      apply auto
-    done
-next
-  case Pc
-  with assms show ?thesis
-    by (cases x) auto
-next
-  case Pinj
-  with assms show ?thesis
-    by (cases x) auto
-qed
-
-text \<open>\<open>mkPX\<close> conserves normalizedness (\<open>_cn\<close>)\<close>
+text \<open>\<open>mkPX\<close> preserves the normal property (\<open>_cn\<close>)\<close>
 lemma mkPX_cn:
   assumes "x \<noteq> 0"
     and "isnorm P"
@@ -124,14 +108,23 @@
   from assms PX have "isnorm P1" "isnorm P2"
     by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2])
   from assms PX Y0 show ?thesis
-    apply (cases x)
-     apply (auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _])
-    apply (cases P2)
-      apply auto
-    done
+  proof (cases P2)
+    case (Pc x1)
+    then show ?thesis
+      using assms gr0_conv_Suc PX by (auto simp add: mkPX_def norm_PXtrans2)
+  next
+    case (Pinj x21 x22)
+    with assms gr0_conv_Suc PX show ?thesis
+      by (auto simp: mkPX_def)
+  next
+    case (PX x31 x32 x33)
+    with assms gr0_conv_Suc \<open>P = PX P1 y P2\<close>
+    show ?thesis
+      using assms PX by (auto simp add: mkPX_def norm_PXtrans2)
+  qed
 qed
 
-text \<open>\<open>add\<close> conserves normalizedness\<close>
+text \<open>\<open>add\<close> preserves the normal property\<close>
 lemma add_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<langle>+\<rangle> Q)"
 proof (induct P Q rule: add.induct)
   case 1
@@ -139,41 +132,23 @@
 next
   case (2 c i P2)
   then show ?case
-    apply (cases P2)
-      apply simp_all
-    apply (cases i)
-     apply simp_all
-    done
+    using isnorm.elims(2) by fastforce
 next
   case (3 i P2 c)
   then show ?case
-    apply (cases P2)
-      apply simp_all
-    apply (cases i)
-     apply simp_all
-    done
+    using isnorm.elims(2) by fastforce
 next
   case (4 c P2 i Q2)
   then have "isnorm P2" "isnorm Q2"
     by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
   with 4 show ?case
-    apply (cases i)
-     apply simp
-    apply (cases P2)
-      apply auto
-    subgoal for \<dots> pol2 by (cases pol2) auto
-    done
+    using isnorm.elims(2) by fastforce
 next
   case (5 P2 i Q2 c)
   then have "isnorm P2" "isnorm Q2"
     by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
   with 5 show ?case
-    apply (cases i)
-     apply simp
-    apply (cases P2)
-      apply auto
-    subgoal for \<dots> pol2 by (cases pol2) auto
-    done
+    using isnorm.elims(2) by fastforce
 next
   case (6 x P2 y Q2)
   then have Y0: "y > 0"
@@ -336,39 +311,19 @@
 next
   case (2 c i P2)
   then show ?case
-    apply (cases P2)
-      apply simp_all
-    apply (cases i)
-     apply (simp_all add: mkPinj_cn)
-    done
+    by (metis mkPinj_cn mul.simps(2) norm_Pinj norm_Pinj_0_False)
 next
   case (3 i P2 c)
   then show ?case
-    apply (cases P2)
-      apply simp_all
-    apply (cases i)
-     apply (simp_all add: mkPinj_cn)
-    done
+    by (metis mkPinj_cn mul.simps(3) norm_Pinj norm_Pinj_0_False)
 next
   case (4 c P2 i Q2)
-  then have "isnorm P2" "isnorm Q2"
-    by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
-  with 4 show ?case
-    apply (cases "c = 0")
-     apply simp_all
-    apply (cases "i = 0")
-     apply (simp_all add: mkPX_cn)
-    done
+  then show ?case
+    by (metis isnorm.simps(6) mkPX_cn mul.simps(4) norm_PX1 norm_PX2)
 next
   case (5 P2 i Q2 c)
-  then have "isnorm P2" "isnorm Q2"
-    by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
-  with 5 show ?case
-    apply (cases "c = 0")
-     apply simp_all
-    apply (cases "i = 0")
-     apply (simp_all add: mkPX_cn)
-    done
+  then show ?case
+    by (metis isnorm.simps(6) mkPX_cn mul.simps(5) norm_PX1 norm_PX2)
 next
   case (6 x P2 y Q2)
   consider "x < y" | "x = y" | "x > y" by arith
@@ -382,11 +337,8 @@
     from 6 have **: "isnorm P2" "isnorm Q2"
       by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
     from 6 xy y have "isnorm (Pinj d Q2)"
-      apply (cases d)
-       apply simp
-      apply (cases Q2)
-        apply auto
-      done
+      apply simp
+      by (smt (verit, ccfv_SIG) "**"(2) Suc_pred isnorm.elims(1) isnorm.simps(2) isnorm.simps(3) isnorm.simps(5))
     with 6 * ** y show ?thesis
       by (simp add: mkPinj_cn)
   next
@@ -405,12 +357,9 @@
       by (cases y) (auto simp add: norm_Pinj_0_False)
     from 6 have **: "isnorm P2" "isnorm Q2"
       by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
-    from 6 xy x have "isnorm (Pinj d P2)"
-      apply (cases d)
+    with 6 xy x have "isnorm (Pinj d P2)"
       apply simp
-      apply (cases P2)
-      apply auto
-      done
+      by (smt (verit, ccfv_SIG) Suc_pred isnorm.elims(1) isnorm.simps(2) isnorm.simps(3) isnorm.simps(5))
     with 6 xy * ** x show ?thesis
       by (simp add: mkPinj_cn)
   qed
@@ -493,7 +442,7 @@
     by (simp add: add_cn)
 qed
 
-text \<open>neg conserves normalizedness\<close>
+text \<open>neg preserves the normal property\<close>
 lemma neg_cn: "isnorm P \<Longrightarrow> isnorm (neg P)"
 proof (induct P)
   case Pc
@@ -520,7 +469,7 @@
   qed (cases x, auto)
 qed
 
-text \<open>sub conserves normalizedness\<close>
+text \<open>sub preserves the normal property\<close>
 lemma sub_cn: "isnorm p \<Longrightarrow> isnorm q \<Longrightarrow> isnorm (p \<langle>-\<rangle> q)"
   by (simp add: sub_def add_cn neg_cn)
 
@@ -532,11 +481,7 @@
 next
   case (Pinj i Q)
   then show ?case
-    apply (cases Q)
-      apply (auto simp add: mkPX_cn mkPinj_cn)
-    apply (cases i)
-     apply (auto simp add: mkPX_cn mkPinj_cn)
-    done
+    by (metis Commutative_Ring.sqr.simps(2) mkPinj_cn norm_Pinj norm_Pinj_0_False)
 next
   case (PX P1 x P2)
   then have "x + x \<noteq> 0" "isnorm P2" "isnorm P1"
@@ -548,7 +493,7 @@
     by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
 qed
 
-text \<open>\<open>pow\<close> conserves normalizedness\<close>
+text \<open>\<open>pow\<close> preserves the normal property\<close>
 lemma pow_cn: "isnorm P \<Longrightarrow> isnorm (pow n P)"
 proof (induct n arbitrary: P rule: less_induct)
   case (less k)
--- a/src/HOL/Decision_Procs/MIR.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -73,10 +73,7 @@
 qed
 
 lemma rdvd_minus: "(real_of_int (d::int) rdvd t) = (real_of_int d rdvd -t)"
-  apply (auto simp add: rdvd_def)
-  apply (rule_tac x="-k" in exI, simp)
-  apply (rule_tac x="-k" in exI, simp)
-  done
+  by (metis equation_minus_iff mult_minus_right of_int_minus rdvd_def)
 
 lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)"
   by (auto simp add: rdvd_def)
@@ -459,24 +456,18 @@
   by (induct ps) (simp_all add: evaldjf_def djf_Or)
 
 lemma evaldjf_bound0:
-  assumes nb: "\<forall> x\<in> set xs. bound0 (f x)"
+  assumes "\<forall>x \<in> set xs. bound0 (f x)"
   shows "bound0 (evaldjf f xs)"
-  using nb
-  apply (induct xs)
-  apply (auto simp add: evaldjf_def djf_def Let_def)
-  apply (case_tac "f a")
-  apply auto
-  done
+  using assms
+  by (induct xs) (auto simp add: evaldjf_def djf_def Let_def split: fm.split)
+
 
 lemma evaldjf_qf:
-  assumes nb: "\<forall> x\<in> set xs. qfree (f x)"
+  assumes "\<forall>x \<in> set xs. qfree (f x)"
   shows "qfree (evaldjf f xs)"
-  using nb
-  apply (induct xs)
-  apply (auto simp add: evaldjf_def djf_def Let_def)
-  apply (case_tac "f a")
-  apply auto
-  done
+  using assms
+  by (induct xs) (auto simp add: evaldjf_def djf_def Let_def split: fm.split)
+
 
 fun disjuncts :: "fm \<Rightarrow> fm list"
 where
@@ -1750,7 +1741,9 @@
   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))"
       by (simp add: nb Let_def split_def isint_Floor isint_neg)
     have "?I (Lt a) = (real_of_int (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
-    also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def)
+    also have "\<dots> = (?I (?l (Lt a)))"
+      unfolding split_int_less_real'[where a="?c*i" and b="?N ?r"]
+      by (simp add: Ia cp cnz Let_def split_def)
     finally have ?case using l by simp}
   moreover
   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))"
@@ -2761,28 +2754,27 @@
   shows "\<forall> b\<in> set (\<beta> p). isint b bs"
 using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub)
 
-lemma cpmi_eq: "0 < D \<Longrightarrow> (\<exists>z::int. \<forall>x. x < z \<longrightarrow> (P x = P1 x))
-\<Longrightarrow> \<forall>x. \<not>(\<exists>(j::int) \<in> {1..D}. \<exists>(b::int) \<in> B. P(b+j)) \<longrightarrow> P (x) \<longrightarrow> P (x - D)
-\<Longrightarrow> (\<forall>(x::int). \<forall>(k::int). ((P1 x)= (P1 (x-k*D))))
-\<Longrightarrow> (\<exists>(x::int). P(x)) = ((\<exists>(j::int) \<in> {1..D} . (P1(j))) | (\<exists>(j::int) \<in> {1..D}. \<exists>(b::int) \<in> B. P (b+j)))"
-apply(rule iffI)
-prefer 2
-apply(drule minusinfinity)
-apply assumption+
-apply(fastforce)
-apply clarsimp
-apply(subgoal_tac "\<And>k. 0<=k \<Longrightarrow> \<forall>x. P x \<longrightarrow> P (x - k*D)")
-apply(frule_tac x = x and z=z in decr_lemma)
-apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)")
-prefer 2
-apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
-prefer 2 apply arith
- apply fastforce
-apply(drule (1)  periodic_finite_ex)
-apply blast
-apply(blast dest:decr_mult_lemma)
-done
-
+lemma cpmi_eq:
+  assumes "0 < D"
+      and "\<exists>z::int. \<forall>x<z. P x = P1 x"
+      and "\<forall>x. \<not> (\<exists>j\<in>{1..D}. \<exists>b\<in>B. P (b + j)) \<longrightarrow> P x \<longrightarrow> P (x - D)"
+      and "\<forall>x k. P1 x = P1 (x - k * D)"
+    shows "(\<exists>x. P x) \<longleftrightarrow> ((\<exists>j\<in>{1..D}. P1 j) \<or> (\<exists>j\<in>{1..D}. \<exists>b\<in>B. P (b + j)))" (is "_=?R")
+proof
+  assume L: "\<exists>x. P x"
+  show "(\<exists>j\<in>{1..D}. P1 j) \<or> (\<exists>j\<in>{1..D}. \<exists>b\<in>B. P (b + j))"
+  proof clarsimp
+    assume \<section>: "\<forall>j\<in>{1..D}. \<forall>b\<in>B. \<not> P (b + j)"
+    then have "\<And>k. 0\<le>k \<Longrightarrow> \<forall>x. P x \<longrightarrow> P (x - k*D)"
+      by (simp add: assms decr_mult_lemma)
+    with L \<section> assms show "\<exists>j\<in>{1..D}. P1 j"
+      using periodic_finite_ex [OF \<open>D>0\<close>, of P1]
+      by (metis abs_1 abs_add_abs abs_ge_zero decr_lemma)
+  qed
+next
+  assume ?R then show "\<exists>x. P x"
+    using decr_lemma assms by blast
+qed
 
 theorem cp_thm:
   assumes lp: "iszlfm p (a #bs)"
@@ -3126,9 +3118,8 @@
   from 6 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric]
     of_int_mult]
   show ?case using 6 dp
-    apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"]
-      algebra_simps del: mult_pos_pos)
-      using order_trans by fastforce
+    by (fastforce simp: numbound0_I[where bs="bs" and b="i - real_of_int d" and b'="i"]
+      algebra_simps intro: order_trans)
 next
   case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)"
     and nob: "\<forall> j\<in> {1 .. c*d}. real_of_int (c*i) \<noteq> - ?N i e + real_of_int j"
@@ -3419,7 +3410,7 @@
     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s).  {?f(p,n,s) j| j. j\<in> {n .. 0}})))"
-    by blast
+    by force
   finally
   have FS: "?SS (Floor a) =
     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
@@ -3576,7 +3567,7 @@
     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s).  {?f(p,n,s) j| j. j\<in> {n .. 0}})))"
-    by blast
+    by force
   finally
   have FS: "?SS (Floor a) =
     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
@@ -4880,7 +4871,8 @@
   have MF: "?M = False"
     apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
     by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))", simp_all)
-  have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
+  have PF: "?P = False" 
+    apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
     by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))", simp_all)
   have "(\<exists> x. ?I x ?q ) =
     ((?I x (minusinf ?rq)) \<or> (?I x (plusinf ?rq )) \<or> (\<exists> (t,n) \<in> set (\<Upsilon> ?rq). \<exists> (s,m) \<in> set (\<Upsilon> ?rq ). ?I x (\<upsilon> ?rq (Add (Mul m t) (Mul n s), 2*n*m))))"
@@ -4936,9 +4928,7 @@
        \<in> (\<lambda>((t, n), s, m).
              (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) `
          (set U \<times> set U)"using mnz nnz th
-    apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def)
-    by (rule_tac x="(s,m)" in bexI,simp_all)
-  (rule_tac x="(t,n)" in bexI,simp_all add: mult.commute)
+    by (force simp add: th add_divide_distrib algebra_simps split_def image_def)
 next
   fix t n s m
   assume tnU: "(t,n) \<in> set U" and smU:"(s,m) \<in> set U"
@@ -5299,9 +5289,8 @@
   from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
 
   from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
-  from ecRo jD px' show ?rhs apply (auto simp: cc')
-    by (rule_tac x="(e', c')" in bexI,simp_all)
-  (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
+  from ecRo jD px' show ?rhs
+    using cc' by blast
 next
   let ?d = "\<delta> p"
   assume ?rhs then obtain e c j where ecR: "(e,c) \<in> set (\<rho> p)" and jD:"j \<in> {1 .. c*?d}"
@@ -5313,9 +5302,8 @@
     and cc':"c = c'" by blast
   from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
   from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
-  from ecRo jD px' show ?lhs apply (auto simp: cc')
-    by (rule_tac x="(e', c')" in bexI,simp_all)
-  (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
+  from ecRo jD px' show ?lhs
+    using cc' by blast
 qed
 
 lemma rl_thm':
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -208,7 +208,7 @@
     and nb: "tmbound m t"
     and nle: "m \<le> length bs"
   shows "Itm vs (removen m bs) (decrtm m t) = Itm vs bs t"
-  using bnd nb nle by (induct t rule: tm.induct) (auto simp add: removen_nth)
+  using bnd nb nle by (induct t rule: tm.induct) (auto simp: removen_nth)
 
 primrec tmsubst0:: "tm \<Rightarrow> tm \<Rightarrow> tm"
   where
@@ -276,25 +276,33 @@
   | "tmadd a b = Add a b"
 
 lemma tmadd [simp]: "Itm vs bs (tmadd t s) = Itm vs bs (Add t s)"
-  apply (induct t s rule: tmadd.induct)
-                      apply (simp_all add: Let_def)
-  apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p")
-   apply (case_tac "n1 \<le> n2")
-    apply simp_all
-   apply (case_tac "n1 = n2")
-    apply (simp_all add: algebra_simps)
-  apply (simp only: distrib_left [symmetric] polyadd [symmetric])
-  apply simp
-  done
+proof (induct t s rule: tmadd.induct)
+  case (1 n1 c1 r1 n2 c2 r2)
+  show ?case
+  proof (cases "c1 +\<^sub>p c2 = 0\<^sub>p")
+    case 0: True
+    show ?thesis
+    proof (cases "n1 \<le> n2")
+      case True
+      with 0 show ?thesis
+        apply (simp add: 1)
+        by (metis INum_int(2) Ipoly.simps(1) comm_semiring_class.distrib mult_eq_0_iff polyadd)
+    qed (use 0 1 in auto)
+  next
+    case False
+    then show ?thesis
+      using 1 comm_semiring_class.distrib by auto
+  qed
+qed auto
 
 lemma tmadd_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmadd t s)"
-  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
+  by (induct t s rule: tmadd.induct) (auto simp: Let_def)
 
 lemma tmadd_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n s \<Longrightarrow> tmbound n (tmadd t s)"
-  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
+  by (induct t s rule: tmadd.induct) (auto simp: Let_def)
 
 lemma tmadd_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n s \<Longrightarrow> tmboundslt n (tmadd t s)"
-  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
+  by (induct t s rule: tmadd.induct) (auto simp: Let_def)
 
 lemma tmadd_allpolys_npoly[simp]:
   "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd t s)"
@@ -316,7 +324,7 @@
   by (induct t arbitrary: n rule: tmmul.induct) auto
 
 lemma tmmul_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmmul t i)"
-  by (induct t arbitrary: i rule: tmmul.induct) (auto simp add: Let_def)
+  by (induct t arbitrary: i rule: tmmul.induct) (auto simp: Let_def)
 
 lemma tmmul_allpolys_npoly[simp]:
   assumes "SORT_CONSTRAINT('a::field_char_0)"
@@ -381,21 +389,19 @@
 lemma polynate_stupid:
   assumes "SORT_CONSTRAINT('a::field_char_0)"
   shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a)"
-  apply (subst polynate[symmetric])
-  apply simp
-  done
+  by (metis INum_int(2) Ipoly.simps(1) polynate)
 
 lemma simptm_ci[simp]: "Itm vs bs (simptm t) = Itm vs bs t"
-  by (induct t rule: simptm.induct) (auto simp add: Let_def polynate_stupid)
+  by (induct t rule: simptm.induct) (auto simp: Let_def polynate_stupid)
 
 lemma simptm_tmbound0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (simptm t)"
-  by (induct t rule: simptm.induct) (auto simp add: Let_def)
+  by (induct t rule: simptm.induct) (auto simp: Let_def)
 
 lemma simptm_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (simptm t)"
-  by (induct t rule: simptm.induct) (auto simp add: Let_def)
+  by (induct t rule: simptm.induct) (auto simp: Let_def)
 
 lemma simptm_nlt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (simptm t)"
-  by (induct t rule: simptm.induct) (auto simp add: Let_def)
+  by (induct t rule: simptm.induct) (auto simp: Let_def)
 
 lemma [simp]: "isnpoly 0\<^sub>p"
   and [simp]: "isnpoly (C (1, 1))"
@@ -404,7 +410,7 @@
 lemma simptm_allpolys_npoly[simp]:
   assumes "SORT_CONSTRAINT('a::field_char_0)"
   shows "allpolys isnpoly (simptm p)"
-  by (induct p rule: simptm.induct) (auto simp add: Let_def)
+  by (induct p rule: simptm.induct) (auto simp: Let_def)
 
 declare let_cong[fundef_cong del]
 
@@ -422,24 +428,15 @@
 declare let_cong[fundef_cong]
 
 lemma split0_stupid[simp]: "\<exists>x y. (x, y) = split0 p"
-  apply (rule exI[where x="fst (split0 p)"])
-  apply (rule exI[where x="snd (split0 p)"])
-  apply simp
-  done
+  using prod.collapse by blast
 
 lemma split0:
   "tmbound 0 (snd (split0 t)) \<and> Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t"
-  apply (induct t rule: split0.induct)
-          apply simp
-         apply (simp add: Let_def split_def field_simps)
-        apply (simp add: Let_def split_def field_simps)
-       apply (simp add: Let_def split_def field_simps)
-      apply (simp add: Let_def split_def field_simps)
-     apply (simp add: Let_def split_def field_simps)
-    apply (simp add: Let_def split_def mult.assoc distrib_left[symmetric])
-   apply (simp add: Let_def split_def field_simps)
-  apply (simp add: Let_def split_def field_simps)
-  done
+proof (induct t rule: split0.induct)
+  case (7 c t)
+  then show ?case
+    by (simp add: Let_def split_def mult.assoc flip: distrib_left)
+qed (auto simp: Let_def split_def field_simps)
 
 lemma split0_ci: "split0 t = (c',t') \<Longrightarrow> Itm vs bs t = Itm vs bs (CNP 0 c' t')"
 proof -
@@ -472,21 +469,21 @@
 lemma split0_nb:
   assumes nb: "tmbound n t"
   shows "tmbound n (snd (split0 t))"
-  using nb by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
+  using nb by (induct t rule: split0.induct) (auto simp: Let_def split_def)
 
 lemma split0_blt:
   assumes nb: "tmboundslt n t"
   shows "tmboundslt n (snd (split0 t))"
-  using nb by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
+  using nb by (induct t rule: split0.induct) (auto simp: Let_def split_def)
 
 lemma tmbound_split0: "tmbound 0 t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0"
-  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
+  by (induct t rule: split0.induct) (auto simp: Let_def split_def)
 
 lemma tmboundslt_split0: "tmboundslt n t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0 \<or> n > 0"
-  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
+  by (induct t rule: split0.induct) (auto simp: Let_def split_def)
 
 lemma tmboundslt0_split0: "tmboundslt 0 t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0"
-  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
+  by (induct t rule: split0.induct) (auto simp: Let_def split_def)
 
 lemma allpolys_split0: "allpolys isnpoly p \<Longrightarrow> allpolys isnpoly (snd (split0 p))"
   by (induct p rule: split0.induct) (auto simp  add: isnpoly_def Let_def split_def)
@@ -764,7 +761,7 @@
     from A(1)[OF bnd nb nle] show ?thesis .
   qed
   then show ?case by auto
-qed (auto simp add: decrtm removen_nth)
+qed (auto simp: decrtm removen_nth)
 
 primrec subst0 :: "tm \<Rightarrow> fm \<Rightarrow> fm"
   where
@@ -849,7 +846,7 @@
       by (simp add: incrtm0[where x="x" and bs="bs" and t="t"])
   qed
   then show ?case by simp
-qed (auto simp add: tmsubst)
+qed (auto simp: tmsubst)
 
 lemma subst_nb:
   assumes "tmbound m t"
@@ -926,36 +923,38 @@
   where "evaldjf f ps \<equiv> foldr (djf f) ps F"
 
 lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)"
-  apply (cases "q = T")
-   apply (simp add: djf_def)
-  apply (cases "q = F")
-   apply (simp add: djf_def)
-  apply (cases "f p")
-              apply (simp_all add: Let_def djf_def)
-  done
+  by (cases "f p") (simp_all add: Let_def djf_def)
 
 lemma evaldjf_ex: "Ifm vs bs (evaldjf f ps) \<longleftrightarrow> (\<exists>p \<in> set ps. Ifm vs bs (f p))"
   by (induct ps) (simp_all add: evaldjf_def djf_Or)
 
 lemma evaldjf_bound0:
-  assumes "\<forall>x\<in> set xs. bound0 (f x)"
+  assumes "\<forall>x \<in> set xs. bound0 (f x)"
   shows "bound0 (evaldjf f xs)"
   using assms
-  apply (induct xs)
-   apply (auto simp add: evaldjf_def djf_def Let_def)
-  apply (case_tac "f a")
-              apply auto
-  done
+proof (induct xs)
+  case Nil
+  then show ?case
+    by (simp add: evaldjf_def)
+next
+  case (Cons a xs)
+  then show ?case
+    by (cases "f a") (simp_all add: evaldjf_def djf_def Let_def)
+qed
 
 lemma evaldjf_qf:
   assumes "\<forall>x\<in> set xs. qfree (f x)"
   shows "qfree (evaldjf f xs)"
   using assms
-  apply (induct xs)
-   apply (auto simp add: evaldjf_def djf_def Let_def)
-  apply (case_tac "f a")
-              apply auto
-  done
+proof (induct xs)
+  case Nil
+  then show ?case
+    by (simp add: evaldjf_def)
+next
+  case (Cons a xs)
+  then show ?case
+    by (cases "f a") (simp_all add: evaldjf_def djf_def Let_def)
+qed
 
 fun disjuncts :: "fm \<Rightarrow> fm list"
   where
@@ -1100,71 +1099,32 @@
 definition "neq p = not (eq p)"
 
 lemma lt: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (lt p) = Ifm vs bs (Lt p)"
-  apply (simp add: lt_def)
-  apply (cases p)
-        apply simp_all
-  apply (rename_tac poly, case_tac poly)
-         apply (simp_all add: isnpoly_def)
-  done
+  by (auto simp: lt_def isnpoly_def split: tm.split poly.split)
 
 lemma le: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (le p) = Ifm vs bs (Le p)"
-  apply (simp add: le_def)
-  apply (cases p)
-        apply simp_all
-  apply (rename_tac poly, case_tac poly)
-         apply (simp_all add: isnpoly_def)
-  done
+  by (auto simp: le_def isnpoly_def split: tm.split poly.split)
 
 lemma eq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (eq p) = Ifm vs bs (Eq p)"
-  apply (simp add: eq_def)
-  apply (cases p)
-        apply simp_all
-  apply (rename_tac poly, case_tac poly)
-         apply (simp_all add: isnpoly_def)
-  done
+  by (auto simp: eq_def isnpoly_def split: tm.split poly.split)
 
 lemma neq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (neq p) = Ifm vs bs (NEq p)"
   by (simp add: neq_def eq)
 
 lemma lt_lin: "tmbound0 p \<Longrightarrow> islin (lt p)"
-  apply (simp add: lt_def)
-  apply (cases p)
-        apply simp_all
-   apply (rename_tac poly, case_tac poly)
-          apply simp_all
-  apply (rename_tac nat a b, case_tac nat)
-   apply simp_all
-  done
+  using islin_stupid
+  by(auto simp: lt_def isnpoly_def split: tm.split poly.split)
 
 lemma le_lin: "tmbound0 p \<Longrightarrow> islin (le p)"
-  apply (simp add: le_def)
-  apply (cases p)
-        apply simp_all
-   apply (rename_tac poly, case_tac poly)
-          apply simp_all
-  apply (rename_tac nat a b, case_tac nat)
-   apply simp_all
-  done
+  using islin_stupid
+  by(auto simp: le_def isnpoly_def split: tm.split poly.split)
 
 lemma eq_lin: "tmbound0 p \<Longrightarrow> islin (eq p)"
-  apply (simp add: eq_def)
-  apply (cases p)
-        apply simp_all
-   apply (rename_tac poly, case_tac poly)
-          apply simp_all
-  apply (rename_tac nat a b, case_tac nat)
-   apply simp_all
-  done
+  using islin_stupid
+  by(auto simp: eq_def isnpoly_def split: tm.split poly.split)
 
 lemma neq_lin: "tmbound0 p \<Longrightarrow> islin (neq p)"
-  apply (simp add: neq_def eq_def)
-  apply (cases p)
-        apply simp_all
-   apply (rename_tac poly, case_tac poly)
-          apply simp_all
-  apply (rename_tac nat a b, case_tac nat)
-   apply simp_all
-  done
+  using islin_stupid
+  by(auto simp: neq_def eq_def isnpoly_def split: tm.split poly.split)
 
 definition "simplt t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then lt s else Lt (CNP 0 c s))"
 definition "simple t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then le s else Le (CNP 0 c s))"
@@ -1176,7 +1136,7 @@
   shows "islin (simplt t)"
   unfolding simplt_def
   using split0_nb0'
-  by (auto simp add: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
+  by (auto simp: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
       islin_stupid allpolys_split0[OF simptm_allpolys_npoly])
 
 lemma simple_islin [simp]:
@@ -1184,7 +1144,7 @@
   shows "islin (simple t)"
   unfolding simple_def
   using split0_nb0'
-  by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
+  by (auto simp: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
       islin_stupid allpolys_split0[OF simptm_allpolys_npoly] le_lin)
 
 lemma simpeq_islin [simp]:
@@ -1192,7 +1152,7 @@
   shows "islin (simpeq t)"
   unfolding simpeq_def
   using split0_nb0'
-  by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
+  by (auto simp: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
       islin_stupid allpolys_split0[OF simptm_allpolys_npoly] eq_lin)
 
 lemma simpneq_islin [simp]:
@@ -1200,7 +1160,7 @@
   shows "islin (simpneq t)"
   unfolding simpneq_def
   using split0_nb0'
-  by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
+  by (auto simp: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
       islin_stupid allpolys_split0[OF simptm_allpolys_npoly] neq_lin)
 
 lemma really_stupid: "\<not> (\<forall>c1 s'. (c1, s') \<noteq> split0 s)"
@@ -1213,7 +1173,7 @@
     and "allpolys isnpoly (snd (split0 t))"
   using *
   by (induct t rule: split0.induct)
-    (auto simp add: Let_def split_def polyadd_norm polymul_norm polyneg_norm
+    (auto simp: Let_def split_def polyadd_norm polymul_norm polyneg_norm
       polysub_norm really_stupid)
 
 lemma simplt[simp]: "Ifm vs bs (simplt t) = Ifm vs bs (Lt t)"
@@ -1291,44 +1251,24 @@
 qed
 
 lemma lt_nb: "tmbound0 t \<Longrightarrow> bound0 (lt t)"
-  apply (simp add: lt_def)
-  apply (cases t)
-        apply auto
-  apply (rename_tac poly, case_tac poly)
-         apply auto
-  done
+  by(auto simp: lt_def split: tm.split poly.split)
 
 lemma le_nb: "tmbound0 t \<Longrightarrow> bound0 (le t)"
-  apply (simp add: le_def)
-  apply (cases t)
-        apply auto
-  apply (rename_tac poly, case_tac poly)
-         apply auto
-  done
+  by(auto simp: le_def split: tm.split poly.split)
 
 lemma eq_nb: "tmbound0 t \<Longrightarrow> bound0 (eq t)"
-  apply (simp add: eq_def)
-  apply (cases t)
-        apply auto
-  apply (rename_tac poly, case_tac poly)
-         apply auto
-  done
+  by(auto simp: eq_def split: tm.split poly.split)
 
 lemma neq_nb: "tmbound0 t \<Longrightarrow> bound0 (neq t)"
-  apply (simp add: neq_def eq_def)
-  apply (cases t)
-        apply auto
-  apply (rename_tac poly, case_tac poly)
-         apply auto
-  done
-
+  by(auto simp: neq_def eq_def split: tm.split poly.split)
+
+(*THE FOLLOWING PROOFS MIGHT BE COMBINED*)
 lemma simplt_nb[simp]:
-  assumes "SORT_CONSTRAINT('a::field_char_0)"
-  shows "tmbound0 t \<Longrightarrow> bound0 (simplt t)"
+  assumes "SORT_CONSTRAINT('a::field_char_0)" and t: "tmbound0 t"
+  shows "bound0 (simplt t)"
 proof (simp add: simplt_def Let_def split_def)
-  assume "tmbound0 t"
-  then have *: "tmbound0 (simptm t)"
-    by simp
+  have *: "tmbound0 (simptm t)"
+    using t by simp
   let ?c = "fst (split0 (simptm t))"
   from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
   have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
@@ -1344,12 +1284,11 @@
 qed
 
 lemma simple_nb[simp]:
-  assumes "SORT_CONSTRAINT('a::field_char_0)"
-  shows "tmbound0 t \<Longrightarrow> bound0 (simple t)"
+  assumes "SORT_CONSTRAINT('a::field_char_0)" and t: "tmbound0 t"
+  shows "bound0 (simple t)"
 proof(simp add: simple_def Let_def split_def)
-  assume "tmbound0 t"
-  then have *: "tmbound0 (simptm t)"
-    by simp
+  have *: "tmbound0 (simptm t)"
+    using t by simp
   let ?c = "fst (split0 (simptm t))"
   from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
   have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
@@ -1365,12 +1304,11 @@
 qed
 
 lemma simpeq_nb[simp]:
-  assumes "SORT_CONSTRAINT('a::field_char_0)"
-  shows "tmbound0 t \<Longrightarrow> bound0 (simpeq t)"
+  assumes "SORT_CONSTRAINT('a::field_char_0)" and t: "tmbound0 t"
+  shows "bound0 (simpeq t)"
 proof (simp add: simpeq_def Let_def split_def)
-  assume "tmbound0 t"
-  then have *: "tmbound0 (simptm t)"
-    by simp
+  have *: "tmbound0 (simptm t)"
+    using t by simp
   let ?c = "fst (split0 (simptm t))"
   from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
   have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
@@ -1386,12 +1324,11 @@
 qed
 
 lemma simpneq_nb[simp]:
-  assumes "SORT_CONSTRAINT('a::field_char_0)"
-  shows "tmbound0 t \<Longrightarrow> bound0 (simpneq t)"
+  assumes "SORT_CONSTRAINT('a::field_char_0)" and t: "tmbound0 t"
+  shows "bound0 (simpneq t)"
 proof (simp add: simpneq_def Let_def split_def)
-  assume "tmbound0 t"
-  then have *: "tmbound0 (simptm t)"
-    by simp
+  have *: "tmbound0 (simptm t)"
+    using t by simp
   let ?c = "fst (split0 (simptm t))"
   from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
   have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
@@ -1419,35 +1356,30 @@
   where "list_disj ps \<equiv> foldr disj ps F"
 
 lemma list_conj: "Ifm vs bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm vs bs p)"
-  by (induct ps) (auto simp add: list_conj_def)
+  by (induct ps) (auto simp: list_conj_def)
 
 lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)"
-  by (induct ps) (auto simp add: list_conj_def)
+  by (induct ps) (auto simp: list_conj_def)
 
 lemma list_disj: "Ifm vs bs (list_disj ps) = (\<exists>p\<in> set ps. Ifm vs bs p)"
-  by (induct ps) (auto simp add: list_disj_def)
+  by (induct ps) (auto simp: list_disj_def)
 
 lemma conj_boundslt: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (conj p q)"
   unfolding conj_def by auto
 
 lemma conjs_nb: "bound n p \<Longrightarrow> \<forall>q\<in> set (conjs p). bound n q"
-  apply (induct p rule: conjs.induct)
-              apply (unfold conjs.simps)
-              apply (unfold set_append)
-              apply (unfold ball_Un)
-              apply (unfold bound.simps)
-              apply auto
-  done
+proof (induct p rule: conjs.induct)
+  case (1 p q)
+  then show ?case 
+    unfolding conjs.simps bound.simps by fastforce
+qed auto
 
 lemma conjs_boundslt: "boundslt n p \<Longrightarrow> \<forall>q\<in> set (conjs p). boundslt n q"
-  apply (induct p rule: conjs.induct)
-              apply (unfold conjs.simps)
-              apply (unfold set_append)
-              apply (unfold ball_Un)
-              apply (unfold boundslt.simps)
-              apply blast
-             apply simp_all
-  done
+proof (induct p rule: conjs.induct)
+  case (1 p q)
+  then show ?case 
+    unfolding conjs.simps bound.simps by fastforce
+qed auto
 
 lemma list_conj_boundslt: " \<forall>p\<in> set ps. boundslt n p \<Longrightarrow> boundslt n (list_conj ps)"
   by (induct ps) (auto simp: list_conj_def)
@@ -1503,7 +1435,7 @@
   also fix y have "\<dots> = (\<exists>x. (Ifm vs (y#bs) ?cyes) \<and> (Ifm vs (x#bs) ?cno))"
     using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast
   also have "\<dots> = (Ifm vs bs (decr0 ?cyes) \<and> Ifm vs bs (E ?cno))"
-    by (auto simp add: decr0[OF yes_nb] simp del: partition_filter_conv)
+    by (auto simp: decr0[OF yes_nb] simp del: partition_filter_conv)
   also have "\<dots> = (Ifm vs bs (conj (decr0 ?cyes) (qe ?cno)))"
     using qe[rule_format, OF no_qf] by auto
   finally have "Ifm vs bs (E p) = Ifm vs bs (CJNB qe p)"
@@ -1546,25 +1478,13 @@
   by (induct p rule: simpfm.induct) auto
 
 lemma lt_qf[simp]: "qfree (lt t)"
-  apply (cases t)
-        apply (auto simp add: lt_def)
-  apply (rename_tac poly, case_tac poly)
-         apply auto
-  done
+  by(auto simp: lt_def split: tm.split poly.split)
 
 lemma le_qf[simp]: "qfree (le t)"
-  apply (cases t)
-        apply (auto simp add: le_def)
-  apply (rename_tac poly, case_tac poly)
-         apply auto
-  done
+  by(auto simp: le_def split: tm.split poly.split)
 
 lemma eq_qf[simp]: "qfree (eq t)"
-  apply (cases t)
-        apply (auto simp add: eq_def)
-  apply (rename_tac poly, case_tac poly)
-         apply auto
-  done
+  by(auto simp: eq_def split: tm.split poly.split)
 
 lemma neq_qf[simp]: "qfree (neq t)"
   by (simp add: neq_def)
@@ -1671,19 +1591,19 @@
   shows "\<exists>z. \<forall>x < z. Ifm vs (x#bs) (minusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
   using assms
 proof (induct p rule: minusinf.induct)
-  case 1
+  case (1 p q)
+  then obtain zp zq where zp: "\<forall>x<zp. Ifm vs (x # bs) (minusinf p) = Ifm vs (x # bs) p"
+       and zq: "\<forall>x<zq. Ifm vs (x # bs) (minusinf q) = Ifm vs (x # bs) q"
+    by force
   then show ?case
-    apply auto
-    apply (rule_tac x="min z za" in exI)
-    apply auto
-    done
+    by (rule_tac x="min zp zq" in exI) auto
 next
-  case 2
+  case (2 p q)
+  then obtain zp zq where zp: "\<forall>x<zp. Ifm vs (x # bs) (minusinf p) = Ifm vs (x # bs) p"
+       and zq: "\<forall>x<zq. Ifm vs (x # bs) (minusinf q) = Ifm vs (x # bs) q"
+    by force
   then show ?case
-    apply auto
-    apply (rule_tac x="min z za" in exI)
-    apply auto
-    done
+    by (rule_tac x="min zp zq" in exI) auto
 next
   case (3 c e)
   then have nbe: "tmbound0 e"
@@ -1876,19 +1796,19 @@
   shows "\<exists>z. \<forall>x > z. Ifm vs (x#bs) (plusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
   using assms
 proof (induct p rule: plusinf.induct)
-  case 1
+  case (1 p q)
+  then obtain zp zq where zp: "\<forall>x>zp. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p"
+       and zq: "\<forall>x>zq. Ifm vs (x # bs) (plusinf q) = Ifm vs (x # bs) q"
+    by force
   then show ?case
-    apply auto
-    apply (rule_tac x="max z za" in exI)
-    apply auto
-    done
+    by (rule_tac x="max zp zq" in exI) auto
 next
-  case 2
+  case (2 p q)
+  then obtain zp zq where zp: "\<forall>x>zp. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p"
+       and zq: "\<forall>x>zq. Ifm vs (x # bs) (plusinf q) = Ifm vs (x # bs) q"
+    by force
   then show ?case
-    apply auto
-    apply (rule_tac x="max z za" in exI)
-    apply auto
-    done
+    by (rule_tac x="max zp zq" in exI) auto
 next
   case (3 c e)
   then have nbe: "tmbound0 e"
@@ -2072,10 +1992,10 @@
 qed auto
 
 lemma minusinf_nb: "islin p \<Longrightarrow> bound0 (minusinf p)"
-  by (induct p rule: minusinf.induct) (auto simp add: eq_nb lt_nb le_nb)
+  by (induct p rule: minusinf.induct) (auto simp: eq_nb lt_nb le_nb)
 
 lemma plusinf_nb: "islin p \<Longrightarrow> bound0 (plusinf p)"
-  by (induct p rule: minusinf.induct) (auto simp add: eq_nb lt_nb le_nb)
+  by (induct p rule: minusinf.induct) (auto simp: eq_nb lt_nb le_nb)
 
 lemma minusinf_ex:
   assumes lp: "islin p"
@@ -2140,16 +2060,14 @@
       Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s \<or>
       Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s"
     using lp nmi ex
-    apply (induct p rule: minusinf.induct)
-                        apply (auto simp add: eq le lt polyneg_norm)
-      apply (auto simp add: linorder_not_less order_le_less)
-    done
+    by (induct p rule: minusinf.induct)
+       (auto simp: eq le lt polyneg_norm linorder_not_less order_le_less)
   then obtain c s where csU: "(c, s) \<in> set (uset p)"
     and x: "(Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or>
       (Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s)" by blast
   then have "x \<ge> (- Itm vs (x#bs) s) / Ipoly vs c"
     using divide_le_eq[of "- Itm vs (x#bs) s" "Ipoly vs c" x]
-    by (auto simp add: mult.commute)
+    by (auto simp: mult.commute)
   then show ?thesis
     using csU by auto
 qed
@@ -2183,10 +2101,8 @@
       Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s \<or>
       Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s"
     using lp nmi ex
-    apply (induct p rule: minusinf.induct)
-                        apply (auto simp add: eq le lt polyneg_norm)
-      apply (auto simp add: linorder_not_less order_le_less)
-    done
+    by (induct p rule: minusinf.induct)
+       (auto simp: eq le lt polyneg_norm linorder_not_less order_le_less)
   then obtain c s
     where c_s: "(c, s) \<in> set (uset p)"
       and "Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s \<or>
@@ -2194,7 +2110,7 @@
     by blast
   then have "x \<le> (- Itm vs (x#bs) s) / Ipoly vs c"
     using le_divide_eq[of x "- Itm vs (x#bs) s" "Ipoly vs c"]
-    by (auto simp add: mult.commute)
+    by (auto simp: mult.commute)
   then show ?thesis
     using c_s by auto
 qed
@@ -2248,7 +2164,7 @@
     case N: 2
     from px pos_less_divide_eq[OF N, where a="x" and b="-?Nt x s"]
     have px': "x < - ?Nt x s / ?N c"
-      by (auto simp add: not_less field_simps)
+      by (auto simp: not_less field_simps)
     from ycs show ?thesis
     proof
       assume y: "y < - ?Nt x s / ?N c"
@@ -2272,7 +2188,7 @@
     case N: 3
     from px neg_divide_less_eq[OF N, where a="x" and b="-?Nt x s"]
     have px': "x > - ?Nt x s / ?N c"
-      by (auto simp add: not_less field_simps)
+      by (auto simp: not_less field_simps)
     from ycs show ?thesis
     proof
       assume y: "y > - ?Nt x s / ?N c"
@@ -2449,7 +2365,7 @@
     show ?thesis
       by (simp add: field_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric])
   qed
-qed (auto simp add: tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"]
+qed (auto simp: tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"]
   bound0_I[where vs=vs and bs="bs" and b="y" and b'="x"])
 
 lemma inf_uset:
@@ -3261,7 +3177,7 @@
     Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) p"
   using lp
   by (induct p rule: islin.induct)
-    (auto simp add: tmbound0_I
+    (auto simp: tmbound0_I
       [where b = "(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2"
         and b' = x and bs = bs and vs = vs]
       msubsteq msubstneq msubstlt [OF nc nd] msubstle [OF nc nd])
@@ -3272,7 +3188,7 @@
     and "tmbound0 s"
   shows "bound0 (msubst p ((c,t),(d,s)))"
   using assms
-  by (induct p rule: islin.induct) (auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
+  by (induct p rule: islin.induct) (auto simp: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
 
 lemma fr_eq_msubst:
   assumes lp: "islin p"
@@ -3315,7 +3231,7 @@
 lemma simpfm_lin:
   assumes "SORT_CONSTRAINT('a::field_char_0)"
   shows "qfree p \<Longrightarrow> islin (simpfm p)"
-  by (induct p rule: simpfm.induct) (auto simp add: conj_lin disj_lin)
+  by (induct p rule: simpfm.induct) (auto simp: conj_lin disj_lin)
 
 definition "ferrack p \<equiv>
   let
@@ -3395,10 +3311,7 @@
     by simp
   also have "\<dots> \<longleftrightarrow> ?rhs"
     using decr0[OF th1, of vs x bs]
-    apply (simp add: ferrack_def Let_def)
-    apply (cases "?mp = T \<or> ?pp = T")
-     apply auto
-    done
+    by (cases "?mp = T \<or> ?pp = T") (auto simp: ferrack_def Let_def)
   finally show ?thesis
     using thqf by blast
 qed
@@ -3468,11 +3381,8 @@
       case z: 4
       from z have ?F
         using ct
-        apply -
-        apply (rule bexI[where x = "(c,t)"])
-         apply simp_all
-        apply (rule bexI[where x = "(d,s)"])
-         apply simp_all
+        apply (intro bexI[where x = "(c,t)"]; simp)
+        apply (intro bexI[where x = "(d,s)"]; simp)
         done
       then show ?thesis by blast
     qed
@@ -3553,7 +3463,7 @@
     Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
   using lp pos
   by (induct p rule: islin.induct)
-    (auto simp add: msubsteq2 msubstltpos[OF pos] msubstlepos[OF pos]
+    (auto simp: msubsteq2 msubstltpos[OF pos] msubstlepos[OF pos]
       tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x]
       bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
 
@@ -3573,7 +3483,7 @@
   shows "Ifm vs (x#bs) (msubstneg p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
   using lp pos
   by (induct p rule: islin.induct)
-    (auto simp add: msubsteq2 msubstltneg[OF pos] msubstleneg[OF pos]
+    (auto simp: msubsteq2 msubstltneg[OF pos] msubstleneg[OF pos]
       tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x]
       bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
 
@@ -3596,12 +3506,12 @@
     case c: 1
     from c msubstneg_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
     show ?thesis
-      by (auto simp add: msubst2_def)
+      by (auto simp: msubst2_def)
   next
     case c: 2
     from c msubstpos_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
     show ?thesis
-      by (auto simp add: msubst2_def)
+      by (auto simp: msubst2_def)
   qed
 qed
 
@@ -3624,7 +3534,7 @@
   shows "bound0 (msubstpos p c t)"
   using lp tnb
   by (induct p c t rule: msubstpos.induct)
-    (auto simp add: msubsteq2_nb msubstltpos_nb msubstlepos_nb)
+    (auto simp: msubsteq2_nb msubstltpos_nb msubstlepos_nb)
 
 lemma msubstneg_nb:
   assumes "SORT_CONSTRAINT('a::field_char_0)"
@@ -3633,7 +3543,7 @@
   shows "bound0 (msubstneg p c t)"
   using lp tnb
   by (induct p c t rule: msubstneg.induct)
-    (auto simp add: msubsteq2_nb msubstltneg_nb msubstleneg_nb)
+    (auto simp: msubsteq2_nb msubstltneg_nb msubstleneg_nb)
 
 lemma msubst2_nb:
   assumes "SORT_CONSTRAINT('a::field_char_0)"
@@ -3652,7 +3562,7 @@
   by simp
 
 lemma islin_qf: "islin p \<Longrightarrow> qfree p"
-  by (induct p rule: islin.induct) (auto simp add: bound0_qf)
+  by (induct p rule: islin.induct) (auto simp: bound0_qf)
 
 lemma fr_eq_msubst2:
   assumes lp: "islin p"
@@ -3689,7 +3599,7 @@
         by (simp add: polyneg_norm nn)
       then have nn2: "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>n \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         using H(2) nn' nn
-        by (auto simp add: msubst2_def lt zero_less_mult_iff mult_less_0_iff)
+        by (auto simp: msubst2_def lt zero_less_mult_iff mult_less_0_iff)
       from msubst2[OF lp nn nn2(1), of x bs t]
       have "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * 2) # bs) p"
         using H(2) nn2 by (simp add: mult_minus2_right)
@@ -3731,7 +3641,7 @@
         by (simp_all add: polyneg_norm nn)
       have nn': "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         using H(3)
-        by (auto simp add: msubst2_def lt[OF stupid(1)]
+        by (auto simp: msubst2_def lt[OF stupid(1)]
             lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
       from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
       have "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and>
@@ -3862,7 +3772,7 @@
     proof
       assume H: ?lhs
       then have z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
-        by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)]
+        by (auto simp: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)]
             mult_less_0_iff zero_less_mult_iff)
       from msubst2[OF lq norm2(1) z(1), of x bs] msubst2[OF lq norm2(2) z(2), of x bs] H
       show ?rhs
@@ -3870,7 +3780,7 @@
     next
       assume H: ?rhs
       then have z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
-        by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)]
+        by (auto simp: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)]
             mult_less_0_iff zero_less_mult_iff)
       from msubst2[OF lq norm2(1) z(1), of x bs] msubst2[OF lq norm2(2) z(2), of x bs] H
       show ?lhs
@@ -3898,13 +3808,16 @@
   also have "\<dots> \<longleftrightarrow> ?I ?R"
     by (simp add: list_disj_def evaldjf_ex split_def)
   also have "\<dots> \<longleftrightarrow> ?rhs"
-    unfolding ferrack2_def
-    apply (cases "?mp = T")
-     apply (simp add: list_disj_def)
-    apply (cases "?pp = T")
-     apply (simp add: list_disj_def)
-    apply (simp_all add: Let_def decr0[OF nb])
-    done
+  proof (cases "?mp = T \<or> ?pp = T")
+    case True
+    then show ?thesis 
+      unfolding ferrack2_def
+      by (force simp add: ferrack2_def list_disj_def)
+  next
+    case False
+    then show ?thesis 
+      by (simp add: ferrack2_def Let_def decr0[OF nb])
+  qed
   finally show ?thesis using decr0_qf[OF nb]
     by (simp add: ferrack2_def Let_def)
 qed
@@ -3915,7 +3828,7 @@
   have this: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)"
     by blast
   from qelim[OF this, of "prep p" bs] show ?thesis
-    unfolding frpar2_def by (auto simp add: prep)
+    unfolding frpar2_def by (auto simp: prep)
 qed
 
 oracle frpar_oracle =
@@ -4092,57 +4005,9 @@
 \<close> "parametric QE for linear Arithmetic over fields, Version 2"
 
 lemma "\<exists>(x::'a::linordered_field). y \<noteq> -1 \<longrightarrow> (y + 1) * x < 0"
-  apply (frpar type: 'a pars: y)
-  apply (simp add: field_simps)
-  apply (rule spec[where x=y])
-  apply (frpar type: 'a pars: "z::'a")
-  apply simp
-  done
+  by (metis mult.commute neg_less_0_iff_less nonzero_divide_eq_eq sum_eq zero_less_two)
 
 lemma "\<exists>(x::'a::linordered_field). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
-  apply (frpar2 type: 'a pars: y)
-  apply (simp add: field_simps)
-  apply (rule spec[where x=y])
-  apply (frpar2 type: 'a pars: "z::'a")
-  apply simp
-  done
-
-text \<open>Collins/Jones Problem\<close>
-(*
-lemma "\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
-proof -
-  have "(\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
-by (simp add: field_simps)
-have "?rhs"
-
-  apply (frpar type: "'a::{linordered_field, number_ring}" pars: "a::'a::{linordered_field, number_ring}" "b::'a::{linordered_field, number_ring}")
-  apply (simp add: field_simps)
-oops
-*)
-(*
-lemma "ALL (x::'a::{linordered_field, number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
-apply (frpar type: "'a::{linordered_field, number_ring}" pars: "t::'a::{linordered_field, number_ring}")
-oops
-*)
-
-text \<open>Collins/Jones Problem\<close>
-
-(*
-lemma "\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
-proof -
-  have "(\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
-by (simp add: field_simps)
-have "?rhs"
-  apply (frpar2 type: "'a::{linordered_field, number_ring}" pars: "a::'a::{linordered_field, number_ring}" "b::'a::{linordered_field, number_ring}")
-  apply simp
-oops
-*)
-
-(*
-lemma "ALL (x::'a::{linordered_field, number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
-apply (frpar2 type: "'a::{linordered_field, number_ring}" pars: "t::'a::{linordered_field, number_ring}")
-apply (simp add: field_simps linorder_neq_iff[symmetric])
-apply ferrack
-oops
-*)
+  by (metis mult.right_neutral mult_minus1_right neg_0_le_iff_le nle_le not_less sum_eq)
+
 end
--- a/src/HOL/Decision_Procs/Polynomial_List.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -5,7 +5,8 @@
 section \<open>Univariate Polynomials as lists\<close>
 
 theory Polynomial_List
-imports Complex_Main
+  imports Complex_Main 
+
 begin
 
 text \<open>Application of polynomial as a function.\<close>
@@ -264,13 +265,9 @@
       by blast
     have "r = 0"
       using p0 by (simp only: Cons qr poly_mult poly_add) simp
-    with Cons qr show ?thesis
-      apply -
-      apply (rule exI[where x = q])
-      apply auto
-      apply (cases q)
-      apply auto
-      done
+    with Cons qr have "p = [- a, 1] *** q"
+      by (simp add: local.padd_commut)
+    then show ?thesis ..
   qed
   ultimately show ?thesis using Cons by blast
 qed
@@ -344,11 +341,8 @@
       by blast
     from y have "y = a \<or> poly q y = 0"
       by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps)
-    with i[of y] y(1) y(2) show ?thesis
-      apply auto
-      apply (erule_tac x = "m" in allE)
-      apply auto
-      done
+    with i[of y] y show ?thesis
+      using le_Suc_eq by auto
   qed
   then show ?case by blast
 qed
@@ -360,12 +354,7 @@
 
 lemma (in idom) poly_roots_finite_lemma1:
   "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>N i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>n::nat. n < N \<and> x = i n)"
-  apply (drule poly_roots_index_length)
-  apply safe
-  apply (rule_tac x = "Suc (length p)" in exI)
-  apply (rule_tac x = i in exI)
-  apply (simp add: less_Suc_eq_le)
-  done
+  by (metis le_imp_less_Suc poly_roots_index_length)
 
 lemma (in idom) idom_finite_lemma:
   assumes "\<forall>x. P x \<longrightarrow> (\<exists>n. n < length j \<and> x = j!n)"
@@ -379,15 +368,8 @@
 
 lemma (in idom) poly_roots_finite_lemma2:
   "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i"
-  apply (drule poly_roots_index_length)
-  apply safe
-  apply (rule_tac x = "map (\<lambda>n. i n) [0 ..< Suc (length p)]" in exI)
-  apply (auto simp add: image_iff)
-  apply (erule_tac x="x" in allE)
-  apply clarsimp
-  apply (case_tac "n = length p")
-  apply (auto simp add: order_le_less)
-  done
+  using poly_roots_index_length atMost_iff atMost_upto imageI set_map
+  by metis
 
 lemma (in ring_char_0) UNIV_ring_char_0_infinte: "\<not> finite (UNIV :: 'a set)"
 proof
@@ -408,21 +390,12 @@
   (is "?lhs \<longleftrightarrow> ?rhs")
 proof
   show ?rhs if ?lhs
-    using that
-    apply -
-    apply (erule contrapos_np)
-    apply (rule ext)
-    apply (rule ccontr)
-    apply (clarify dest!: poly_roots_finite_lemma2)
-    using finite_subset
   proof -
-    fix x i
-    assume F: "\<not> finite {x. poly p x = 0}"
-      and P: "\<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i"
-    from P have "{x. poly p x = 0} \<subseteq> set i"
-      by auto
-    with finite_subset F show False
-      by auto
+    have False if  F: "\<not> finite {x. poly p x = 0}"
+      and P: "\<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i" for  i
+      by (smt (verit, del_insts) in_set_conv_nth local.idom_finite_lemma that)
+    with that show ?thesis
+      using local.poly_roots_finite_lemma2 by blast
   qed
   show ?lhs if ?rhs
     using UNIV_ring_char_0_infinte that by auto
@@ -473,18 +446,15 @@
 qed
 
 lemma (in idom) poly_exp_eq_zero[simp]: "poly (p %^ n) = poly [] \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0"
-  apply (simp only: fun_eq_iff add: HOL.all_simps [symmetric])
-  apply (rule arg_cong [where f = All])
-  apply (rule ext)
-  apply (induct n)
-  apply (auto simp add: poly_exp poly_mult)
-  done
+  by (simp add: local.poly_exp fun_eq_iff)
 
 lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a, 1] \<noteq> poly []"
-  apply (simp add: fun_eq_iff)
-  apply (rule_tac x = "minus one a" in exI)
-  apply (simp add: add.commute [of a])
-  done
+proof -
+  have "\<exists>x. a + x \<noteq> 0"
+    by (metis add_cancel_left_right zero_neq_one)
+  then show ?thesis
+    by (simp add: fun_eq_iff)
+qed
 
 lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \<noteq> poly []"
   by auto
@@ -492,26 +462,18 @@
 
 text \<open>A more constructive notion of polynomials being trivial.\<close>
 
-lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] \<Longrightarrow> h = 0 \<and> poly t = poly []"
-  apply (simp add: fun_eq_iff)
-  apply (case_tac "h = zero")
-  apply (drule_tac [2] x = zero in spec)
-  apply auto
-  apply (cases "poly t = poly []")
-  apply simp
+lemma (in idom_char_0) poly_zero_lemma': 
+  assumes "poly (h # t) = poly []" shows "h = 0 \<and> poly t = poly []"
 proof -
-  fix x
-  assume H: "\<forall>x. x = 0 \<or> poly t x = 0"
-  assume pnz: "poly t \<noteq> poly []"
-  let ?S = "{x. poly t x = 0}"
-  from H have "\<forall>x. x \<noteq> 0 \<longrightarrow> poly t x = 0"
-    by blast
-  then have th: "?S \<supseteq> UNIV - {0}"
-    by auto
-  from poly_roots_finite pnz have th': "finite ?S"
-    by blast
-  from finite_subset[OF th th'] UNIV_ring_char_0_infinte show "poly t x = 0"
-    by simp
+  have "poly t x = 0" if H: "\<forall>x. x = 0 \<or> poly t x = 0" and pnz: "poly t \<noteq> poly []" for x
+  proof -
+    from H have "{x. poly t x = 0} \<supseteq> UNIV - {0}"
+      by auto
+    then show ?thesis
+      using finite_subset local.poly_roots_finite pnz by fastforce
+  qed
+  with assms show ?thesis
+    by (simp add: fun_eq_iff) (metis add_cancel_right_left mult_eq_0_iff)
 qed
 
 lemma (in idom_char_0) poly_zero: "poly p = poly [] \<longleftrightarrow> (\<forall>c \<in> set p. c = 0)"
@@ -520,12 +482,8 @@
   then show ?case by simp
 next
   case Cons
-  show ?case
-    apply (rule iffI)
-    apply (drule poly_zero_lemma')
-    using Cons
-    apply auto
-    done
+  then show ?case
+    by (smt (verit) list.set_intros pmult_by_x poly_entire poly_zero_lemma' set_ConsD)
 qed
 
 lemma (in idom_char_0) poly_0: "\<forall>c \<in> set p. c = 0 \<Longrightarrow> poly p x = 0"
@@ -535,41 +493,33 @@
 text \<open>Basics of divisibility.\<close>
 
 lemma (in idom) poly_primes: "[a, 1] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
-  apply (auto simp add: divides_def fun_eq_iff poly_mult poly_add poly_cmult distrib_right [symmetric])
-  apply (drule_tac x = "uminus a" in spec)
-  apply (simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
-  apply (cases "p = []")
-  apply (rule exI[where x="[]"])
-  apply simp
-  apply (cases "q = []")
-  apply (erule allE[where x="[]"])
-  apply simp
-
-  apply clarsimp
-  apply (cases "\<exists>q. p = a %* q +++ (0 # q)")
-  apply (clarsimp simp add: poly_add poly_cmult)
-  apply (rule_tac x = qa in exI)
-  apply (simp add: distrib_right [symmetric])
-  apply clarsimp
-
-  apply (auto simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
-  apply (rule_tac x = "pmult qa q" in exI)
-  apply (rule_tac [2] x = "pmult p qa" in exI)
-  apply (auto simp add: poly_add poly_mult poly_cmult ac_simps)
-  done
+proof -
+  have "\<exists>q. \<forall>x. poly p x = (a + x) * poly q x"
+    if "poly p (uminus a) * poly q (uminus a) = (a + (uminus a)) * poly qa (uminus a)"
+      and "\<forall>qa. \<exists>x. poly q x \<noteq> (a + x) * poly qa x"
+    for qa 
+    using that   
+    apply (simp add: poly_linear_divides poly_add)
+    by (metis add_cancel_left_right combine_common_factor mult_eq_0_iff poly.poly_Cons poly.poly_Nil poly_add poly_cmult)
+  moreover have "\<exists>qb. \<forall>x. (a + x) * poly qa x * poly q x = (a + x) * poly qb x" for qa
+    by (metis local.poly_mult mult_assoc)
+  moreover have "\<exists>q. \<forall>x. poly p x * ((a + x) * poly qa x) = (a + x) * poly q x" for qa 
+    by (metis mult.left_commute local.poly_mult)
+  ultimately show ?thesis
+    by (auto simp: divides_def divisors_zero fun_eq_iff poly_mult poly_add poly_cmult simp flip: distrib_right)
+qed
 
 lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p"
-  apply (simp add: divides_def)
-  apply (rule_tac x = "[one]" in exI)
-  apply (auto simp add: poly_mult fun_eq_iff)
-  done
+proof -
+  have "poly p = poly (p *** [1])"
+    by (auto simp add: poly_mult fun_eq_iff)
+  then show ?thesis
+    using local.dividesI by blast
+qed
 
 lemma (in comm_semiring_1) poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r"
-  apply (simp add: divides_def)
-  apply safe
-  apply (rule_tac x = "pmult qa qaa" in exI)
-  apply (auto simp add: poly_mult fun_eq_iff mult.assoc)
-  done
+  unfolding divides_def
+  by (metis ext local.poly_mult local.poly_mult_assoc)
 
 lemma (in comm_semiring_1) poly_divides_exp: "m \<le> n \<Longrightarrow> (p %^ m) divides (p %^ n)"
   by (auto simp: le_iff_add divides_def poly_exp_add fun_eq_iff)
@@ -577,34 +527,35 @@
 lemma (in comm_semiring_1) poly_exp_divides: "(p %^ n) divides q \<Longrightarrow> m \<le> n \<Longrightarrow> (p %^ m) divides q"
   by (blast intro: poly_divides_exp poly_divides_trans)
 
-lemma (in comm_semiring_0) poly_divides_add: "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)"
-  apply (auto simp add: divides_def)
-  apply (rule_tac x = "padd qa qaa" in exI)
-  apply (auto simp add: poly_add fun_eq_iff poly_mult distrib_left)
-  done
+lemma (in comm_semiring_0) poly_divides_add:
+  assumes "p divides q" and "p divides r" shows "p divides (q +++ r)"
+proof -
+  have "\<And>qa qb. \<lbrakk>poly q = poly (p *** qa); poly r = poly (p *** qb)\<rbrakk>
+       \<Longrightarrow> poly (q +++ r) = poly (p *** (qa +++ qb))"
+    by (auto simp add: poly_add fun_eq_iff poly_mult distrib_left)
+  with assms show ?thesis
+    by (auto simp add: divides_def)
+qed
 
-lemma (in comm_ring_1) poly_divides_diff: "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r"
-  apply (auto simp add: divides_def)
-  apply (rule_tac x = "padd qaa (poly_minus qa)" in exI)
-  apply (auto simp add: poly_add fun_eq_iff poly_mult poly_minus algebra_simps)
-  done
+lemma (in comm_ring_1) poly_divides_diff:
+  assumes "p divides q" and "p divides (q +++ r)"
+  shows "p divides r"
+proof -
+  have "\<And>qa qb. \<lbrakk>poly q = poly (p *** qa); poly (q +++ r) = poly (p *** qb)\<rbrakk>
+         \<Longrightarrow> poly r = poly (p *** (qb +++ -- qa))"
+  by (auto simp add: poly_add fun_eq_iff poly_mult poly_minus algebra_simps)
+  with assms show ?thesis
+    by (auto simp add: divides_def)
+qed
 
 lemma (in comm_ring_1) poly_divides_diff2: "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q"
-  apply (erule poly_divides_diff)
-  apply (auto simp add: poly_add fun_eq_iff poly_mult divides_def ac_simps)
-  done
+  by (metis local.padd_commut local.poly_divides_diff)
 
 lemma (in semiring_0) poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p"
-  apply (simp add: divides_def)
-  apply (rule exI[where x = "[]"])
-  apply (auto simp add: fun_eq_iff poly_mult)
-  done
+  by (metis ext dividesI poly.poly_Nil poly_mult_Nil2)
 
 lemma (in semiring_0) poly_divides_zero2 [simp]: "q divides []"
-  apply (simp add: divides_def)
-  apply (rule_tac x = "[]" in exI)
-  apply (auto simp add: fun_eq_iff)
-  done
+  using local.poly_divides_zero by force
 
 
 text \<open>At last, we can consider the order of a root.\<close>
@@ -629,11 +580,7 @@
     from True[unfolded poly_linear_divides] pN obtain q where q: "p = [-a, 1] *** q"
       by blast
     from q h True have qh: "length q = n" "poly q \<noteq> poly []"
-      apply simp_all
-      apply (simp only: fun_eq_iff)
-      apply (rule ccontr)
-      apply (simp add: fun_eq_iff poly_add poly_cmult)
-      done
+      using h(2) local.poly_entire q by fastforce+
     from Suc.hyps[OF qh] obtain m r where mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0"
       by blast
     from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0"
@@ -641,12 +588,8 @@
     then show ?thesis by blast
   next
     case False
-    then show ?thesis
-      using Suc.prems
-      apply simp
-      apply (rule exI[where x="0::nat"])
-      apply simp
-      done
+    with Suc.prems show ?thesis
+      by (smt (verit, best) local.mulexp.mulexp_zero)
   qed
 qed
 
@@ -719,12 +662,7 @@
 
 lemma (in idom_char_0) poly_order:
   "poly p \<noteq> poly [] \<Longrightarrow> \<exists>!n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)"
-  apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
-  apply (cut_tac x = y and y = n in less_linear)
-  apply (drule_tac m = n in poly_exp_divides)
-  apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
-    simp del: pmult_Cons pexp_Suc)
-  done
+  by (meson Suc_le_eq linorder_neqE_nat local.poly_exp_divides poly_order_exists)
 
 
 text \<open>Order\<close>
@@ -736,10 +674,7 @@
   "([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p) \<longleftrightarrow>
     n = order a p \<and> poly p \<noteq> poly []"
   unfolding order_def
-  apply (rule iffI)
-  apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
-  apply (blast intro!: poly_order [THEN [2] some1_equalityD])
-  done
+  by (metis (no_types, lifting) local.poly_divides_zero local.poly_order someI)
 
 lemma (in idom_char_0) order2:
   "poly p \<noteq> poly [] \<Longrightarrow>
@@ -767,97 +702,89 @@
   by (induct n arbitrary: a p) (auto simp add: divides_def poly_mult simp del: pmult_Cons)
 
 lemma (in idom_char_0) order_root: "poly p a = 0 \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0"
-  apply (cases "poly p = poly []")
-  apply auto
-  apply (simp add: poly_linear_divides del: pmult_Cons)
-  apply safe
-  apply (drule_tac [!] a = a in order2)
-  apply (rule ccontr)
-  apply (simp add: divides_def poly_mult fun_eq_iff del: pmult_Cons)
-  apply blast
-  using neq0_conv apply (blast intro: lemma_order_root)
-  done
+proof (cases "poly p = poly []")
+  case False
+  then show ?thesis 
+    by (metis (mono_tags, lifting) dividesI lemma_order_root order2 pexp_one poly_linear_divides neq0_conv)
+qed auto
 
 lemma (in idom_char_0) order_divides:
   "([-a, 1] %^ n) divides p \<longleftrightarrow> poly p = poly [] \<or> n \<le> order a p"
-  apply (cases "poly p = poly []")
-  apply auto
-  apply (simp add: divides_def fun_eq_iff poly_mult)
-  apply (rule_tac x = "[]" in exI)
-  apply (auto dest!: order2 [where a=a] intro: poly_exp_divides simp del: pexp_Suc)
-  done
+proof (cases "poly p = poly []")
+  case True
+  then show ?thesis 
+    using local.poly_divides_zero by force
+next
+  case False
+  then show ?thesis 
+    by (meson local.order2 local.poly_exp_divides not_less_eq_eq)
+qed
 
 lemma (in idom_char_0) order_decomp:
-  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>q. poly p = poly (([-a, 1] %^ order a p) *** q) \<and> \<not> [-a, 1] divides q"
-  unfolding divides_def
-  apply (drule order2 [where a = a])
-  apply (simp add: divides_def del: pexp_Suc pmult_Cons)
-  apply safe
-  apply (rule_tac x = q in exI)
-  apply safe
-  apply (drule_tac x = qa in spec)
-  apply (auto simp add: poly_mult fun_eq_iff poly_exp ac_simps simp del: pmult_Cons)
-  done
+  assumes "poly p \<noteq> poly []"
+  shows "\<exists>q. poly p = poly (([-a, 1] %^ order a p) *** q) \<and> \<not> [-a, 1] divides q"
+proof -
+  obtain q where q: "poly p = poly ([- a, 1] %^ order a p *** q)"
+    using assms local.order2 divides_def by blast
+  have False if "poly q = poly ([- a, 1] *** qa)" for qa
+  proof -
+    have "poly p \<noteq> poly ([- a, 1] %^ Suc (order a p) *** qa)"
+      using assms local.divides_def local.order2 by blast
+    with q that show False
+      by (auto simp add: poly_mult ac_simps simp del: pmult_Cons)
+  qed 
+  with q show ?thesis
+    unfolding divides_def by blast
+qed
 
 text \<open>Important composition properties of orders.\<close>
 lemma order_mult:
   fixes a :: "'a::idom_char_0"
-  shows "poly (p *** q) \<noteq> poly [] \<Longrightarrow> order a (p *** q) = order a p + order a q"
-  apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order)
-  apply (auto simp add: poly_entire simp del: pmult_Cons)
-  apply (drule_tac a = a in order2)+
-  apply safe
-  apply (simp add: divides_def fun_eq_iff poly_exp_add poly_mult del: pmult_Cons, safe)
-  apply (rule_tac x = "qa *** qaa" in exI)
-  apply (simp add: poly_mult ac_simps del: pmult_Cons)
-  apply (drule_tac a = a in order_decomp)+
-  apply safe
-  apply (subgoal_tac "[-a, 1] divides (qa *** qaa) ")
-  apply (simp add: poly_primes del: pmult_Cons)
-  apply (auto simp add: divides_def simp del: pmult_Cons)
-  apply (rule_tac x = qb in exI)
-  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) =
-    poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
-  apply (drule poly_mult_left_cancel [THEN iffD1])
-  apply force
-  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) =
-    poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
-  apply (drule poly_mult_left_cancel [THEN iffD1])
-  apply force
-  apply (simp add: fun_eq_iff poly_exp_add poly_mult ac_simps del: pmult_Cons)
-  done
-
-lemma (in idom_char_0) order_mult:
   assumes "poly (p *** q) \<noteq> poly []"
   shows "order a (p *** q) = order a p + order a q"
-  using assms
-  apply (cut_tac a = a and p = "pmult p q" and n = "order a p + order a q" in order)
-  apply (auto simp add: poly_entire simp del: pmult_Cons)
-  apply (drule_tac a = a in order2)+
-  apply safe
-  apply (simp add: divides_def fun_eq_iff poly_exp_add poly_mult del: pmult_Cons)
-  apply safe
-  apply (rule_tac x = "pmult qa qaa" in exI)
-  apply (simp add: poly_mult ac_simps del: pmult_Cons)
-  apply (drule_tac a = a in order_decomp)+
-  apply safe
-  apply (subgoal_tac "[uminus a, one] divides pmult qa qaa")
-  apply (simp add: poly_primes del: pmult_Cons)
-  apply (auto simp add: divides_def simp del: pmult_Cons)
-  apply (rule_tac x = qb in exI)
-  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa)) =
-    poly (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb))")
-  apply (drule poly_mult_left_cancel [THEN iffD1], force)
-  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a q))
-      (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) =
-    poly (pmult (pexp [uminus a, one] (order a q))
-      (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb)))")
-  apply (drule poly_mult_left_cancel [THEN iffD1], force)
-  apply (simp add: fun_eq_iff poly_exp_add poly_mult ac_simps del: pmult_Cons)
-  done
+proof -
+  have p: "poly p \<noteq> poly []" and q: "poly q \<noteq> poly []"
+    using assms poly_entire by auto
+  obtain p' where p': 
+          "\<And>x. poly p x = poly ([- a, 1] %^ order a p) x * poly p' x"
+          "\<not> [- a, 1] divides p'"
+    by (metis order_decomp p poly_mult)
+  obtain q' where q': 
+          "\<And>x. poly q x = poly ([- a, 1] %^ order a q) x * poly q' x"
+          "\<not> [- a, 1] divides q'"
+    by (metis order_decomp q poly_mult)
+  have "[- a, 1] %^ (order a p + order a q) divides (p *** q)"
+  proof -
+    have *: "poly p x * poly q x =
+          poly ([- a, 1] %^ order a p) x * poly ([- a, 1] %^ order a q) x * poly (p' *** q') x" for x
+      using p' q' by (simp add: poly_mult)
+    then show ?thesis
+      unfolding divides_def  poly_exp_add poly_mult using * by blast
+  qed
+  moreover have False
+    if pq: "order a (p *** q) \<noteq> order a p + order a q"
+      and dv: "[- a, 1] *** [- a, 1] %^ (order a p + order a q) divides (p *** q)"
+  proof -
+    obtain pq' :: "'a list"
+      where pq': "poly (p *** q) = poly ([- a, 1] *** [- a, 1] %^ (order a p + order a q) *** pq')"
+      using dv unfolding divides_def by auto
+    have "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (p' *** q'))) =
+          poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** pq')))"
+      using p' q' pq pq'
+      by (simp add: fun_eq_iff poly_exp_add poly_mult ac_simps del: pmult_Cons)
+    then have "poly ([-a, 1] %^ (order a p) *** (p' *** q')) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** pq'))"
+      by (simp add: poly_mult_left_cancel)
+    then have "[-a, 1] divides (p' *** q')"
+      unfolding divides_def by (meson poly_exp_prime_eq_zero poly_mult_left_cancel)
+    with p' q' show ?thesis
+      by (simp add: poly_primes)
+  qed
+  ultimately show ?thesis
+    by (metis order pexp_Suc)
+qed
 
 lemma (in idom_char_0) order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order a p \<noteq> 0"
-  by (rule order_root [THEN ssubst]) auto
+  using order_root by presburger
 
 lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p"
   by auto
@@ -866,21 +793,18 @@
   by (simp add: fun_eq_iff)
 
 lemma (in idom_char_0) rsquarefree_decomp:
-  "rsquarefree p \<Longrightarrow> poly p a = 0 \<Longrightarrow> \<exists>q. poly p = poly ([-a, 1] *** q) \<and> poly q a \<noteq> 0"
-  apply (simp add: rsquarefree_def)
-  apply safe
-  apply (frule_tac a = a in order_decomp)
-  apply (drule_tac x = a in spec)
-  apply (drule_tac a = a in order_root2 [symmetric])
-  apply (auto simp del: pmult_Cons)
-  apply (rule_tac x = q in exI, safe)
-  apply (simp add: poly_mult fun_eq_iff)
-  apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
-  apply (simp add: divides_def del: pmult_Cons, safe)
-  apply (drule_tac x = "[]" in spec)
-  apply (auto simp add: fun_eq_iff)
-  done
-
+  assumes "rsquarefree p" and "poly p a = 0"
+  shows "\<exists>q. poly p = poly ([-a, 1] *** q) \<and> poly q a \<noteq> 0"
+proof -
+  have "order a p = Suc 0"
+    using assms local.order_root2 rsquarefree_def by force
+  moreover
+  obtain q where "poly p = poly ([- a, 1] %^ order a p *** q)" 
+                 "\<not> [- a, 1] divides q"
+    using assms(1) order_decomp rsquarefree_def by blast
+  ultimately show ?thesis
+    using dividesI poly_linear_divides by auto
+qed
 
 text \<open>Normalization of a polynomial.\<close>
 
@@ -952,28 +876,12 @@
   then show ?case
   proof (induct p)
     case Nil
-    then have "poly [] = poly (c # cs)"
-      by blast
-    then have "poly (c#cs) = poly []"
-      by simp
     then show ?case
-      by (simp only: poly_zero lemma_degree_zero) simp
+      by (metis local.poly_zero_lemma')
   next
     case (Cons d ds)
-    then have eq: "poly (d # ds) = poly (c # cs)"
-      by blast
-    then have eq': "\<And>x. poly (d # ds) x = poly (c # cs) x"
-      by simp
-    then have "poly (d # ds) 0 = poly (c # cs) 0"
-      by blast
-    then have dc: "d = c"
-      by auto
-    with eq have "poly ds = poly cs"
-      unfolding  poly_Cons_eq by simp
-    with Cons.prems have "pnormalize ds = pnormalize cs"
-      by blast
-    with dc show ?case
-      by simp
+    then show ?case
+      by (metis pnormalize.pnormalize_Cons local.poly_Cons_eq)
   qed
 qed
 
@@ -987,15 +895,16 @@
 
 lemma (in semiring_0) last_linear_mul_lemma:
   "last ((a %* p) +++ (x # (b %* p))) = (if p = [] then x else b * last p)"
-  apply (induct p arbitrary: a x b)
-  apply auto
-  subgoal for a p c x b
-    apply (subgoal_tac "padd (cmult c p) (times b a # cmult b p) \<noteq> []")
-    apply simp
-    apply (induct p)
-    apply auto
-    done
-  done
+proof (induct p arbitrary: a x b)
+  case Nil
+  then show ?case by auto
+next
+  case (Cons a p c x b)
+  then have "padd (cmult c p) (times b a # cmult b p) \<noteq> []"
+    by (metis local.padd.padd_Nil local.padd_Cons_Cons neq_Nil_conv)
+  then show ?case
+    by (simp add: local.Cons)
+qed
 
 lemma (in semiring_1) last_linear_mul:
   assumes p: "p \<noteq> []"
@@ -1051,18 +960,11 @@
     case True
     then show ?thesis
       using degree_unique[OF True] by (simp add: degree_def)
-  next
-    case False
-    then show ?thesis
-      by (auto simp add: poly_Nil_ext)
-  qed
+  qed (auto simp add: poly_Nil_ext)
 next
   case (Suc n a p)
   have eq: "poly ([a, 1] %^(Suc n) *** p) = poly ([a, 1] %^ n *** ([a, 1] *** p))"
-    apply (rule ext)
-    apply (simp add: poly_mult poly_add poly_cmult)
-    apply (simp add: ac_simps distrib_left)
-    done
+    by (force simp add: poly_mult poly_add poly_cmult ac_simps distrib_left)
   note deq = degree_unique[OF eq]
   show ?case
   proof (cases "poly p = poly []")
@@ -1080,9 +982,7 @@
     from ap have ap': "poly ([a, 1] *** p) = poly [] \<longleftrightarrow> False"
       by blast
     have th0: "degree ([a, 1]%^n *** ([a, 1] *** p)) = degree ([a, 1] *** p) + n"
-      apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap')
-      apply simp
-      done
+      unfolding Suc.hyps[of a "pmult [a,one] p"] ap' by simp
     from degree_unique[OF eq] ap False th0 linear_mul_degree[OF False, of a]
     show ?thesis
       by (auto simp del: poly.simps)
@@ -1118,12 +1018,12 @@
   then show ?case by simp
 next
   case (Cons a p)
-  then show ?case
-    apply auto
-    apply (rule_tac y = "\<bar>a\<bar> + \<bar>x * poly p x\<bar>" in order_trans)
-    apply (rule abs_triangle_ineq)
-    apply (auto intro!: mult_mono simp add: abs_mult)
-    done
+  have "\<bar>a + x * poly p x\<bar> \<le> \<bar>a\<bar> + \<bar>x * poly p x\<bar>"
+    using abs_triangle_ineq by blast
+  also have "\<dots> \<le> \<bar>a\<bar> + k * poly (map abs p) k"
+    by (simp add: Cons.hyps Cons.prems abs_mult mult_mono')
+  finally show ?case
+    using Cons by auto
 qed
 
 lemma (in semiring_0) poly_Sing: "poly [c] x = c"
--- a/src/HOL/Decision_Procs/Rat_Pair.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Decision_Procs/Rat_Pair.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -140,10 +140,7 @@
 qed
 
 lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)"
-  apply (simp add: Ninv_def isnormNum_def split_def)
-  apply (cases "fst x = 0")
-  apply (auto simp add: gcd.commute)
-  done
+  by (simp add: Ninv_def isnormNum_def split_def gcd.commute split: if_split_asm)
 
 lemma isnormNum_int[simp]: "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \<noteq> 0 \<Longrightarrow> isnormNum (i)\<^sub>N"
   by (simp_all add: isnormNum_def)
@@ -201,12 +198,7 @@
     then have "coprime a b" "coprime b a" "coprime a' b'" "coprime b' a'"
       by (simp_all add: coprime_iff_gcd_eq_1)
     from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'"
-      apply -
-      apply algebra
-      apply algebra
-      apply simp
-      apply algebra
-      done
+      by (algebra|simp)+
     then have eq1: "b = b'"
       using pos \<open>coprime b a\<close> \<open>coprime b' a'\<close>
       by (auto simp add: coprime_dvd_mult_left_iff intro: associated_eqI)
@@ -282,15 +274,7 @@
   proof cases
     case 1
     then show ?thesis
-      apply (cases "a = 0")
-      apply (simp_all add: x y Nadd_def)
-      apply (cases "b = 0")
-      apply (simp_all add: INum_def)
-      apply (cases "a'= 0")
-      apply simp_all
-      apply (cases "b'= 0")
-      apply simp_all
-      done
+      by (auto simp: x y INum_def Nadd_def normNum_def Let_def of_int_div)
   next
     case neq: 2
     show ?thesis
@@ -480,13 +464,7 @@
     and "(a, 0) +\<^sub>N y = normNum y"
     and "x +\<^sub>N (0, b) = normNum x"
     and "x +\<^sub>N (a, 0) = normNum x"
-  apply (simp add: Nadd_def split_def)
-  apply (simp add: Nadd_def split_def)
-  apply (subst Nadd_commute)
-  apply (simp add: Nadd_def split_def)
-  apply (subst Nadd_commute)
-  apply (simp add: Nadd_def split_def)
-  done
+  by (simp_all add: Nadd_def normNum_def split_def)
 
 lemma normNum_nilpotent_aux[simp]:
   assumes "SORT_CONSTRAINT('a::field_char_0)"
@@ -599,18 +577,13 @@
   obtain a' b' where y: "y = (a', b')" by (cases y)
   have n0: "isnormNum 0\<^sub>N" by simp
   show ?thesis using nx ny
-    apply (simp only: isnormNum_unique[where ?'a = 'a, OF  Nmul_normN[OF nx ny] n0, symmetric]
-      Nmul[where ?'a = 'a])
-    apply (simp add: x y INum_def split_def isnormNum_def split: if_split_asm)
-    done
+    by (metis (no_types, opaque_lifting) INum_int(2) Nmul Nmul0(1) Nmul0(2) isnormNum0 mult_eq_0_iff)
 qed
 
 lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c"
   by (simp add: Nneg_def split_def)
 
 lemma Nmul1[simp]: "isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c" "isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c"
-  apply (simp_all add: Nmul_def Let_def split_def isnormNum_def)
-  apply (cases "fst c = 0", simp_all, cases c, simp_all)+
-  done
+  by (simp add: Nmul_def Let_def split_def isnormNum_def, metis div_0 div_by_1 surjective_pairing)+
 
 end
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -893,12 +893,7 @@
   by (induct p) auto
 
 lemma decrpoly_normh: "isnpolyh p n0 \<Longrightarrow> polybound0 p \<Longrightarrow> isnpolyh (decrpoly p) (n0 - 1)"
-  apply (induct p arbitrary: n0)
-         apply auto
-  apply atomize
-  apply (rename_tac nat a b, erule_tac x = "Suc nat" in allE)
-  apply auto
-  done
+  by (induct p arbitrary: n0) force+
 
 lemma head_polybound0: "isnpolyh p n0 \<Longrightarrow> polybound0 (head p)"
   by (induct p  arbitrary: n0 rule: head.induct) (auto intro: isnpolyh_polybound0)
@@ -1014,22 +1009,18 @@
   by (rule exI[where x="replicate (n - length xs) z" for z]) simp
 
 lemma isnpolyh_Suc_const: "isnpolyh p (Suc n) \<Longrightarrow> isconstant p"
-  apply (cases p)
-         apply auto
-  apply (rename_tac nat a, case_tac "nat")
-   apply simp_all
-  done
+  by (simp add: isconstant_polybound0 isnpolyh_polybound0)
 
 lemma wf_bs_polyadd: "wf_bs bs p \<and> wf_bs bs q \<longrightarrow> wf_bs bs (p +\<^sub>p q)"
   by (induct p q rule: polyadd.induct) (auto simp add: Let_def wf_bs_def)
 
 lemma wf_bs_polyul: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p *\<^sub>p q)"
-  apply (induct p q arbitrary: bs rule: polymul.induct)
-                      apply (simp_all add: wf_bs_polyadd wf_bs_def)
-  apply clarsimp
-  apply (rule wf_bs_polyadd[unfolded wf_bs_def, rule_format])
-  apply auto
-  done
+proof (induct p q rule: polymul.induct)
+  case (4 c n p c' n' p')
+  then show ?case
+    apply (simp add: wf_bs_def)
+    by (metis Suc_eq_plus1 max.bounded_iff max_0L maxindex.simps(2) maxindex.simps(8) wf_bs_def wf_bs_polyadd)
+qed (simp_all add: wf_bs_def)
 
 lemma wf_bs_polyneg: "wf_bs bs p \<Longrightarrow> wf_bs bs (~\<^sub>p p)"
   by (induct p rule: polyneg.induct) (auto simp: wf_bs_def)
@@ -1100,11 +1091,7 @@
   shows "polypoly bs p = [Ipoly bs p]"
   using assms
   unfolding polypoly_def
-  apply (cases p)
-         apply auto
-  apply (rename_tac nat a, case_tac nat)
-   apply auto
-  done
+  by (cases p; use gr0_conv_Suc in force)
 
 lemma head_isnpolyh: "isnpolyh p n0 \<Longrightarrow> isnpolyh (head p) n0"
   by (induct p rule: head.induct) auto
@@ -1291,12 +1278,11 @@
 lemma degree_polyneg:
   assumes "isnpolyh p n"
   shows "degree (polyneg p) = degree p"
-  apply (induct p rule: polyneg.induct)
-  using assms
-         apply simp_all
-  apply (case_tac na)
-   apply auto
-  done
+proof (induct p rule: polyneg.induct)
+  case (2 c n p)
+  then show ?case
+  by simp (smt (verit) degree.elims degree.simps(1) poly.inject(8))
+qed auto
 
 lemma degree_polyadd:
   assumes np: "isnpolyh p n0"
@@ -1458,24 +1444,24 @@
 
 lemma polyadd_eq_const_degree:
   "isnpolyh p n0 \<Longrightarrow> isnpolyh q n1 \<Longrightarrow> polyadd p q = C c \<Longrightarrow> degree p = degree q"
-  using polyadd_eq_const_degreen degree_eq_degreen0 by simp
+  by (metis degree_eq_degreen0 polyadd_eq_const_degreen)
+
+lemma polyadd_head_aux:
+  assumes "isnpolyh p n0" "isnpolyh q n1"
+  shows "head (p +\<^sub>p q)
+           = (if degree p < degree q then head q 
+              else if degree p > degree q then head p else head (p +\<^sub>p q))"
+  using assms
+proof (induct p q rule: polyadd.induct)
+  case (4 c n p c' n' p')
+  then show ?case
+    by (auto simp add: degree_eq_degreen0 Let_def; metis head_nz)
+qed (auto simp: degree_eq_degreen0)
 
 lemma polyadd_head:
-  assumes np: "isnpolyh p n0"
-    and nq: "isnpolyh q n1"
-    and deg: "degree p \<noteq> degree q"
-  shows "head (p +\<^sub>p q) = (if degree p < degree q then head q else head p)"
-  using np nq deg
-  apply (induct p q arbitrary: n0 n1 rule: polyadd.induct)
-                      apply simp_all
-    apply (case_tac n', simp, simp)
-   apply (case_tac n, simp, simp)
-  apply (case_tac n, case_tac n', simp add: Let_def)
-    apply (auto simp add: polyadd_eq_const_degree)[2]
-    apply (metis head_nz)
-   apply (metis head_nz)
-  apply (metis degree.simps(9) gr0_conv_Suc head.simps(1) less_Suc0 not_less_eq)
-  done
+  assumes "isnpolyh p n0" and "isnpolyh q n1" and "degree p \<noteq> degree q"
+  shows "head (p +\<^sub>p q) = (if degree p < degree q then head q  else head p)"
+  by (meson assms nat_neq_iff polyadd_head_aux)
 
 lemma polymul_head_polyeq:
   assumes "SORT_CONSTRAINT('a::field_char_0)"
@@ -1503,22 +1489,12 @@
     case nn': 1
     then show ?thesis
       using norm "4.hyps"(2)[OF norm(1,6)] "4.hyps"(1)[OF norm(2,6)]
-      apply simp
-      apply (cases n)
-       apply simp
-      apply (cases n')
-       apply simp_all
-      done
+      by (simp add: head_eq_headn0)
   next
     case nn': 2
     then show ?thesis
       using norm "4.hyps"(6) [OF norm(5,3)] "4.hyps"(5)[OF norm(5,4)]
-      apply simp
-      apply (cases n')
-       apply simp
-      apply (cases n)
-       apply auto
-      done
+      by (simp add: head_eq_headn0)
   next
     case nn': 3
     from nn' polymul_normh[OF norm(5,4)]
@@ -1626,9 +1602,7 @@
     case True
     with np show ?thesis
       apply (clarsimp simp: polydivide_aux.simps)
-      apply (rule exI[where x="0\<^sub>p"])
-      apply simp
-      done
+      by (metis polyadd_0(1) polymul_0(1) zero_normh)
   next
     case sz: False
     show ?thesis
@@ -1636,10 +1610,8 @@
       case True
       then show ?thesis
         using ns ndp np polydivide_aux.simps
-        apply auto
-        apply (rule exI[where x="0\<^sub>p"])
-        apply simp
-        done
+        apply clarsimp
+        by (metis polyadd_0(2) polymul_0(1) polymul_1(2) zero_normh)
     next
       case dn': False
       then have dn: "degree s \<ge> n"
@@ -1896,11 +1868,7 @@
     (\<exists>nr. isnpolyh r nr) \<and> (\<exists>q n1. isnpolyh q n1 \<and> head p ^\<^sub>p k - 0 *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)"
     by blast
   with kr show ?thesis
-    apply -
-    apply (rule exI[where x="k"])
-    apply (rule exI[where x="r"])
-    apply simp
-    done
+    by auto
 qed
 
 
@@ -1936,11 +1904,7 @@
 
 lemma isnonconstant_coefficients_length: "isnonconstant p \<Longrightarrow> length (coefficients p) > 1"
   unfolding isnonconstant_def
-  apply (cases p)
-  apply simp_all
-  apply (rename_tac nat a, case_tac nat)
-  apply auto
-  done
+  using isconstant.elims(3) by fastforce
 
 lemma isnonconstant_nonconstant:
   assumes "isnonconstant p"
@@ -1967,11 +1931,7 @@
 qed
 
 lemma pnormal_length: "p \<noteq> [] \<Longrightarrow> pnormal p \<longleftrightarrow> length (pnormalize p) = length p"
-  apply (induct p)
-   apply (simp_all add: pnormal_def)
-  apply (case_tac "p = []")
-   apply simp_all
-  done
+  by (induct p) (auto simp: pnormal_id)
 
 lemma degree_degree:
   assumes "isnonconstant p"
--- a/src/HOL/Groups.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Groups.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -207,26 +207,27 @@
 subsection \<open>Semigroups and Monoids\<close>
 
 class semigroup_add = plus +
-  assumes add_assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps]:
-    "(a + b) + c = a + (b + c)"
+  assumes add_assoc: "(a + b) + c = a + (b + c)"
 begin
 
 sublocale add: semigroup plus
   by standard (fact add_assoc)
 
+declare add.assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps]
+
 end
 
 hide_fact add_assoc
 
 class ab_semigroup_add = semigroup_add +
-  assumes add_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]:
-    "a + b = b + a"
+  assumes add_commute: "a + b = b + a"
 begin
 
 sublocale add: abel_semigroup plus
   by standard (fact add_commute)
 
-declare add.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]
+declare add.commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]
+  add.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]
 
 lemmas add_ac = add.assoc add.commute add.left_commute
 
@@ -237,26 +238,27 @@
 lemmas add_ac = add.assoc add.commute add.left_commute
 
 class semigroup_mult = times +
-  assumes mult_assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps]:
-    "(a * b) * c = a * (b * c)"
+  assumes mult_assoc: "(a * b) * c = a * (b * c)"
 begin
 
 sublocale mult: semigroup times
   by standard (fact mult_assoc)
 
+declare mult.assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps]
+
 end
 
 hide_fact mult_assoc
 
 class ab_semigroup_mult = semigroup_mult +
-  assumes mult_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]:
-    "a * b = b * a"
+  assumes mult_commute: "a * b = b * a"
 begin
 
 sublocale mult: abel_semigroup times
   by standard (fact mult_commute)
 
-declare mult.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]
+declare mult.commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]
+  mult.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]
 
 lemmas mult_ac = mult.assoc mult.commute mult.left_commute
 
--- a/src/HOL/Homology/Homology_Groups.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Homology/Homology_Groups.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -124,8 +124,7 @@
 lemma subgroup_singular_relboundary_relcycle:
   "subgroup (singular_relboundary_set p X S) (relcycle_group p X S)"
   unfolding relcycle_group_def
-  apply (rule group.subgroup_of_subgroup_generated)
-  by (auto simp: singular_relcycle subgroup_singular_relboundary intro: singular_relboundary_imp_relcycle)
+  by (simp add: Collect_mono group.subgroup_of_subgroup_generated singular_relboundary_imp_relcycle subgroup_singular_relboundary)
 
 lemma normal_subgroup_singular_relboundary_relcycle:
    "(singular_relboundary_set p X S) \<lhd> (relcycle_group p X S)"
@@ -288,7 +287,7 @@
 lemma hom_boundary2:
   "\<exists>d. (\<forall>p X S.
            (d p X S) \<in> hom (relative_homology_group p X S)
-                           (homology_group (p - 1) (subtopology X S)))
+                           (homology_group (p-1) (subtopology X S)))
      \<and> (\<forall>p X S c. singular_relcycle p X S c \<and> Suc 0 \<le> p
             \<longrightarrow> d p X S (homologous_rel_set p X S c)
               = homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c))"
@@ -323,7 +322,7 @@
   proof -
     obtain d:: "[int, 'a topology, 'a set, ('a chain) set] \<Rightarrow> ('a chain) set"
       where 1: "\<And>p X S. d p X S \<in> hom (relative_homology_group p X S)
-                                      (homology_group (p - 1) (subtopology X S))"
+                                      (homology_group (p-1) (subtopology X S))"
         and 2: "\<And>n X S c. singular_relcycle n X S c \<and> Suc 0 \<le> n
                   \<Longrightarrow> d n X S (homologous_rel_set n X S c)
                     = homologous_rel_set (n - Suc 0) (subtopology X S) {} (chain_boundary n c)"
@@ -337,13 +336,10 @@
       apply (rule_tac x="\<lambda>p X S c.
                if c \<in> carrier(relative_homology_group p X S)
                then d p X (topspace X \<inter> S) c
-               else one(homology_group (p - 1) (subtopology X S))" in exI)
+               else one(homology_group (p-1) (subtopology X S))" in exI)
       apply (simp add: Int_left_absorb subtopology_restrict carrier_relative_homology_group
           group.is_monoid group.restrict_hom_iff 4 cong: if_cong)
-      apply (rule conjI)
-       apply (metis 1 relative_homology_group_restrict subtopology_restrict)
-      apply (metis 2 homologous_rel_restrict singular_relcycle_def subtopology_restrict)
-      done
+      by (metis "1" "2" homologous_rel_restrict relative_homology_group_restrict singular_relcycle_def subtopology_restrict)
   qed
   ultimately show ?thesis
     by auto
@@ -408,25 +404,24 @@
   proof -
     let ?f = "(#>\<^bsub>relcycle_group p Y T\<^esub>) (singular_relboundary_set p Y T) \<circ> chain_map p f"
     let ?F = "\<lambda>x. singular_relboundary_set p X S #>\<^bsub>relcycle_group p X S\<^esub> x"
-    have 1: "?f \<in> hom (relcycle_group p X S) (relative_homology_group (int p) Y T)"
-      apply (rule hom_compose [where H = "relcycle_group p Y T"])
-       apply (metis contf fim hom_chain_map relcycle_group_restrict)
-      by (simp add: nontrivial_relative_homology_group normal.r_coset_hom_Mod normal_subgroup_singular_relboundary_relcycle)
+    have "chain_map p f \<in> hom (relcycle_group p X S) (relcycle_group p Y T)"
+      by (metis contf fim hom_chain_map relcycle_group_restrict)
+    then have 1: "?f \<in> hom (relcycle_group p X S) (relative_homology_group (int p) Y T)"
+      by (simp add: hom_compose normal.r_coset_hom_Mod normal_subgroup_singular_relboundary_relcycle relative_homology_group_def)
     have 2: "singular_relboundary_set p X S \<lhd> relcycle_group p X S"
       using normal_subgroup_singular_relboundary_relcycle by blast
     have 3: "?f x = ?f y"
       if "singular_relcycle p X S x" "singular_relcycle p X S y" "?F x = ?F y" for x y
     proof -
-      have "singular_relboundary p Y T (chain_map p f (x - y))"
-        apply (rule singular_relboundary_chain_map [OF _ contf fim])
-        by (metis homologous_rel_def homologous_rel_eq mem_Collect_eq right_coset_singular_relboundary singular_relboundary_restrict that(3))
+      have "homologous_rel p X S x y"
+        by (metis (no_types) homologous_rel_set_eq right_coset_singular_relboundary that(3))
+      then have "singular_relboundary p Y T (chain_map p f (x - y))"
+        using  singular_relboundary_chain_map [OF _ contf fim] by (simp add: homologous_rel_def)
       then have "singular_relboundary p Y T (chain_map p f x - chain_map p f y)"
         by (simp add: chain_map_diff)
       with that
       show ?thesis
-        apply (simp add: right_coset_singular_relboundary homologous_rel_set_eq)
-        apply (simp add: homologous_rel_def)
-        done
+        by (metis comp_apply homologous_rel_def homologous_rel_set_eq right_coset_singular_relboundary)
     qed
     obtain g where "g \<in> hom (relcycle_group p X S Mod singular_relboundary_set p X S)
                             (relative_homology_group (int p) Y T)"
@@ -501,9 +496,7 @@
                    homologous_rel_set p Y T (chain_map p f c)"
         and 3: "(\<forall>p. p < 0 \<longrightarrow> hom_relmap p = (\<lambda>X S Y T f c. undefined))"
       using hom_induced2 [where ?'a='a and ?'b='b]
-      apply clarify
-      apply (rule_tac hom_relmap=hom_relmap in that, auto)
-      done
+      by (metis (mono_tags, lifting))
     have 4: "\<lbrakk>continuous_map X Y f; f ` (topspace X \<inter> S) \<subseteq> T; c \<in> carrier (relative_homology_group p X S)\<rbrakk> \<Longrightarrow>
         hom_relmap p X (topspace X \<inter> S) Y (topspace Y \<inter> T) f c
            \<in> carrier (relative_homology_group p Y T)"
@@ -676,9 +669,7 @@
 
 lemma abelian_relative_homology_group [simp]:
    "comm_group(relative_homology_group p X S)"
-  apply (simp add: relative_homology_group_def)
-  apply (metis comm_group.abelian_FactGroup abelian_relcycle_group subgroup_singular_relboundary_relcycle)
-  done
+  by (simp add: comm_group.abelian_FactGroup relative_homology_group_def subgroup_singular_relboundary_relcycle)
 
 lemma abelian_homology_group: "comm_group(homology_group p X)"
   by simp
@@ -783,10 +774,7 @@
             continuous_map_from_subtopology)
       show ?thesis
         unfolding q using assms p1 a
-        apply (simp add: ceq assms hom_induced_chain_map hom_boundary_chain_boundary
-                         hom_boundary_chain_boundary [OF sr] singular_relcycle_def mod_subset_def)
-        apply (simp add: p1_eq contf sb cbm hom_induced_chain_map)
-        done
+        by (simp add: cbm ceq contf hom_boundary_chain_boundary hom_induced_chain_map p1_eq sb sr)
     next
       case False
       with assms show ?thesis
@@ -850,10 +838,8 @@
     \<Longrightarrow> f ` A = f ` B" for f A B
           by blast
         have "?lhs = homologous_rel_set (Suc n) X S ` singular_relcycle_set (Suc n) X {}"
-          apply (rule image_cong [OF refl])
-          apply (simp add: o_def hom_induced_chain_map chain_map_ident [of _ X] singular_relcycle
-                 del: of_nat_Suc)
-          done
+          using hom_induced_chain_map chain_map_ident [of _ X] singular_relcycle 
+          by (smt (verit) bot.extremum comp_apply continuous_map_id image_cong image_empty mem_Collect_eq)
         also have "\<dots> = homologous_rel_set (Suc n) X S `
                          (singular_relcycle_set (Suc n) X S \<inter>
                           {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)})"
@@ -861,10 +847,9 @@
           fix c
           assume "c \<in> singular_relcycle_set (Suc n) X {}"
           then show "\<exists>y. y \<in> singular_relcycle_set (Suc n) X S \<inter>
-                 {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)} \<and>
-            homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S y"
-            apply (rule_tac x=c in exI)
-            by (simp add: singular_boundary) (metis chain_boundary_0 singular_cycle singular_relcycle singular_relcycle_0)
+                         {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)} \<and>
+                    homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S y"
+            using singular_cycle singular_relcycle by (fastforce simp: singular_boundary)
         next
           fix c
           assume c: "c \<in> singular_relcycle_set (Suc n) X S \<inter>
@@ -905,9 +890,9 @@
 proof -
   consider (neg) "p \<le> 0" | (int) n where "p = int (Suc n)"
     by (metis linear not0_implies_Suc of_nat_0 zero_le_imp_eq_int)
-  then have "kernel (relative_homology_group (p - 1) (subtopology X S) {})
-                     (relative_homology_group (p - 1) X {})
-                     (hom_induced (p - 1) (subtopology X S) {} X {} id)
+  then have "kernel (relative_homology_group (p-1) (subtopology X S) {})
+                     (relative_homology_group (p-1) X {})
+                     (hom_induced (p-1) (subtopology X S) {} X {} id)
             = hom_boundary p X S ` carrier (relative_homology_group p X S)"
   proof cases
     case neg
@@ -937,9 +922,7 @@
         by (force simp: singular_relcycle singular_boundary chain_boundary_boundary_alt)
       also have "\<dots> = ?rhs"
         unfolding carrier_relative_homology_group vimage_def
-        apply (rule 2)
-        apply (auto simp: hom_induced_chain_map chain_map_ident homologous_rel_set_eq_relboundary singular_relcycle)
-        done
+        by (intro 2) (auto simp: hom_induced_chain_map chain_map_ident homologous_rel_set_eq_relboundary singular_relcycle)
       finally show ?thesis .
     qed
     then show ?thesis
@@ -957,9 +940,8 @@
 proof (cases "p < 0")
   case True
   then show ?thesis
-    apply (simp add: relative_homology_group_def hom_induced_trivial group_hom_def group_hom_axioms_def)
-    apply (auto simp: kernel_def singleton_group_def)
-    done
+    unfolding relative_homology_group_def
+    by (simp add: group_hom.kernel_to_trivial_group group_hom_axioms_def group_hom_def hom_induced_trivial)
 next
   case False
   then obtain n where peq: "p = int n"
@@ -975,11 +957,7 @@
     \<Longrightarrow> f ` A = f ` B" for f A B
       by blast
     have "?lhs = homologous_rel_set n X {} ` (singular_relcycle_set n (subtopology X S) {})"
-      unfolding image_comp o_def
-      apply (rule image_cong [OF refl])
-      apply (simp add: hom_induced_chain_map singular_relcycle)
-       apply (metis chain_map_ident)
-      done
+      by (smt (verit) chain_map_ident continuous_map_id_subt empty_subsetI hom_induced_chain_map image_cong image_empty image_image mem_Collect_eq singular_relcycle)
     also have "\<dots> = homologous_rel_set n X {} ` (singular_relcycle_set n X {} \<inter> singular_relboundary_set n X S)"
     proof (rule 2)
       fix c
@@ -1074,10 +1052,8 @@
       case True
       then obtain a where a: "c = homologous_rel_set n X S a" "singular_relcycle n X S a"
         unfolding carrier_relative_homology_group peq by auto
-      then show ?thesis
-        apply (simp add: peq hom_induced_chain_map_gen cont im homologous_rel_set_eq)
-        apply (blast intro: assms homotopic_imp_homologous_rel_chain_maps)
-        done
+      with assms homotopic_imp_homologous_rel_chain_maps show ?thesis
+        by (force simp add: peq hom_induced_chain_map_gen cont im homologous_rel_set_eq)
     qed (simp add: hom_induced_default)
   qed
 qed
@@ -1166,11 +1142,10 @@
             then show "chain_boundary (Suc n) c1 = c + c2"
             unfolding c1_def c2_def
               by (simp add: algebra_simps chain_boundary_diff)
-            have "singular_chain n (subtopology X (S - U)) c2" "singular_chain n (subtopology X T) c2"
+            obtain "singular_chain n (subtopology X (S - U)) c2" "singular_chain n (subtopology X T) c2"
               using singular_chain_diff c c1 *
               unfolding c1_def c2_def
-               apply (metis add_diff_cancel_left' singular_chain_boundary_alt)
-              by (simp add: e g hSuc singular_chain_boundary_alt singular_chain_diff)
+              by (metis add_diff_cancel_left' e g hSuc singular_chain_boundary_alt)
             then show "singular_chain n (subtopology (subtopology X (S - U)) (T - U)) c2"
               by (fastforce simp add: singular_chain_subtopology)
         qed
@@ -1194,18 +1169,21 @@
       then have scc': "singular_chain n (subtopology X S) c'"
         using homologous_rel_singular_chain singular_relcycle that by blast
       then show ?thesis
-        apply (rule_tac x="c'" in image_eqI)
-         apply (auto simp: scc' chain_map_ident [of _ "subtopology X S"] c' homologous_rel_set_eq)
-        done
+        using scc' chain_map_ident [of _ "subtopology X S"] c' homologous_rel_set_eq
+        by fastforce
     qed
+    have "(\<lambda>x. homologous_rel_set n (subtopology X S) T (chain_map n id x)) `
+          singular_relcycle_set n (subtopology X (S - U)) (T - U) =
+          homologous_rel_set n (subtopology X S) T `
+          singular_relcycle_set n (subtopology X S) T"
+      by (force simp: cont h singular_relcycle_chain_map)
+    then
     show "hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id `
           homologous_rel_set n (subtopology X (S - U)) (T - U) `
           singular_relcycle_set n (subtopology X (S - U)) (T - U)
         = homologous_rel_set n (subtopology X S) T ` singular_relcycle_set n (subtopology X S) T"
-      apply (simp add: image_comp o_def hom_induced_chain_map_gen cont TU topspace_subtopology
+      by (simp add: image_comp o_def hom_induced_chain_map_gen cont TU topspace_subtopology
                        cong: image_cong_simp)
-      apply (force simp: cont h singular_relcycle_chain_map)
-      done
   qed
 qed
 
@@ -1263,10 +1241,10 @@
 qed
 
 lemma relcycle_group_0_eq_chain_group: "relcycle_group 0 X {} = chain_group 0 X"
-  apply (rule monoid.equality, simp)
-     apply (simp_all add: relcycle_group_def chain_group_def)
-  by (metis chain_boundary_def singular_cycle)
-
+proof (rule monoid.equality)
+  show "carrier (relcycle_group 0 X {}) = carrier (chain_group 0 X)"
+    by (simp add: Collect_mono chain_boundary_def singular_cycle subset_antisym)
+qed (simp_all add: relcycle_group_def chain_group_def)
 
 proposition iso_cycle_group_sum:
   assumes disj: "pairwise disjnt \<U>" and UU: "\<Union>\<U> = topspace X"
@@ -1336,10 +1314,8 @@
     have "singular_relcycle_set p X {} \<subseteq> carrier (chain_group p X)"
       using subgroup.subset subgroup_singular_relcycle by blast
     then show "(\<lambda>f. sum' f \<U>) ` (carrier ?SG \<inter> ?PI) = singular_relcycle_set p X {}"
-      using iso B
-      apply (auto simp: iso_def bij_betw_def)
-      apply (force simp: singular_relcycle)
-      done
+      using iso B unfolding Group.iso_def
+      by (smt (verit, del_insts) Int_iff bij_betw_def image_iff mem_Collect_eq subset_antisym subset_iff)  
   qed (auto simp: assms iso_chain_group_sum)
   then show ?thesis
     by (simp add: relcycle_group_def sum_group_subgroup_generated subgroup_singular_relcycle)
@@ -1388,15 +1364,11 @@
           by (simp add: PiE_def Pi_def image_def) metis
         let ?f = "\<lambda>S\<in>\<U>. if singular_relboundary n (subtopology X S) {} (c S) then 0 else c S"
         have "z = (\<lambda>S\<in>\<U>. homologous_rel_set n (subtopology X S) {} (?f S))"
-          apply (simp_all add: c fun_eq_iff PiE_arb [OF z])
-          apply (metis homologous_rel_eq_relboundary singular_boundary singular_relboundary_0)
-          done
+          by (smt (verit) PiE_restrict c homologous_rel_eq_relboundary restrict_apply restrict_ext singular_relboundary_0 z)
         moreover have "?f \<in> (\<Pi>\<^sub>E i\<in>\<U>. singular_relcycle_set n (subtopology X i) {})"
           by (simp add: c fun_eq_iff PiE_arb [OF z])
         moreover have "finite {i \<in> \<U>. ?f i \<noteq> 0}"
-          apply (rule finite_subset [OF _ fin])
-          using z apply (clarsimp simp: PiE_def Pi_def image_def)
-          by (metis c homologous_rel_set_eq_relboundary singular_boundary)
+          using z c  by (intro finite_subset [OF _ fin]) auto 
         ultimately
         show "z \<in> (\<lambda>x. \<lambda>S\<in>\<U>. homologous_rel_set n (subtopology X S) {} (x S)) `
              {x \<in> \<Pi>\<^sub>E i\<in>\<U>. singular_relcycle_set n (subtopology X i) {}. finite {i \<in> \<U>. x i \<noteq> 0}}"
@@ -1413,24 +1385,26 @@
       if z: "z \<in> carrier (sum_group \<U> (\<lambda>S. relcycle_group n (subtopology X S) {}))" for z
     proof -
       have hom_pi: "(\<lambda>S. homologous_rel_set n X {} (z S)) \<in> \<U> \<rightarrow> carrier (homology_group p X)"
-        apply (rule Pi_I)
         using z
-        apply (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle)
-        done
+        by (intro Pi_I) (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle)
       have fin: "finite {S \<in> \<U>. z S \<noteq> 0}"
         using that by (force simp: carrier_sum_group)
       have "?lhs = gfinprod (homology_group p X) (\<lambda>S. homologous_rel_set n X {} (z S)) \<U>"
-        apply (rule gfinprod_cong [OF refl Pi_I])
-         apply (simp add: hom_induced_carrier peq)
-        using that
-           apply (auto simp: peq simp_implies_def carrier_sum_group PiE_def Pi_def chain_map_ident singular_cycle hom_induced_chain_map)
-        done
+      proof (rule gfinprod_cong [OF refl Pi_I])
+        fix i 
+        show "i \<in> \<U> =simp=> hom_induced (int n) (subtopology X i) {} X {} id ((\<lambda>S\<in>\<U>. homologous_rel_set n (subtopology X S) {} (z S)) i)
+                 = homologous_rel_set n X {} (z i)"
+          using that
+          by (auto simp: peq simp_implies_def carrier_sum_group PiE_def Pi_def chain_map_ident singular_cycle hom_induced_chain_map)
+      qed (simp add: hom_induced_carrier peq)
       also have "\<dots> = gfinprod (homology_group p X)
                                 (\<lambda>S. homologous_rel_set n X {} (z S)) {S \<in> \<U>. z S \<noteq> 0}"
-        apply (rule gfinprod_mono_neutral_cong_right, simp_all add: hom_pi)
-        apply (simp add: relative_homology_group_def peq)
-        apply (metis homologous_rel_eq_relboundary singular_relboundary_0)
-        done
+      proof -
+        have "homologous_rel_set n X {} 0 = singular_relboundary_set n X {}"
+          by (metis homologous_rel_eq_relboundary singular_relboundary_0)
+        with hom_pi peq show ?thesis
+          by (intro gfinprod_mono_neutral_cong_right) auto
+      qed
       also have "\<dots> = ?rhs"
       proof -
         have "gfinprod (homology_group p X) (\<lambda>S. homologous_rel_set n X {} (z S)) \<F>
@@ -1440,8 +1414,7 @@
         proof (induction \<F>)
           case empty
           have "\<one>\<^bsub>homology_group p X\<^esub> = homologous_rel_set n X {} 0"
-            apply (simp add: relative_homology_group_def peq)
-            by (metis diff_zero homologous_rel_def homologous_rel_sym)
+            by (metis homologous_rel_eq_relboundary one_relative_homology_group peq singular_relboundary_0)
           then show ?case
             by simp
         next
@@ -1455,10 +1428,13 @@
           show ?case
             using insert z
           proof (simp add: pi)
+            have "\<And>x. homologous_rel n X {} (z S + sum z \<F>) x
+            \<Longrightarrow> \<exists>u v. homologous_rel n X {} (z S) u \<and> homologous_rel n X {} (sum z \<F>) v \<and> x = u + v"
+              by (metis (no_types, lifting) diff_add_cancel diff_diff_eq2 homologous_rel_def homologous_rel_refl)
+            with insert z
             show "homologous_rel_set n X {} (z S) \<otimes>\<^bsub>homology_group p X\<^esub> homologous_rel_set n X {} (sum z \<F>)
               = homologous_rel_set n X {} (z S + sum z \<F>)"
-              using insert z apply (auto simp: peq homologous_rel_add mult_relative_homology_group)
-              by (metis (no_types, lifting) diff_add_cancel diff_diff_eq2 homologous_rel_def homologous_rel_refl)
+              using insert z by (auto simp: peq homologous_rel_add mult_relative_homology_group)
           qed
         qed
         with fin show ?thesis
@@ -1516,10 +1492,8 @@
             by force
         qed
         show ?thesis
-          apply (rule restrict_ext)
           using that *
-          apply (simp add: singular_boundary relative_homology_group_def homologous_rel_set_eq_relboundary peq)
-          done
+          by (force intro!: restrict_ext simp add: singular_boundary relative_homology_group_def homologous_rel_set_eq_relboundary peq)
       qed
       show "x = (\<lambda>S\<in>\<U>. \<one>\<^bsub>homology_group p (subtopology X S)\<^esub>)"
         using x 1 carrSG gf
@@ -1557,12 +1531,10 @@
     "path_connectedin X C" and
     "T \<in> \<U>" and
     "\<not> disjnt C T"
-  then have "C \<subseteq> topspace X"
-    and *: "\<And>B. \<lbrakk>openin X T; T \<inter> B \<inter> C = {}; C \<subseteq> T \<union> B; openin X B\<rbrakk> \<Longrightarrow> B \<inter> C = {}"
-     apply (auto simp: connectedin disjnt_def dest!: path_connectedin_imp_connectedin, blast)
-    done
+  then have  *: "\<And>B. \<lbrakk>openin X T; T \<inter> B \<inter> C = {}; C \<subseteq> T \<union> B; openin X B\<rbrakk> \<Longrightarrow> B \<inter> C = {}"
+    by (meson connectedin disjnt_def disjnt_sym path_connectedin_imp_connectedin)
   have "C \<subseteq> Union \<U>"
-    using \<open>C \<subseteq> topspace X\<close> UU by blast
+    by (simp add: UU \<open>compactin X C\<close> compactin_subset_topspace)
   moreover have "\<Union> (\<U> - {T}) \<inter> C = {}"
   proof (rule *)
     show "T \<inter> \<Union> (\<U> - {T}) \<inter> C = {}"
@@ -1596,11 +1568,12 @@
   have "singular_chain (Suc 0) X (frag_of (restrict (g \<circ> (\<lambda>x. x 0)) (standard_simplex 1)))"
   proof -
     have "continuous_map (subtopology (powertop_real UNIV) (standard_simplex (Suc 0)))
-                         (top_of_set {0..1}) (\<lambda>x. x 0)"
-      apply (auto simp: continuous_map_in_subtopology g)
-        apply (metis (mono_tags) UNIV_I continuous_map_from_subtopology continuous_map_product_projection)
-       apply (simp_all add: standard_simplex_def)
-      done
+                         euclideanreal (\<lambda>x. x 0)"
+      by (metis (mono_tags) UNIV_I continuous_map_from_subtopology continuous_map_product_projection)
+    then have "continuous_map (subtopology (powertop_real UNIV) (standard_simplex (Suc 0)))
+                              (top_of_set {0..1}) (\<lambda>x. x 0)"
+      unfolding continuous_map_in_subtopology g
+      by (auto simp: continuous_map_in_subtopology standard_simplex_def g)
     moreover have "continuous_map (top_of_set {0..1}) X g"
       using contg by blast
     ultimately show ?thesis
@@ -1692,9 +1665,7 @@
       by (simp add: relcycle_group_def chain_group_def group.int_pow_subgroup_generated f)
     show "pow ?H ?q n = homologous_rel_set 0 X {} (frag_cmul n (frag_of f))"
       apply (rule subst [OF right_coset_singular_relboundary])
-      apply (simp add: relative_homology_group_def)
-      apply (simp add: srf ff normal.FactGroup_int_pow normal_subgroup_singular_relboundary_relcycle)
-      done
+      by (simp add: ff normal.FactGroup_int_pow normal_subgroup_singular_relboundary_relcycle relative_homology_group_def srf)
   qed
   show ?thesis
     apply (subst GH.iso_iff)
@@ -1817,17 +1788,13 @@
   assume x: "x \<in> carrier (relative_homology_group p X S)"
   then show "hom_induced p Y T X S g (hom_induced p X S Y T f x) = x"
     using homology_homotopy_axiom [OF gf, of p]
-    apply (simp add: hom_induced_compose [OF contf fim contg gim])
-    apply (metis comp_apply hom_induced_id)
-    done
+    by (simp add: contf contg fim gim hom_induced_compose' hom_induced_id)
 next
   fix y
   assume "y \<in> carrier (relative_homology_group p Y T)"
   then show "hom_induced p X S Y T f (hom_induced p Y T X S g y) = y"
     using homology_homotopy_axiom [OF fg, of p]
-    apply (simp add: hom_induced_compose [OF contg gim contf fim])
-    apply (metis comp_apply hom_induced_id)
-    done
+    by (simp add: contf contg fim gim hom_induced_compose' hom_induced_id)
 qed (auto simp: hom_induced_hom)
 
 
@@ -1846,8 +1813,7 @@
       and "homotopic_with (\<lambda>h. True) X X (g \<circ> f) id"
       and "homotopic_with (\<lambda>k. True) Y Y (f \<circ> g) id"
     shows "(hom_induced p X {} Y {} f) \<in> iso (homology_group p X) (homology_group p Y)"
-  apply (rule homotopy_equivalence_relative_homology_group_isomorphism)
-  using assms by auto
+  using assms by (intro homotopy_equivalence_relative_homology_group_isomorphism) auto
 
 lemma homotopy_equivalent_space_imp_isomorphic_relative_homology_groups:
   assumes "continuous_map X Y f" and fim: "f ` S \<subseteq> T"
@@ -1885,8 +1851,8 @@
 next
   show "hom_induced p (subtopology X S) {} X {} id
         \<in> iso (homology_group p (subtopology X S)) (homology_group p X)"
-       "hom_induced (p - 1) (subtopology X S) {} X {} id
-        \<in> iso (homology_group (p - 1) (subtopology X S)) (homology_group (p - 1) X)"
+       "hom_induced (p-1) (subtopology X S) {} X {} id
+        \<in> iso (homology_group (p-1) (subtopology X S)) (homology_group (p-1) X)"
     using assms
     by (auto intro: homotopy_equivalence_relative_homology_group_isomorphism)
 qed
@@ -1956,10 +1922,7 @@
       by (metis (no_types, lifting) continuous_map_eq continuous_map_id gf id_apply)
     with x show "hom_induced p Y {} X {} g (hom_induced p X {} Y {} f x) = x"
       using hom_induced_compose_empty [OF cont, symmetric]
-      apply (simp add: o_def fun_eq_iff)
-      apply (rule hom_induced_id_gen)
-      apply (auto simp: gf)
-      done
+      by (metis comp_apply cont continuous_map_compose gf hom_induced_id_gen)
   qed
 qed
 
@@ -1983,10 +1946,7 @@
       by (metis (no_types, lifting) continuous_map_eq continuous_map_id fg id_apply)
     with x show ?thesis
       using hom_induced_compose_empty [OF cont, symmetric]
-      apply (simp add: o_def fun_eq_iff)
-      apply (rule hom_induced_id_gen [symmetric])
-        apply (auto simp: fg)
-      done
+      by (metis comp_def cont continuous_map_compose fg hom_induced_id_gen)
   qed
   moreover
   have "(hom_induced p Y {} X {} g x) \<in> carrier (homology_group p X)"
@@ -2011,9 +1971,7 @@
 lemma homeomorphic_map_homology_iso:
   assumes "homeomorphic_map X Y f"
   shows "(hom_induced p X {} Y {} f) \<in> iso (homology_group p X) (homology_group p Y)"
-  using assms
-  apply (simp add: iso_def bij_betw_def flip: section_and_retraction_eq_homeomorphic_map)
-  by (metis inj_on_hom_induced_section_map surj_hom_induced_retraction_map hom_induced_hom)
+  using assms by (simp add: homeomorphic_map_relative_homology_iso)
 
 (*See also hom_hom_induced_inclusion*)
 lemma inj_on_hom_induced_inclusion:
@@ -2037,20 +1995,14 @@
   shows "trivial_homomorphism
              (relative_homology_group p X S) (homology_group (p-1) (subtopology X S))
              (hom_boundary p X S)"
-  apply (rule iffD1 [OF exact_seq_mon_eq_triviality inj_on_hom_induced_inclusion [OF assms]])
-  apply (rule exact_seq.intros)
-     apply (rule homology_exactness_axiom_1 [of p])
-  using homology_exactness_axiom_2 [of p]
-  by auto
+  using exact_seq_mon_eq_triviality inj_on_hom_induced_inclusion [OF assms]
+  by (metis exact_seq_cons_iff homology_exactness_axiom_1 homology_exactness_axiom_2)
 
 lemma epi_hom_induced_relativization:
   assumes "S = {} \<or> S retract_of_space X"
   shows "(hom_induced p X {} X S id) ` carrier (homology_group p X) = carrier (relative_homology_group p X S)"
-  apply (rule iffD2 [OF exact_seq_epi_eq_triviality trivial_homomorphism_hom_boundary_inclusion])
-   apply (rule exact_seq.intros)
-      apply (rule homology_exactness_axiom_1 [of p])
-  using homology_exactness_axiom_2 [of p] apply (auto simp: assms)
-  done
+  using exact_seq_epi_eq_triviality trivial_homomorphism_hom_boundary_inclusion
+  by (metis assms exact_seq_cons_iff homology_exactness_axiom_1 homology_exactness_axiom_2)
 
 (*different in HOL Light because we don't need short_exact_sequence*)
 lemmas short_exact_sequence_hom_induced_inclusion = homology_exactness_axiom_3
@@ -2222,54 +2174,49 @@
 definition hom_relboundary  :: "[int,'a topology,'a set,'a set,'a chain set] \<Rightarrow> 'a chain set"
   where
   "hom_relboundary p X S T =
-    hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id \<circ>
+    hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id \<circ>
     hom_boundary p X S"
 
 lemma group_homomorphism_hom_relboundary:
    "hom_relboundary p X S T
-  \<in> hom (relative_homology_group p X S) (relative_homology_group (p - 1) (subtopology X S) T)"
+  \<in> hom (relative_homology_group p X S) (relative_homology_group (p-1) (subtopology X S) T)"
   unfolding hom_relboundary_def
   proof (rule hom_compose)
-    show "hom_boundary p X S \<in> hom (relative_homology_group p X S) (homology_group(p - 1) (subtopology X S))"
+    show "hom_boundary p X S \<in> hom (relative_homology_group p X S) (homology_group(p-1) (subtopology X S))"
       by (simp add: hom_boundary_hom)
-  show "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id
-      \<in> hom (homology_group(p - 1) (subtopology X S)) (relative_homology_group (p - 1) (subtopology X S) T)"
+  show "hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id
+      \<in> hom (homology_group(p-1) (subtopology X S)) (relative_homology_group (p-1) (subtopology X S) T)"
     by (simp add: hom_induced_hom)
 qed
 
 lemma hom_relboundary:
-   "hom_relboundary p X S T c \<in> carrier (relative_homology_group (p - 1) (subtopology X S) T)"
+   "hom_relboundary p X S T c \<in> carrier (relative_homology_group (p-1) (subtopology X S) T)"
   by (simp add: hom_relboundary_def hom_induced_carrier)
 
 lemma hom_relboundary_empty: "hom_relboundary p X S {} = hom_boundary p X S"
-  apply (simp add: hom_relboundary_def o_def)
-  apply (subst hom_induced_id)
-  apply (metis hom_boundary_carrier, auto)
-  done
+  by (simp add: ext hom_boundary_carrier hom_induced_id hom_relboundary_def)
 
 lemma naturality_hom_induced_relboundary:
   assumes "continuous_map X Y f" "f ` S \<subseteq> U" "f ` T \<subseteq> V"
   shows "hom_relboundary p Y U V \<circ>
          hom_induced p X S Y (U) f =
-         hom_induced (p - 1) (subtopology X S) T (subtopology Y U) V f \<circ>
+         hom_induced (p-1) (subtopology X S) T (subtopology Y U) V f \<circ>
          hom_relboundary p X S T"
 proof -
   have [simp]: "continuous_map (subtopology X S) (subtopology Y U) f"
     using assms continuous_map_from_subtopology continuous_map_in_subtopology topspace_subtopology by fastforce
-  have "hom_induced (p - 1) (subtopology Y U) {} (subtopology Y U) V id \<circ>
-        hom_induced (p - 1) (subtopology X S) {} (subtopology Y U) {} f
-      = hom_induced (p - 1) (subtopology X S) T (subtopology Y U) V f \<circ>
-        hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id"
+  have "hom_induced (p-1) (subtopology Y U) {} (subtopology Y U) V id \<circ>
+        hom_induced (p-1) (subtopology X S) {} (subtopology Y U) {} f
+      = hom_induced (p-1) (subtopology X S) T (subtopology Y U) V f \<circ>
+        hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id"
     using assms by (simp flip: hom_induced_compose)
-  then show ?thesis
-    apply (simp add: hom_relboundary_def comp_assoc naturality_hom_induced assms)
-    apply (simp flip: comp_assoc)
-    done
+  with assms show ?thesis
+    by (metis (no_types, lifting) fun.map_comp hom_relboundary_def naturality_hom_induced)
 qed
 
 proposition homology_exactness_triple_1:
   assumes "T \<subseteq> S"
-  shows "exact_seq ([relative_homology_group(p - 1) (subtopology X S) T,
+  shows "exact_seq ([relative_homology_group(p-1) (subtopology X S) T,
                      relative_homology_group p X S,
                      relative_homology_group p X T],
                     [hom_relboundary p X S T, hom_induced p X T X S id])"
@@ -2279,43 +2226,43 @@
     using assms by auto
   have "?h2 B \<in> kernel ?G2 ?G1 ?h1" for B
   proof -
-    have "hom_boundary p X T B \<in> carrier (relative_homology_group (p - 1) (subtopology X T) {})"
+    have "hom_boundary p X T B \<in> carrier (relative_homology_group (p-1) (subtopology X T) {})"
       by (metis (no_types) hom_boundary)
-    then have *: "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id
-         (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id
+    then have *: "hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id
+         (hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id
          (hom_boundary p X T B))
        = \<one>\<^bsub>?G1\<^esub>"
       using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T]
       by (auto simp: subtopology_subtopology kernel_def)
     show ?thesis
-      apply (simp add: kernel_def hom_induced_carrier hom_relboundary_def flip: *)
-      by (metis comp_def naturality_hom_induced [OF continuous_map_id iTS])
+      using naturality_hom_induced [OF continuous_map_id iTS]
+      by (smt (verit, best) * comp_apply hom_induced_carrier hom_relboundary_def kernel_def mem_Collect_eq)
   qed
   moreover have "B \<in> ?h2 ` carrier ?G3" if "B \<in> kernel ?G2 ?G1 ?h1" for B
   proof -
     have Bcarr: "B \<in> carrier ?G2"
       and Beq: "?h1 B = \<one>\<^bsub>?G1\<^esub>"
       using that by (auto simp: kernel_def)
-    have "\<exists>A' \<in> carrier (homology_group (p - 1) (subtopology X T)). hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id A' = A"
-      if "A \<in> carrier (homology_group (p - 1) (subtopology X S))"
-        "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id A =
+    have "\<exists>A' \<in> carrier (homology_group (p-1) (subtopology X T)). hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id A' = A"
+      if "A \<in> carrier (homology_group (p-1) (subtopology X S))"
+        "hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id A =
       \<one>\<^bsub>?G1\<^esub>" for A
       using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T] that
       by (simp add: kernel_def subtopology_subtopology image_iff set_eq_iff) meson
-    then obtain C where Ccarr: "C \<in> carrier (homology_group (p - 1) (subtopology X T))"
-      and Ceq: "hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id C = hom_boundary p X S B"
+    then obtain C where Ccarr: "C \<in> carrier (homology_group (p-1) (subtopology X T))"
+      and Ceq: "hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id C = hom_boundary p X S B"
       using Beq by (simp add: hom_relboundary_def) (metis hom_boundary_carrier)
-    let ?hi_XT = "hom_induced (p - 1) (subtopology X T) {} X {} id"
+    let ?hi_XT = "hom_induced (p-1) (subtopology X T) {} X {} id"
     have "?hi_XT
-        = hom_induced (p - 1) (subtopology X S) {} X {} id
-            \<circ> (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id)"
+        = hom_induced (p-1) (subtopology X S) {} X {} id
+            \<circ> (hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id)"
       by (metis assms comp_id continuous_map_id_subt hom_induced_compose_empty inf.absorb_iff2 subtopology_subtopology)
     then have "?hi_XT C
-        = hom_induced (p - 1) (subtopology X S) {} X {} id (hom_boundary p X S B)"
+        = hom_induced (p-1) (subtopology X S) {} X {} id (hom_boundary p X S B)"
       by (simp add: Ceq)
-    also have eq: "\<dots> = \<one>\<^bsub>homology_group (p - 1) X\<^esub>"
+    also have eq: "\<dots> = \<one>\<^bsub>homology_group (p-1) X\<^esub>"
       using homology_exactness_axiom_2 [of p X S] Bcarr by (auto simp: kernel_def)
-    finally have "?hi_XT C = \<one>\<^bsub>homology_group (p - 1) X\<^esub>" .
+    finally have "?hi_XT C = \<one>\<^bsub>homology_group (p-1) X\<^esub>" .
     then obtain D where Dcarr: "D \<in> carrier ?G3" and Deq: "hom_boundary p X T D = C"
       using homology_exactness_axiom_2 [of p X T] Ccarr
       by (auto simp: kernel_def image_iff set_eq_iff) meson
@@ -2325,7 +2272,7 @@
     let ?A = "B \<otimes>\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> ?h2 D"
     have "\<exists>A' \<in> carrier (homology_group p X). hom_induced p X {} X S id A' = A"
       if "A \<in> carrier ?G2"
-         "hom_boundary p X S A = one (homology_group (p - 1) (subtopology X S))" for A
+         "hom_boundary p X S A = one (homology_group (p-1) (subtopology X S))" for A
       using that homology_exactness_axiom_1 [of p X S]
       by (simp add: kernel_def subtopology_subtopology image_iff set_eq_iff) meson
     moreover
@@ -2333,7 +2280,7 @@
       by (simp add: Bcarr abelian_relative_homology_group comm_groupE(1) hom_induced_carrier)
     moreover have "hom_boundary p X S (?h2 D) = hom_boundary p X S B"
       by (metis (mono_tags, lifting) Ceq Deq comp_eq_dest continuous_map_id iTS naturality_hom_induced)
-    then have "hom_boundary p X S ?A = one (homology_group (p - 1) (subtopology X S))"
+    then have "hom_boundary p X S ?A = one (homology_group (p-1) (subtopology X S))"
       by (simp add: hom_induced_carrier Bcarr)
     ultimately obtain W where Wcarr: "W \<in> carrier (homology_group p X)"
       and Weq: "hom_induced p X {} X S id W = ?A"
@@ -2343,17 +2290,14 @@
     proof
       interpret comm_group "?G2"
         by (rule abelian_relative_homology_group)
-      have "B = (?h2 \<circ> hom_induced p X {} X T id) W \<otimes>\<^bsub>?G2\<^esub> ?h2 D"
-        apply (simp add: hom_induced_compose [symmetric] assms)
-        by (metis Bcarr Weq hb.G.inv_solve_right hom_induced_carrier)
-      then have "B \<otimes>\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> ?h2 D
-          = ?h2 (hom_induced p X {} X T id W)"
-        by (simp add: hb.G.m_assoc hom_induced_carrier)
+      have "hom_induced p X T X S id (hom_induced p X {} X T id W) = hom_induced p X {} X S id W"
+        by (simp add: assms hom_induced_compose')
+      then have "B = (?h2 \<circ> hom_induced p X {} X T id) W \<otimes>\<^bsub>?G2\<^esub> ?h2 D"
+        by (simp add: Bcarr Weq hb.G.m_assoc hom_induced_carrier)
       then show "B = ?h2 ?W"
-        apply (simp add: Dcarr hom_induced_carrier hom_mult [OF hom_induced_hom])
-        by (metis Bcarr hb.G.inv_solve_right hom_induced_carrier m_comm)
+        by (metis hom_mult [OF hom_induced_hom] Dcarr comp_apply hom_induced_carrier m_comm)
       show "?W \<in> carrier ?G3"
-        by (simp add: Dcarr abelian_relative_homology_group comm_groupE(1) hom_induced_carrier)
+        by (simp add: Dcarr comm_groupE(1) hom_induced_carrier)
     qed
   qed
   ultimately show ?thesis
@@ -2363,19 +2307,19 @@
 
 proposition homology_exactness_triple_2:
   assumes "T \<subseteq> S"
-  shows "exact_seq ([relative_homology_group(p - 1) X T,
-                     relative_homology_group(p - 1) (subtopology X S) T,
+  shows "exact_seq ([relative_homology_group(p-1) X T,
+                     relative_homology_group(p-1) (subtopology X S) T,
                      relative_homology_group p X S],
-                    [hom_induced (p - 1) (subtopology X S) T X T id, hom_relboundary p X S T])"
+                    [hom_induced (p-1) (subtopology X S) T X T id, hom_relboundary p X S T])"
     (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])")
 proof -
-  let ?H2 = "homology_group (p - 1) (subtopology X S)"
+  let ?H2 = "homology_group (p-1) (subtopology X S)"
   have iTS: "id ` T \<subseteq> S" and [simp]: "S \<inter> T = T"
     using assms by auto
   have "?h2 C \<in> kernel ?G2 ?G1 ?h1" for C
   proof -
     have "?h1 (?h2 C)
-       = (hom_induced (p - 1) X {} X T id \<circ> hom_induced (p - 1) (subtopology X S) {} X {} id \<circ> hom_boundary p X S) C"
+       = (hom_induced (p-1) X {} X T id \<circ> hom_induced (p-1) (subtopology X S) {} X {} id \<circ> hom_boundary p X S) C"
       unfolding hom_relboundary_def
       by (metis (no_types, lifting) comp_apply continuous_map_id continuous_map_id_subt empty_subsetI hom_induced_compose id_apply image_empty image_id order_refl)
     also have "\<dots> = \<one>\<^bsub>?G1\<^esub>"
@@ -2387,8 +2331,8 @@
         apply (simp add: kernel_def set_eq_iff)
         by (metis group_relative_homology_group hom_boundary_default hom_one image_eqI)
       ultimately
-      have 1: "hom_induced (p - 1) (subtopology X S) {} X {} id (hom_boundary p X S C)
-             = \<one>\<^bsub>homology_group (p - 1) X\<^esub>"
+      have 1: "hom_induced (p-1) (subtopology X S) {} X {} id (hom_boundary p X S C)
+             = \<one>\<^bsub>homology_group (p-1) X\<^esub>"
         using homology_exactness_axiom_2 [of p X S] by (simp add: kernel_def) blast
       show ?thesis
         by (simp add: 1 hom_one [OF hom_induced_hom])
@@ -2399,58 +2343,60 @@
   qed
   moreover have "x \<in> ?h2 ` carrier ?G3" if "x \<in> kernel ?G2 ?G1 ?h1" for x
   proof -
-    let ?homX = "hom_induced (p - 1) (subtopology X S) {} X {} id"
-    let ?homXS = "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id"
-    have "x \<in> carrier (relative_homology_group (p - 1) (subtopology X S) T)"
+    let ?homX = "hom_induced (p-1) (subtopology X S) {} X {} id"
+    let ?homXS = "hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id"
+    have "x \<in> carrier (relative_homology_group (p-1) (subtopology X S) T)"
       using that by (simp add: kernel_def)
     moreover
     have "hom_boundary (p-1) X T \<circ> hom_induced (p-1) (subtopology X S) T X T id = hom_boundary (p-1) (subtopology X S) T"
       by (metis Int_lower2 \<open>S \<inter> T = T\<close> continuous_map_id_subt hom_relboundary_def hom_relboundary_empty id_apply image_id naturality_hom_induced subtopology_subtopology)
-    then have "hom_boundary (p - 1) (subtopology X S) T x = \<one>\<^bsub>homology_group (p - 2) (subtopology (subtopology X S) T)\<^esub>"
+    then have "hom_boundary (p-1) (subtopology X S) T x = \<one>\<^bsub>homology_group (p - 2) (subtopology (subtopology X S) T)\<^esub>"
       using naturality_hom_induced [of "subtopology X S" X id T T "p-1"] that
         hom_one [OF hom_boundary_hom group_relative_homology_group group_relative_homology_group, of "p-1" X T]
-      apply (simp add: kernel_def subtopology_subtopology)
-      by (metis comp_apply)
+      by (smt (verit) assms comp_apply inf.absorb_iff2 kernel_def mem_Collect_eq subtopology_subtopology)
     ultimately
     obtain y where ycarr: "y \<in> carrier ?H2"
                and yeq: "?homXS y = x"
       using homology_exactness_axiom_1 [of "p-1" "subtopology X S" T]
       by (simp add: kernel_def image_def set_eq_iff) meson
-    have "?homX y \<in> carrier (homology_group (p - 1) X)"
+    have "?homX y \<in> carrier (homology_group (p-1) X)"
       by (simp add: hom_induced_carrier)
     moreover
-    have "(hom_induced (p - 1) X {} X T id \<circ> ?homX) y = \<one>\<^bsub>relative_homology_group (p - 1) X T\<^esub>"
-      apply (simp flip: hom_induced_compose)
-      using hom_induced_compose [of "subtopology X S" "subtopology X S" id "{}" T X id T "p-1"]
-      apply simp
-      by (metis (mono_tags, lifting) kernel_def mem_Collect_eq that yeq)
-    then have "hom_induced (p - 1) X {} X T id (?homX y) = \<one>\<^bsub>relative_homology_group (p - 1) X T\<^esub>"
+    have "(hom_induced (p-1) X {} X T id \<circ> ?homX) y = \<one>\<^bsub>relative_homology_group (p-1) X T\<^esub>"
+      using that 
+      apply (simp add: kernel_def flip: hom_induced_compose)
+      using hom_induced_compose [of "subtopology X S" "subtopology X S" id "{}" T X id T "p-1"] yeq 
+        by auto
+    then have "hom_induced (p-1) X {} X T id (?homX y) = \<one>\<^bsub>relative_homology_group (p-1) X T\<^esub>"
       by simp
-    ultimately obtain z where zcarr: "z \<in> carrier (homology_group (p - 1) (subtopology X T))"
-               and zeq: "hom_induced (p - 1) (subtopology X T) {} X {} id z = ?homX y"
+    ultimately obtain z where zcarr: "z \<in> carrier (homology_group (p-1) (subtopology X T))"
+               and zeq: "hom_induced (p-1) (subtopology X T) {} X {} id z = ?homX y"
       using homology_exactness_axiom_3 [of "p-1" X T]
       by (auto simp: kernel_def dest!: equalityD1 [of "Collect _"])
     have *: "\<And>t. \<lbrakk>t \<in> carrier ?H2;
-                  hom_induced (p - 1) (subtopology X S) {} X {} id t = \<one>\<^bsub>homology_group (p - 1) X\<^esub>\<rbrakk>
+                  hom_induced (p-1) (subtopology X S) {} X {} id t = \<one>\<^bsub>homology_group (p-1) X\<^esub>\<rbrakk>
                   \<Longrightarrow> t \<in> hom_boundary p X S ` carrier ?G3"
       using homology_exactness_axiom_2 [of p X S]
       by (auto simp: kernel_def dest!: equalityD1 [of "Collect _"])
     interpret comm_group "?H2"
       by (rule abelian_relative_homology_group)
-    interpret gh: group_hom ?H2 "homology_group (p - 1) X" "hom_induced (p-1) (subtopology X S) {} X {} id"
+    interpret gh: group_hom ?H2 "homology_group (p-1) X" "hom_induced (p-1) (subtopology X S) {} X {} id"
       by (meson group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced)
-    let ?yz = "y \<otimes>\<^bsub>?H2\<^esub> inv\<^bsub>?H2\<^esub> hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id z"
+    let ?yz = "y \<otimes>\<^bsub>?H2\<^esub> inv\<^bsub>?H2\<^esub> hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id z"
     have yzcarr: "?yz \<in> carrier ?H2"
       by (simp add: hom_induced_carrier ycarr)
-    have yzeq: "hom_induced (p - 1) (subtopology X S) {} X {} id ?yz = \<one>\<^bsub>homology_group (p - 1) X\<^esub>"
-      apply (simp add: hom_induced_carrier ycarr gh.inv_solve_right')
+    have "hom_induced (p-1) (subtopology X S) {} X {} id y =
+          hom_induced (p-1) (subtopology X S) {} X {} id
+            (hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id z)"
       by (metis assms continuous_map_id_subt hom_induced_compose_empty inf.absorb_iff2 o_apply o_id subtopology_subtopology zeq)
+    then have yzeq: "hom_induced (p-1) (subtopology X S) {} X {} id ?yz = \<one>\<^bsub>homology_group (p-1) X\<^esub>"
+      by (simp add: hom_induced_carrier ycarr gh.inv_solve_right')
     obtain w where wcarr: "w \<in> carrier ?G3" and weq: "hom_boundary p X S w = ?yz"
       using * [OF yzcarr yzeq] by blast
     interpret gh2: group_hom ?H2 ?G2 ?homXS
       by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom)
-    have "?homXS (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id z)
-        = \<one>\<^bsub>relative_homology_group (p - 1) (subtopology X S) T\<^esub>"
+    have "?homXS (hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id z)
+        = \<one>\<^bsub>relative_homology_group (p-1) (subtopology X S) T\<^esub>"
       using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T] zcarr
       by (auto simp: kernel_def subtopology_subtopology)
     then show ?thesis
@@ -2500,10 +2446,10 @@
     interpret G2: comm_group "?G2"
       by (rule abelian_relative_homology_group)
     let ?b = "hom_boundary p X T x"
-    have bcarr: "?b \<in> carrier(homology_group(p - 1) (subtopology X T))"
+    have bcarr: "?b \<in> carrier(homology_group(p-1) (subtopology X T))"
       by (simp add: hom_boundary_carrier)
     have "hom_boundary p X S (hom_induced p X T X S id x)
-        = hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id
+        = hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id
             (hom_boundary p X T x)"
       using naturality_hom_induced [of X X id T S p]  by (simp add: assms o_def) meson
     with bcarr have "hom_boundary p X T x \<in> hom_boundary p (subtopology X S) T ` carrier ?G3"
@@ -2518,11 +2464,11 @@
       using x by (simp add: y_def kernel_def hom_induced_carrier)
     interpret hb: group_hom ?G2 "homology_group (p-1) (subtopology X T)" "hom_boundary p X T"
       by (simp add: group_hom_axioms_def group_hom_def hom_boundary_hom)
-    have yyy: "hom_boundary p X T y = \<one>\<^bsub>homology_group (p - 1) (subtopology X T)\<^esub>"
+    have yyy: "hom_boundary p X T y = \<one>\<^bsub>homology_group (p-1) (subtopology X T)\<^esub>"
       apply (simp add: y_def bcarr xcarr hom_induced_carrier hom_boundary_carrier hb.inv_solve_right')
       using naturality_hom_induced [of concl: p X T "subtopology X S" T id]
-      apply (simp add: o_def fun_eq_iff subtopology_subtopology)
-      by (metis hom_boundary_carrier hom_induced_id ueq)
+      by (smt (verit, best) \<open>S \<inter> T = T\<close> bcarr comp_eq_dest continuous_map_id_subt hom_induced_id id_apply 
+          image_subset_iff subtopology_subtopology ueq)
     then have "y \<in> hom_induced p X {} X T id ` carrier (homology_group p X)"
       using homology_exactness_axiom_1 [of p X T] x ycarr by (auto simp: kernel_def)
     then obtain z where zcarr: "z \<in> carrier (homology_group p X)"
@@ -2552,10 +2498,8 @@
             = hom_induced p X {} X T id \<circ> hom_induced p (subtopology X S) {} X {} id"
         by (simp flip: hom_induced_compose)
       show "x = hom_induced p (subtopology X S) T X T id ?u"
-        apply (simp add: hom_mult [OF hom_induced_hom] hom_induced_carrier ucarr)
-        apply (rule *)
-        using eq apply (simp_all add: fun_eq_iff hom_induced_carrier flip: y_def zeq weq)
-        done
+        using hom_mult [OF hom_induced_hom] hom_induced_carrier *
+        by (smt (verit, best) comp_eq_dest eq ucarr weq y_def zeq)
       show "?u \<in> carrier (relative_homology_group p (subtopology X S) T)"
         by (simp add: abelian_relative_homology_group comm_groupE(1) hom_induced_carrier ucarr)
     qed
--- a/src/HOL/Homology/Simplices.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Homology/Simplices.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -1133,11 +1133,7 @@
     using assms by (force simp: simplicial_simplex_def)
   moreover   
   have "singular_face p k f = oriented_simplex (p - Suc 0) (\<lambda>i. if i < k then m i else m (Suc i))"
-    unfolding feq singular_face_def oriented_simplex_def
-    apply (simp add: simplical_face_in_standard_simplex [OF p] restrict_compose_left subset_eq)
-    using sum.zero_middle [OF p, where 'a=real, symmetric]  unfolding simplical_face_def o_def
-    apply (simp add: if_distrib [of "\<lambda>x. _ * x"] if_distrib [of "\<lambda>f. f _ * _"] atLeast0AtMost cong: if_cong)
-    done
+    using feq p singular_face_oriented_simplex by auto
   ultimately
   show ?thesis
     using p simplicial_simplex_def singular_simplex_singular_face by blast
@@ -1174,10 +1170,10 @@
     "\<And>p v l. simplex_cone p v (oriented_simplex p l) =
           oriented_simplex (Suc p) (\<lambda>i. if i = 0 then v else l(i -1))"
 proof -
-  have *: "\<And>x. \<exists>y. \<forall>v. (\<lambda>l. oriented_simplex (Suc x) (\<lambda>i. if i = 0 then v else l (i -1)))
-                  = (y v \<circ> (oriented_simplex x))"
-    apply (subst choice_iff [symmetric])
-    by (simp add: oriented_simplex_eq  choice_iff [symmetric] function_factors_left [symmetric])
+  have *: "\<And>x. \<forall>xv. \<exists>y. (\<lambda>l. oriented_simplex (Suc x)
+                            (\<lambda>i. if i = 0 then xv else l (i - 1))) =
+                  y \<circ> oriented_simplex x"
+    by (simp add: oriented_simplex_eq flip: choice_iff function_factors_left)
   then show ?thesis
     unfolding o_def by (metis(no_types))
 qed
@@ -1528,12 +1524,12 @@
 proof (induction p arbitrary: d c f x y)
   case (Suc p)
   define Sigp where "Sigp \<equiv> \<lambda>f:: (nat \<Rightarrow> real) \<Rightarrow> nat \<Rightarrow> real. \<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j f i) / real (p + 2)"
-  let ?CB = "\<lambda>f. chain_boundary (Suc p) (frag_of f)"
+  define CB where "CB \<equiv> \<lambda>f::(nat \<Rightarrow> real) \<Rightarrow> nat \<Rightarrow> real. chain_boundary (Suc p) (frag_of f)"
   have *: "Poly_Mapping.keys
              (simplicial_cone p (Sigp f)
-               (simplicial_subdivision p (?CB f)))
+               (simplicial_subdivision p (CB f)))
            \<subseteq> {f. \<forall>x\<in>standard_simplex (Suc p). \<forall>y\<in>standard_simplex (Suc p).
-                      \<bar>f x k - f y k\<bar> \<le> real (Suc p) / (real p + 2) * d}" (is "?lhs \<subseteq> ?rhs")
+                      \<bar>f x k - f y k\<bar> \<le> Suc p / (real p + 2) * d}" (is "?lhs \<subseteq> ?rhs")
     if f: "f \<in> Poly_Mapping.keys c" for f
   proof -
     have ssf: "simplicial_simplex (Suc p) S f"
@@ -1541,22 +1537,24 @@
     have 2: "\<And>x y. \<lbrakk>x \<in> standard_simplex (Suc p); y \<in> standard_simplex (Suc p)\<rbrakk> \<Longrightarrow> \<bar>f x k - f y k\<bar> \<le> d"
       by (meson Suc.prems(2) f subsetD le_Suc_eq order_refl standard_simplex_mono)
     have sub: "Poly_Mapping.keys ((frag_of \<circ> simplex_cone p (Sigp f)) g) \<subseteq> ?rhs"
-      if "g \<in> Poly_Mapping.keys (simplicial_subdivision p (?CB f))" for g
+      if "g \<in> Poly_Mapping.keys (simplicial_subdivision p (CB f))" for g
     proof -
-      have 1: "simplicial_chain p S (?CB f)"
+      have 1: "simplicial_chain p S (CB f)"
+        unfolding CB_def
         using ssf simplicial_chain_boundary simplicial_chain_of by fastforce
       have "simplicial_chain (Suc p) (f ` standard_simplex(Suc p)) (frag_of f)"
         by (metis simplicial_chain_of simplicial_simplex ssf subset_refl)
-      then have sc_sub: "Poly_Mapping.keys (?CB f)
+      then have sc_sub: "Poly_Mapping.keys (CB f)
                          \<subseteq> Collect (simplicial_simplex p (f ` standard_simplex (Suc p)))"
-        by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_def)
-      have led: "\<And>h x y. \<lbrakk>h \<in> Poly_Mapping.keys (chain_boundary (Suc p) (frag_of f));
+        by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_def CB_def)
+      have led: "\<And>h x y. \<lbrakk>h \<in> Poly_Mapping.keys (CB f);
                           x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk> \<Longrightarrow> \<bar>h x k - h y k\<bar> \<le> d"
         using Suc.prems(2) f sc_sub
         by (simp add: simplicial_simplex subset_iff image_iff) metis
-      have "\<And>f' x y. \<lbrakk>f' \<in> Poly_Mapping.keys (simplicial_subdivision p (?CB f)); x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk>
+      have "\<And>f' x y. \<lbrakk>f' \<in> Poly_Mapping.keys (simplicial_subdivision p (CB f)); 
+                       x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk>
             \<Longrightarrow> \<bar>f' x k - f' y k\<bar> \<le> (p / (Suc p)) * d"
-        by (blast intro: led Suc.IH [of "chain_boundary (Suc p) (frag_of f)", OF 1])
+        by (blast intro: led Suc.IH [of "CB f", OF 1])
       then have g: "\<And>x y. \<lbrakk>x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk> \<Longrightarrow> \<bar>g x k - g y k\<bar> \<le> (p / (Suc p)) * d"
         using that by blast
       have "d \<ge> 0"
@@ -1668,16 +1666,16 @@
       unfolding simplicial_chain_def simplicial_cone_def
       by (simp add: order_trans [OF keys_frag_extend] sub UN_subset_iff)
   qed
-  show ?case
-    using Suc
-    apply (simp del: sum.atMost_Suc)
-    apply (drule subsetD [OF keys_frag_extend])
-    apply (simp del: sum.atMost_Suc)
-    apply clarify (*OBTAIN?*)
-    apply (rename_tac FFF)
-    using *
-    apply (simp add: add.commute Sigp_def subset_iff)
-    done
+  obtain ff where "ff \<in> Poly_Mapping.keys c"
+            "f \<in> Poly_Mapping.keys
+                  (simplicial_cone p
+                    (\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j ff i) /
+                          (real p + 2))
+                    (simplicial_subdivision p (CB ff)))"
+    using Suc.prems(3) subsetD [OF keys_frag_extend]
+    by (force simp: CB_def simp del: sum.atMost_Suc)
+  then show ?case
+    using Suc * by (simp add: add.commute Sigp_def subset_iff)
 qed (auto simp: standard_simplex_0)
 
 
@@ -1715,22 +1713,27 @@
                  (subtopology (powertop_real UNIV) S) id"
     using continuous_map_from_subtopology_mono continuous_map_id by blast
   moreover have "\<exists>l. restrict id (standard_simplex p) = oriented_simplex p l"
-    apply (rule_tac x="\<lambda>i j. if i = j then 1 else 0" in exI)
-    apply (force simp: oriented_simplex_def standard_simplex_def if_distrib [of "\<lambda>u. u * _"] cong: if_cong)
-    done
+  proof
+    show "restrict id (standard_simplex p) = oriented_simplex p (\<lambda>i j. if i = j then 1 else 0)"
+      by (force simp: oriented_simplex_def standard_simplex_def if_distrib [of "\<lambda>u. u * _"] cong: if_cong)
+  qed
   ultimately show ?lhs
     by (simp add: simplicial_simplex_def singular_simplex_def)
 qed
 
 lemma singular_chain_singular_subdivision:
-   "singular_chain p X c
-        \<Longrightarrow> singular_chain p X (singular_subdivision p c)"
+  assumes "singular_chain p X c"
+  shows "singular_chain p X (singular_subdivision p c)"
   unfolding singular_subdivision_def
-  apply (rule singular_chain_extend)
-  apply (rule singular_chain_chain_map [where X = "subtopology (powertop_real UNIV)
-                          (standard_simplex p)"])
-  apply (simp add: simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain)
-  by (simp add: singular_chain_def singular_simplex_def subset_iff)
+proof (rule singular_chain_extend)
+  fix ca 
+  assume "ca \<in> Poly_Mapping.keys c"
+  with assms have "singular_simplex p X ca"
+    by (simp add: singular_chain_def subset_iff)
+  then show "singular_chain p X (chain_map p ca (simplicial_subdivision p (frag_of (restrict id (standard_simplex p)))))"
+    unfolding singular_simplex_def
+    by (metis order_refl simplicial_chain_of simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain simplicial_simplex_id singular_chain_chain_map)
+qed
 
 lemma naturality_singular_subdivision:
    "singular_chain p X c
@@ -1783,11 +1786,8 @@
   qed
   then show ?case
     using f one
-    apply (auto simp: simplicial_simplex_def)
-    apply (rule singular_simplex_simplex_map
-        [where X = "subtopology (powertop_real UNIV) (standard_simplex q)"])
-    unfolding singular_simplex_def apply (fastforce simp add:)+
-    done
+    apply (simp add: simplicial_simplex_def)
+    using singular_simplex_def singular_simplex_simplex_map by blast
 next
   case (diff a b)
   then show ?case
@@ -1847,9 +1847,15 @@
                 (simplicial_subdivision p
                    (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))
            = simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))"
-      apply (simp add: chain_boundary_chain_map [OF scp] del: chain_map_of
-          flip: singular_simplex_chain_map_id [OF simplicial_imp_singular_simplex [OF ssf]])
-      by (metis (no_types) scp singular_chain_boundary_alt Suc.IH [OF scps] Suc.IH [OF scpf] naturality_singular_subdivision)
+    proof -
+      have "singular_simplex (Suc p) (subtopology (powertop_real UNIV) S) f"
+        using simplicial_simplex_def ssf by blast
+      then  have "chain_map (Suc p) f (frag_of (restrict id (standard_simplex (Suc p)))) = frag_of f"
+        using singular_simplex_chain_map_id by blast
+      then show ?thesis
+        by (metis (no_types) Suc.IH chain_boundary_chain_map diff_Suc_Suc diff_zero 
+             naturality_singular_subdivision scp scpf scps simplicial_imp_singular_chain)
+    qed
     show ?case
       apply (simp add: singular_subdivision_def del: sum.atMost_Suc)
       apply (simp only: ssf 1 2 3 4 chain_map_simplicial_cone [of "Suc p" S _ p "Suc p"])
@@ -1861,8 +1867,8 @@
 lemma naturality_simplicial_subdivision:
    "\<lbrakk>simplicial_chain p (standard_simplex q) c; simplicial_simplex q S g\<rbrakk>
     \<Longrightarrow> simplicial_subdivision p (chain_map p g c) = chain_map p g (simplicial_subdivision p c)"
-apply (simp flip: singular_subdivision_simplicial_simplex)
-  by (metis naturality_singular_subdivision simplicial_chain_chain_map simplicial_imp_singular_chain singular_subdivision_simplicial_simplex)
+  by (metis naturality_singular_subdivision simplicial_chain_chain_map simplicial_imp_singular_chain
+        singular_subdivision_simplicial_simplex)
 
 lemma chain_boundary_singular_subdivision:
    "singular_chain p X c
@@ -1925,12 +1931,21 @@
   by (metis diff_0 subd_0 subd_diff)
 
 lemma subd_power_uminus: "subd p (frag_cmul ((-1) ^ k) c) = frag_cmul ((-1) ^ k) (subd p c)"
-  apply (induction k, simp_all)
-  by (metis minus_frag_cmul subd_uminus)
+proof (induction k)
+  case 0
+  then show ?case by simp
+next
+  case (Suc k)
+  then show ?case
+    by (metis frag_cmul_cmul frag_cmul_minus_one power_Suc subd_uminus)
+qed
 
 lemma subd_power_sum: "subd p (sum f I) = sum (subd p \<circ> f) I"
-  apply (induction I rule: infinite_finite_induct)
-  by auto (metis add_diff_cancel_left' diff_add_cancel subd_diff)
+proof (induction I rule: infinite_finite_induct)
+  case (insert i I)
+  then show ?case
+    by (metis (no_types, lifting) comp_apply diff_minus_eq_add subd_diff subd_uminus sum.insert)
+qed auto
 
 lemma subd: "simplicial_chain p (standard_simplex s) c
      \<Longrightarrow> (\<forall>r g. simplicial_simplex s (standard_simplex r) g \<longrightarrow> chain_map (Suc p) g (subd p c) = subd p (chain_map p g c))
@@ -2007,16 +2022,13 @@
              = subd p (chain_boundary (Suc p) (frag_of (simplex_map (Suc p) g f)))"
         by (metis (no_types) One_nat_def scf Suc.IH chain_boundary_chain_map chain_map_of diff_Suc_Suc diff_zero g simplicial_chain_boundary simplicial_imp_singular_chain)
       show "chain_map (Suc (Suc p)) g (subd (Suc p) (frag_of f)) = subd (Suc p) (chain_map (Suc p) g (frag_of f))"
-        using g
-        apply (simp only: subd.simps frag_extend_of)
+        unfolding subd.simps frag_extend_of
+        using g 
         apply (subst chain_map_simplicial_cone [of s "standard_simplex r" _ "Suc p" s], assumption)
-           apply (intro simplicial_chain_diff)
-        using "1" apply auto[1]
-        using one.hyps apply auto[1]
-        apply (metis Suc.IH diff_Suc_1 mem_Collect_eq one.hyps simplicial_chain_boundary simplicial_chain_of)
+        apply (metis "1" Suc.IH diff_Suc_1 scf simplicial_chain_boundary simplicial_chain_diff)
         using "**" apply auto[1]
          apply (rule order_refl)
-         apply (simp only: chain_map_of frag_extend_of)
+        unfolding chain_map_of frag_extend_of
         apply (rule arg_cong2 [where f = "simplicial_cone (Suc p)"])
          apply (simp add: geq sum_distrib_left oriented_simplex_def ** del: sum.atMost_Suc flip: sum_divide_distrib)
         using 2  apply (simp only: oriented_simplex_def sum.swap [where A = "{..s}"])
@@ -2033,26 +2045,27 @@
         by (metis (no_types) Suc.IH diff_Suc_1 scf simplicial_chain_boundary)
       show "simplicial_chain (Suc (Suc p)) (standard_simplex s) (subd (Suc p) (frag_of f))"
         using one
-        apply (simp only: subd.simps frag_extend_of)
+        unfolding subd.simps frag_extend_of
         apply (rule_tac S="standard_simplex s" in simplicial_chain_simplicial_cone)
-         apply (intro simplicial_chain_diff ff)
-        using sc apply (simp add: algebra_simps)
-        using "**" convex_standard_simplex  apply force+
-        done
+        apply (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision)
+        using "**" convex_standard_simplex by force
       have "simplicial_chain p (standard_simplex s) (chain_boundary (Suc p) (frag_of f))"
         using scf simplicial_chain_boundary by fastforce
       then have "chain_boundary (Suc p) (simplicial_subdivision (Suc p) (frag_of f) - frag_of f
                                          - subd p (chain_boundary (Suc p) (frag_of f))) = 0"
-        apply (simp only: chain_boundary_diff)
-        using Suc.IH chain_boundary_boundary [of "Suc p" "subtopology (powertop_real UNIV)
-                                (standard_simplex s)" "frag_of f"]
-        by (metis One_nat_def add_diff_cancel_left' subd_0 chain_boundary_simplicial_subdivision plus_1_eq_Suc scf simplicial_imp_singular_chain)
-      then show "chain_boundary (Suc (Suc p)) (subd (Suc p) (frag_of f))
+        unfolding chain_boundary_diff
+        using Suc.IH chain_boundary_boundary
+        by (metis One_nat_def add_diff_cancel_left' chain_boundary_simplicial_subdivision diff_Suc_1 scf 
+              simplicial_imp_singular_chain subd_0)
+      moreover have "simplicial_chain (Suc p) (standard_simplex s)
+                       (simplicial_subdivision (Suc p) (frag_of f) - frag_of f -
+                        subd p (chain_boundary (Suc p) (frag_of f)))"
+        by (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision)
+        ultimately show "chain_boundary (Suc (Suc p)) (subd (Suc p) (frag_of f))
           + subd (Suc p - Suc 0) (chain_boundary (Suc p) (frag_of f))
           = simplicial_subdivision (Suc p) (frag_of f) - frag_of f"
-        apply (simp only: subd.simps frag_extend_of)
-        apply (subst chain_boundary_simplicial_cone [of "Suc p" "standard_simplex s"])
-         apply (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision)
+        unfolding subd.simps frag_extend_of
+        apply (simp add: chain_boundary_simplicial_cone )
         apply (simp add: simplicial_cone_def del: sum.atMost_Suc simplicial_subdivision.simps)
         done
     qed
@@ -2062,9 +2075,7 @@
       apply safe
         apply (metis chain_map_diff subd_diff)
        apply (metis simplicial_chain_diff subd_diff)
-      apply (auto simp:  simplicial_subdivision_diff chain_boundary_diff
-          simp del: simplicial_subdivision.simps subd.simps)
-      by (metis (no_types, lifting) add_diff_add add_uminus_conv_diff diff_0 diff_diff_add)
+      by (smt (verit, ccfv_threshold) add_diff_add chain_boundary_diff diff_add_cancel simplicial_subdivision_diff subd_diff)
   qed auto
 qed simp
 
@@ -2169,9 +2180,9 @@
             then show ?case
                 using one
               apply (simp add: chain_map_of singular_simplex_def simplicial_simplex_def, auto)
-                apply (rule_tac f=frag_of in arg_cong, rule)
-                apply (simp add: simplex_map_def)
-                by (simp add: continuous_map_in_subtopology image_subset_iff singular_face_def)
+                apply (rule arg_cong [where f=frag_of])
+                by (auto simp: image_subset_iff simplex_map_def simplicial_simplex singular_face_def)
+
           qed (auto simp: chain_map_diff)
           have "?lhs
                 = chain_map p f
@@ -2347,8 +2358,7 @@
     show "chain_boundary (Suc p) (?h p c) + ?h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ Suc n) c - c"
       using chain_boundary_singular_subdivision [of "Suc p" X]
       apply (simp add: chain_boundary_add f5 h k algebra_simps)
-      apply (smt (verit, ccfv_threshold) add.commute add_diff_eq diff_add_cancel diff_diff_eq2 h(4) k(4) sc singular_subdivision_add)
-      done
+      by (smt (verit, del_insts) add.commute add.left_commute diff_add_cancel h(4) k(4) sc singular_subdivision_add)
   qed (auto simp: k h singular_subdivision_diff)
 qed
 
@@ -2642,8 +2652,14 @@
   show ?thesis
   proof
     show "homologous_rel p X S c ?c'"
-      apply (induction n, simp_all)
-      by (metis p homologous_rel_singular_subdivision homologous_rel_singular_relcycle homologous_rel_trans homologous_rel_sym)
+    proof (induction n)
+      case 0
+      then show ?case by auto
+    next
+      case (Suc n)
+      then show ?case
+        by simp (metis homologous_rel_eq p  homologous_rel_singular_subdivision homologous_rel_singular_relcycle)
+    qed
     then show "singular_relcycle p X S ?c'"
       by (metis homologous_rel_singular_relcycle p)
   next
@@ -2760,8 +2776,7 @@
         case (Suc n)
         then show ?case
           apply simp
-          apply (rule homologous_rel_trans)
-          using c homologous_rel_singular_relcycle_1 homologous_rel_singular_subdivision homologous_rel_sym by blast
+          by (metis c homologous_rel_eq homologous_rel_singular_relcycle_1 homologous_rel_singular_subdivision)
       qed auto
       show "homologous_rel p (subtopology X S) T 0 e"
         unfolding homologous_rel_def using e
@@ -2871,8 +2886,7 @@
     next
       case (diff a b)
       then show ?case
-        apply (simp add: frag_extend_diff keys_diff)
-        using singular_chain_def singular_chain_diff by blast
+        by (simp add: frag_extend_diff singular_chain_diff)
     qed auto
   next
     show "singular_chain (Suc q) (subtopology U V) ((frag_extend \<circ> pr) q c)"
@@ -3079,8 +3093,7 @@
             qed
             have "simp (q - Suc 0) (i - Suc 0) x \<circ> Suc \<in> standard_simplex (q - Suc 0)"
               using ss_ss [OF iq] \<open>i \<le> q\<close> False \<open>i > 0\<close>
-              apply (simp add: simplicial_simplex image_subset_iff)
-              using \<open>x \<in> standard_simplex q\<close> by blast
+              by (simp add: image_subset_iff simplicial_simplex x)
             then show "((\<lambda>z. h (z 0, singular_face q j m (z \<circ> Suc))) \<circ> simp (q - Suc 0) (i - Suc 0)) x
                 = ((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> (simp q i \<circ> simplical_face j)) x"
               by (simp add: singular_face_def \<alpha> \<beta>)
@@ -3189,19 +3202,18 @@
                 then show ?thesis
                   by simp
               qed
+              have xq: "x q = (\<Sum>j\<le>q. if j \<le> i then if q - Suc 0 = j then x j else 0
+                                       else if q = j then x j else 0)" if "q\<noteq>j"
+                using ij that
+                by (force simp flip: ivl_disj_un(2) intro: sum.neutral)
               show ?thesis
-                apply (rule ext)
-                unfolding simplical_face_def using ij
-                apply (auto simp: sum.atMost_Suc cong: if_cong)
-                 apply (force simp flip: ivl_disj_un(2) intro: sum.neutral)
-                 apply (auto simp: *)
-                done
+                using ij unfolding simplical_face_def
+                by (intro ext) (auto simp: * sum.atMost_Suc xq cong: if_cong)
             qed
             show ?thesis
               using False that iq
               unfolding oriented_simplex_def simp_def vv_def ww_def
-              apply (simp add: if_distribR cong: if_cong)
-              apply (simp add: simplical_face_def if_distrib [of "\<lambda>u. u * _"] o_def cong: if_cong)
+              apply (simp add: if_distribR simplical_face_def if_distrib [of "\<lambda>u. u * _"] o_def cong: if_cong)
               apply (simp add: singular_face_def fm ss QQ WW)
               done
           qed
--- a/src/HOL/Library/Groups_Big_Fun.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Library/Groups_Big_Fun.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -22,19 +22,14 @@
 lemma expand_superset:
   assumes "finite A" and "{a. g a \<noteq> \<^bold>1} \<subseteq> A"
   shows "G g = F.F g A"
-  apply (simp add: expand_set)
-  apply (rule F.same_carrierI [of A])
-  apply (simp_all add: assms)
-  done
+  using F.mono_neutral_right assms expand_set by fastforce
 
 lemma conditionalize:
   assumes "finite A"
   shows "F.F g A = G (\<lambda>a. if a \<in> A then g a else \<^bold>1)"
   using assms
-  apply (simp add: expand_set)
-  apply (rule F.same_carrierI [of A])
-  apply auto
-  done
+  by (smt (verit, ccfv_threshold) Diff_iff F.mono_neutral_cong_right expand_set mem_Collect_eq subsetI)
+
 
 lemma neutral [simp]:
   "G (\<lambda>a. \<^bold>1) = \<^bold>1"
@@ -140,30 +135,13 @@
     by (auto simp add: finite_cartesian_product_iff)
   have *: "G (\<lambda>a. G (g a)) =
     (F.F (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) {a. \<exists>b. g a b \<noteq> \<^bold>1})"
-    apply (subst expand_superset [of "?B"])
-    apply (rule \<open>finite ?B\<close>)
-    apply auto
-    apply (subst expand_superset [of "?A"])
-    apply (rule \<open>finite ?A\<close>)
-    apply auto
-    apply (erule F.not_neutral_contains_not_neutral)
-    apply auto
-    done
-  have "{p. (case p of (a, b) \<Rightarrow> g a b) \<noteq> \<^bold>1} \<subseteq> ?A \<times> ?B"
+    using \<open>finite ?A\<close> \<open>finite ?B\<close> expand_superset
+    by (smt (verit, del_insts) Collect_mono local.cong not_neutral_obtains_not_neutral)
+  have **: "{p. (case p of (a, b) \<Rightarrow> g a b) \<noteq> \<^bold>1} \<subseteq> ?A \<times> ?B"
     by auto
-  with subset have **: "{p. (case p of (a, b) \<Rightarrow> g a b) \<noteq> \<^bold>1} \<subseteq> C"
-    by blast
   show ?thesis
-    apply (simp add: *)
-    apply (simp add: F.cartesian_product)
-    apply (subst expand_superset [of C])
-    apply (rule \<open>finite C\<close>)
-    apply (simp_all add: **)
-    apply (rule F.same_carrierI [of C])
-    apply (rule \<open>finite C\<close>)
-    apply (simp_all add: subset)
-    apply auto
-    done
+    using \<open>finite C\<close> expand_superset
+    using "*" ** F.cartesian_product fin_prod by force
 qed
 
 lemma cartesian_product2:
@@ -231,37 +209,23 @@
   fixes r :: "'a :: semiring_0"
   assumes "finite {a. g a \<noteq> 0}"
   shows "Sum_any g * r = (\<Sum>n. g n * r)"
-proof -
-  note assms
-  moreover have "{a. g a * r \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
-  ultimately show ?thesis
-    by (simp add: sum_distrib_right Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
-qed  
+  by (metis (mono_tags, lifting) Collect_mono Sum_any.expand_superset assms mult_zero_left sum_distrib_right)
 
 lemma Sum_any_right_distrib:
   fixes r :: "'a :: semiring_0"
   assumes "finite {a. g a \<noteq> 0}"
   shows "r * Sum_any g = (\<Sum>n. r * g n)"
-proof -
-  note assms
-  moreover have "{a. r * g a \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
-  ultimately show ?thesis
-    by (simp add: sum_distrib_left Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
-qed
+  by (metis (mono_tags, lifting) Collect_mono Sum_any.expand_superset assms mult_zero_right sum_distrib_left)
 
 lemma Sum_any_product:
   fixes f g :: "'b \<Rightarrow> 'a::semiring_0"
   assumes "finite {a. f a \<noteq> 0}" and "finite {b. g b \<noteq> 0}"
   shows "Sum_any f * Sum_any g = (\<Sum>a. \<Sum>b. f a * g b)"
 proof -
-  have subset_f: "{a. (\<Sum>b. f a * g b) \<noteq> 0} \<subseteq> {a. f a \<noteq> 0}"
-    by rule (simp, rule, auto)
-  moreover have subset_g: "\<And>a. {b. f a * g b \<noteq> 0} \<subseteq> {b. g b \<noteq> 0}"
-    by rule (simp, rule, auto)
-  ultimately show ?thesis using assms
-    by (auto simp add: Sum_any.expand_set [of f] Sum_any.expand_set [of g]
-      Sum_any.expand_superset [of "{a. f a \<noteq> 0}"] Sum_any.expand_superset [of "{b. g b \<noteq> 0}"]
-      sum_product)
+  have "\<forall>a. (\<Sum>b. a * g b) = a * Sum_any g"
+    by (simp add: Sum_any_right_distrib assms(2))
+  then show ?thesis
+    by (simp add: Sum_any_left_distrib assms(1))
 qed
 
 lemma Sum_any_eq_zero_iff [simp]: 
--- a/src/HOL/Library/Multiset.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Library/Multiset.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -559,10 +559,7 @@
 
 lemma mset_subset_eq_exists_conv: "(A::'a multiset) \<subseteq># B \<longleftrightarrow> (\<exists>C. B = A + C)"
   unfolding subseteq_mset_def
-  apply (rule iffI)
-   apply (rule exI [where x = "B - A"])
-   apply (auto intro: multiset_eq_iff [THEN iffD2])
-  done
+  by (metis add_diff_cancel_left' count_diff count_union le_Suc_ex le_add_same_cancel1 multiset_eq_iff zero_le)
 
 interpretation subset_mset: ordered_cancel_comm_monoid_diff "(+)" 0 "(\<subseteq>#)" "(\<subset>#)" "(-)"
   by standard (simp, fact mset_subset_eq_exists_conv)
@@ -629,14 +626,18 @@
   by (metis mset_subset_eqD subsetI)
 
 lemma mset_subset_eq_insertD:
-  "add_mset x A \<subseteq># B \<Longrightarrow> x \<in># B \<and> A \<subset># B"
-apply (rule conjI)
- apply (simp add: mset_subset_eqD)
- apply (clarsimp simp: subset_mset_def subseteq_mset_def)
- apply safe
-  apply (erule_tac x = a in allE)
-  apply (auto split: if_split_asm)
-done
+  assumes "add_mset x A \<subseteq># B"
+  shows "x \<in># B \<and> A \<subset># B"
+proof 
+  show "x \<in># B"
+    using assms by (simp add: mset_subset_eqD)
+  have "A \<subseteq># add_mset x A"
+    by (metis (no_types) add_mset_add_single mset_subset_eq_add_left)
+  then have "A \<subset># add_mset x A"
+    by (meson multi_self_add_other_not_self subset_mset.le_imp_less_or_eq)
+  then show "A \<subset># B"
+    using assms subset_mset.strict_trans2 by blast
+qed
 
 lemma mset_subset_insertD:
   "add_mset x A \<subset># B \<Longrightarrow> x \<in># B \<and> A \<subset># B"
@@ -653,11 +654,7 @@
 
 lemma insert_subset_eq_iff:
   "add_mset a A \<subseteq># B \<longleftrightarrow> a \<in># B \<and> A \<subseteq># B - {#a#}"
-  using le_diff_conv2 [of "Suc 0" "count B a" "count A a"]
-  apply (auto simp add: subseteq_mset_def not_in_iff Suc_le_eq)
-  apply (rule ccontr)
-  apply (auto simp add: not_in_iff)
-  done
+  using mset_subset_eq_insertD subset_mset.le_diff_conv2 by fastforce
 
 lemma insert_union_subset_iff:
   "add_mset a A \<subset># B \<longleftrightarrow> a \<in># B \<and> A \<subset># B - {#a#}"
@@ -723,11 +720,12 @@
   by (simp add: union_mset_def)
 
 global_interpretation subset_mset: semilattice_neutr_order \<open>(\<union>#)\<close> \<open>{#}\<close> \<open>(\<supseteq>#)\<close> \<open>(\<supset>#)\<close>
-  apply standard
-      apply (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def max_def)
-   apply (auto simp add: le_antisym dest: sym)
-   apply (metis nat_le_linear)+
-  done
+proof 
+  show "\<And>a b. (b \<subseteq># a) = (a = a \<union># b)"
+    by (simp add: Diff_eq_empty_iff_mset union_mset_def)
+  show "\<And>a b. (b \<subset># a) = (a = a \<union># b \<and> a \<noteq> b)"
+    by (metis Diff_eq_empty_iff_mset add_cancel_left_right subset_mset_def union_mset_def)
+qed (auto simp: multiset_eqI union_mset_def)
 
 interpretation subset_mset: semilattice_sup \<open>(\<union>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close>
 proof -
@@ -774,30 +772,8 @@
 qed
 
 lemma disjunct_not_in:
-  "A \<inter># B = {#} \<longleftrightarrow> (\<forall>a. a \<notin># A \<or> a \<notin># B)" (is "?P \<longleftrightarrow> ?Q")
-proof
-  assume ?P
-  show ?Q
-  proof
-    fix a
-    from \<open>?P\<close> have "min (count A a) (count B a) = 0"
-      by (simp add: multiset_eq_iff)
-    then have "count A a = 0 \<or> count B a = 0"
-      by (cases "count A a \<le> count B a") (simp_all add: min_def)
-    then show "a \<notin># A \<or> a \<notin># B"
-      by (simp add: not_in_iff)
-  qed
-next
-  assume ?Q
-  show ?P
-  proof (rule multiset_eqI)
-    fix a
-    from \<open>?Q\<close> have "count A a = 0 \<or> count B a = 0"
-      by (auto simp add: not_in_iff)
-    then show "count (A \<inter># B) a = count {#} a"
-      by auto
-  qed
-qed
+  "A \<inter># B = {#} \<longleftrightarrow> (\<forall>a. a \<notin># A \<or> a \<notin># B)"
+  by (metis disjoint_iff set_mset_eq_empty_iff set_mset_inter)
 
 lemma inter_mset_empty_distrib_right: "A \<inter># (B + C) = {#} \<longleftrightarrow> A \<inter># B = {#} \<and> A \<inter># C = {#}"
   by (meson disjunct_not_in union_iff)
@@ -1148,15 +1124,7 @@
   fix X :: "'a multiset" and A
   assume "X \<in> A"
   show "Inf A \<subseteq># X"
-  proof (rule mset_subset_eqI)
-    fix x
-    from \<open>X \<in> A\<close> have "A \<noteq> {}" by auto
-    hence "count (Inf A) x = (INF X\<in>A. count X x)"
-      by (simp add: count_Inf_multiset_nonempty)
-    also from \<open>X \<in> A\<close> have "\<dots> \<le> count X x"
-      by (intro cInf_lower) simp_all
-    finally show "count (Inf A) x \<le> count X x" .
-  qed
+    by (metis \<open>X \<in> A\<close> count_Inf_multiset_nonempty empty_iff image_eqI mset_subset_eqI wellorder_Inf_le1)
 next
   fix X :: "'a multiset" and A
   assume nonempty: "A \<noteq> {}" and le: "\<And>Y. Y \<in> A \<Longrightarrow> X \<subseteq># Y"
@@ -1240,8 +1208,7 @@
       using that by (auto simp: subseteq_mset_def algebra_simps not_in_iff)
     hence "Sup A \<subseteq># Sup A - {#x#}" by (intro subset_mset.cSup_least nonempty)
     with \<open>x \<in># Sup A\<close> show False
-      by (auto simp: subseteq_mset_def simp flip: count_greater_zero_iff
-               dest!: spec[of _ x])
+      using mset_subset_diff_self by fastforce
   qed
 next
   fix x X assume "x \<in> set_mset X" "X \<in> A"
@@ -1253,20 +1220,12 @@
 lemma in_Sup_multiset_iff:
   assumes "subset_mset.bdd_above A"
   shows   "x \<in># Sup A \<longleftrightarrow> (\<exists>X\<in>A. x \<in># X)"
-proof -
-  from assms have "set_mset (Sup A) = (\<Union>X\<in>A. set_mset X)" by (rule set_mset_Sup)
-  also have "x \<in> \<dots> \<longleftrightarrow> (\<exists>X\<in>A. x \<in># X)" by simp
-  finally show ?thesis .
-qed
+  by (simp add: assms set_mset_Sup)
 
 lemma in_Sup_multisetD:
   assumes "x \<in># Sup A"
   shows   "\<exists>X\<in>A. x \<in># X"
-proof -
-  have "subset_mset.bdd_above A"
-    by (rule ccontr) (insert assms, simp_all add: Sup_multiset_unbounded)
-  with assms show ?thesis by (simp add: in_Sup_multiset_iff)
-qed
+  using Sup_multiset_unbounded assms in_Sup_multiset_iff by fastforce
 
 interpretation subset_mset: distrib_lattice "(\<inter>#)" "(\<subseteq>#)" "(\<subset>#)" "(\<union>#)"
 proof
@@ -1328,11 +1287,7 @@
 lemma multiset_filter_mono:
   assumes "A \<subseteq># B"
   shows "filter_mset f A \<subseteq># filter_mset f B"
-proof -
-  from assms[unfolded mset_subset_eq_exists_conv]
-  obtain C where B: "B = A + C" by auto
-  show ?thesis unfolding B by auto
-qed
+  by (metis assms filter_sup_mset subset_mset.order_iff)
 
 lemma filter_mset_eq_conv:
   "filter_mset P M = N \<longleftrightarrow> N \<subseteq># M \<and> (\<forall>b\<in>#N. P b) \<and> (\<forall>a\<in>#M - N. \<not> P a)" (is "?P \<longleftrightarrow> ?Q")
@@ -1421,16 +1376,16 @@
   size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified]
 
 lemma size_multiset_empty [simp]: "size_multiset f {#} = 0"
-by (simp add: size_multiset_def)
+  by (simp add: size_multiset_def)
 
 lemma size_empty [simp]: "size {#} = 0"
-by (simp add: size_multiset_overloaded_def)
+  by (simp add: size_multiset_overloaded_def)
 
 lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)"
-by (simp add: size_multiset_eq)
+  by (simp add: size_multiset_eq)
 
 lemma size_single: "size {#b#} = 1"
-by (simp add: size_multiset_overloaded_def size_multiset_single)
+  by (simp add: size_multiset_overloaded_def size_multiset_single)
 
 lemma sum_wcount_Int:
   "finite A \<Longrightarrow> sum (wcount f N) (A \<inter> set_mset N) = sum (wcount f N) A"
@@ -1439,20 +1394,18 @@
 
 lemma size_multiset_union [simp]:
   "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N"
-apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union)
-apply (subst Int_commute)
-apply (simp add: sum_wcount_Int)
-done
+  apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union)
+  by (metis add_implies_diff finite_set_mset inf.commute sum_wcount_Int)
 
 lemma size_multiset_add_mset [simp]:
   "size_multiset f (add_mset a M) = Suc (f a) + size_multiset f M"
-  unfolding add_mset_add_single[of _ M] size_multiset_union by (auto simp: size_multiset_single)
+  by (metis add.commute add_mset_add_single size_multiset_single size_multiset_union)
 
 lemma size_add_mset [simp]: "size (add_mset a A) = Suc (size A)"
-by (simp add: size_multiset_overloaded_def wcount_add_mset)
+  by (simp add: size_multiset_overloaded_def wcount_add_mset)
 
 lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N"
-by (auto simp add: size_multiset_overloaded_def)
+  by (auto simp add: size_multiset_overloaded_def)
 
 lemma size_multiset_eq_0_iff_empty [iff]:
   "size_multiset f M = 0 \<longleftrightarrow> M = {#}"
@@ -1462,23 +1415,15 @@
 by (auto simp add: size_multiset_overloaded_def)
 
 lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)"
-by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty)
+  by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty)
 
 lemma size_eq_Suc_imp_elem: "size M = Suc n \<Longrightarrow> \<exists>a. a \<in># M"
-apply (unfold size_multiset_overloaded_eq)
-apply (drule sum_SucD)
-apply auto
-done
+  using all_not_in_conv by fastforce
 
 lemma size_eq_Suc_imp_eq_union:
   assumes "size M = Suc n"
   shows "\<exists>a N. M = add_mset a N"
-proof -
-  from assms obtain a where "a \<in># M"
-    by (erule size_eq_Suc_imp_elem [THEN exE])
-  then have "M = add_mset a (M - {#a#})" by simp
-  then show ?thesis by blast
-qed
+  by (metis assms insert_DiffM size_eq_Suc_imp_elem)
 
 lemma size_mset_mono:
   fixes A B :: "'a multiset"
@@ -1491,7 +1436,7 @@
 qed
 
 lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \<le> size M"
-by (rule size_mset_mono[OF multiset_filter_subset])
+  by (rule size_mset_mono[OF multiset_filter_subset])
 
 lemma size_Diff_submset:
   "M \<subseteq># M' \<Longrightarrow> size (M' - M) = size M' - size(M::'a multiset)"
@@ -1560,22 +1505,21 @@
 qed (simp add: empty)
 
 lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = add_mset a A"
-by (induct M) auto
+  by (induct M) auto
 
 lemma multiset_cases [cases type]:
-  obtains (empty) "M = {#}"
-    | (add) x N where "M = add_mset x N"
+  obtains (empty) "M = {#}" | (add) x N where "M = add_mset x N"
   by (induct M) simp_all
 
 lemma multi_drop_mem_not_eq: "c \<in># B \<Longrightarrow> B - {#c#} \<noteq> B"
-by (cases "B = {#}") (auto dest: multi_member_split)
+  by (cases "B = {#}") (auto dest: multi_member_split)
 
 lemma union_filter_mset_complement[simp]:
   "\<forall>x. P x = (\<not> Q x) \<Longrightarrow> filter_mset P M + filter_mset Q M = M"
-by (subst multiset_eq_iff) auto
+  by (subst multiset_eq_iff) auto
 
 lemma multiset_partition: "M = {#x \<in># M. P x#} + {#x \<in># M. \<not> P x#}"
-by simp
+  by simp
 
 lemma mset_subset_size: "A \<subset># B \<Longrightarrow> size A < size B"
 proof (induct A arbitrary: B)
@@ -1587,8 +1531,7 @@
   have "add_mset x A \<subseteq># B"
     by (meson add.prems subset_mset_def)
   then show ?case
-    by (metis (no_types) add.prems add.right_neutral add_diff_cancel_left' leD nat_neq_iff
-      size_Diff_submset size_eq_0_iff_empty size_mset_mono subset_mset.le_iff_add subset_mset_def)
+    using add.prems subset_mset.less_eqE by fastforce
 qed
 
 lemma size_1_singleton_mset: "size M = 1 \<Longrightarrow> \<exists>a. M = {#a#}"
@@ -1605,19 +1548,17 @@
 text \<open>Well-foundedness of strict subset relation\<close>
 
 lemma wf_subset_mset_rel: "wf {(M, N :: 'a multiset). M \<subset># N}"
-apply (rule wf_measure [THEN wf_subset, where f1=size])
-apply (clarsimp simp: measure_def inv_image_def mset_subset_size)
-done
+  using mset_subset_size wfP_def wfP_if_convertible_to_nat by blast
 
 lemma wfP_subset_mset[simp]: "wfP (\<subset>#)"
   by (rule wf_subset_mset_rel[to_pred])
 
 lemma full_multiset_induct [case_names less]:
-assumes ih: "\<And>B. \<forall>(A::'a multiset). A \<subset># B \<longrightarrow> P A \<Longrightarrow> P B"
-shows "P B"
-apply (rule wf_subset_mset_rel [THEN wf_induct])
-apply (rule ih, auto)
-done
+  assumes ih: "\<And>B. \<forall>(A::'a multiset). A \<subset># B \<longrightarrow> P A \<Longrightarrow> P B"
+  shows "P B"
+  apply (rule wf_subset_mset_rel [THEN wf_induct])
+  apply (rule ih, auto)
+  done
 
 lemma multi_subset_induct [consumes 2, case_names empty add]:
   assumes "F \<subseteq># A"
@@ -1912,9 +1853,7 @@
   proof (cases "x \<in># B")
     case True
     with add.prems have "image_mset f A = image_mset f (B - {#x#}) + C"
-      by (smt (verit, del_insts) add.left_commute add_cancel_right_left diff_union_cancelL
-          diff_union_single_conv image_mset_union union_mset_add_mset_left
-          union_mset_add_mset_right)
+      by (smt (verit) add_mset_add_mset_same_iff image_mset_add_mset insert_DiffM union_mset_add_mset_left)
     with add.IH have "\<exists>M3'. A = B - {#x#} + M3' \<and> image_mset f M3' = C"
       by (smt (verit, del_insts) True Un_insert_left Un_insert_right add.prems(2) inj_on_insert
           insert_DiffM set_mset_add_mset_insert)
@@ -2035,12 +1974,7 @@
 
 lemma set_eq_iff_mset_remdups_eq:
   \<open>set x = set y \<longleftrightarrow> mset (remdups x) = mset (remdups y)\<close>
-apply (rule iffI)
-apply (simp add: set_eq_iff_mset_eq_distinct[THEN iffD1])
-apply (drule distinct_remdups [THEN distinct_remdups
-      [THEN set_eq_iff_mset_eq_distinct [THEN iffD2]]])
-apply simp
-done
+  using set_eq_iff_mset_eq_distinct by fastforce
 
 lemma mset_eq_imp_distinct_iff:
   \<open>distinct xs \<longleftrightarrow> distinct ys\<close> if \<open>mset xs = mset ys\<close>
@@ -2255,14 +2189,14 @@
 
 lemma set_sorted_list_of_multiset [simp]:
   "set (sorted_list_of_multiset M) = set_mset M"
-by (induct M) (simp_all add: set_insort_key)
+  by (induct M) (simp_all add: set_insort_key)
 
 lemma sorted_list_of_mset_set [simp]:
   "sorted_list_of_multiset (mset_set A) = sorted_list_of_set A"
-by (cases "finite A") (induct A rule: finite_induct, simp_all)
+  by (cases "finite A") (induct A rule: finite_induct, simp_all)
 
 lemma mset_upt [simp]: "mset [m..<n] = mset_set {m..<n}"
-  by (induction n) (simp_all add: atLeastLessThanSuc)
+  by (metis distinct_upt mset_set_set set_upt)
 
 lemma image_mset_map_of:
   "distinct (map fst xs) \<Longrightarrow> {#the (map_of xs i). i \<in># mset (map fst xs)#} = mset (map snd xs)"
@@ -2308,7 +2242,7 @@
   unfolding replicate_mset_def by (induct n) auto
 
 lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})"
-  by (auto split: if_splits)
+  by auto
 
 lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n"
   by (induct n, simp_all)
@@ -2530,8 +2464,8 @@
   shows "size {#x \<in># M. P x#} < size M"
 proof -
   have "size (filter_mset P M) \<noteq> size M"
-    using assms by (metis add.right_neutral add_diff_cancel_left' count_filter_mset mem_Collect_eq
-      multiset_partition nonempty_has_size set_mset_def size_union)
+    using assms
+    by (metis dual_order.strict_iff_order filter_mset_eq_conv mset_subset_size subset_mset.nless_le)
   then show ?thesis
     by (meson leD nat_neq_iff size_filter_mset_lesseq)
 qed
@@ -3956,17 +3890,11 @@
   by (force simp: mult1_def)
 
 lemma union_le_mono2: "B < D \<Longrightarrow> C + B < C + (D::'a::preorder multiset)"
-apply (unfold less_multiset_def multp_def mult_def)
-apply (erule trancl_induct)
- apply (blast intro: mult1_union)
-apply (blast intro: mult1_union trancl_trans)
-done
+  unfolding less_multiset_def multp_def mult_def
+  by (induction rule: trancl_induct; blast intro: mult1_union trancl_trans)
 
 lemma union_le_mono1: "B < D \<Longrightarrow> B + C < D + (C::'a::preorder multiset)"
-apply (subst add.commute [of B C])
-apply (subst add.commute [of D C])
-apply (erule union_le_mono2)
-done
+  by (metis add.commute union_le_mono2)
 
 lemma union_less_mono:
   fixes A B C D :: "'a::preorder multiset"
@@ -3991,8 +3919,8 @@
 definition "ms_weak = ms_strict \<union> Id"
 
 lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)"
-unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def
-by (auto intro: wf_mult1 wf_trancl simp: mult_def)
+  unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def
+  by (auto intro: wf_mult1 wf_trancl simp: mult_def)
 
 lemma smsI:
   "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z + B) \<in> ms_strict"
@@ -4268,11 +4196,7 @@
 
 lemma [code]: \<comment> \<open>not very efficient, but representation-ignorant!\<close>
   "mset_set A = mset (sorted_list_of_set A)"
-  apply (cases "finite A")
-  apply simp_all
-  apply (induct A rule: finite_induct)
-  apply simp_all
-  done
+  by (metis mset_sorted_list_of_multiset sorted_list_of_mset_set)
 
 declare size_mset [code]
 
@@ -4313,18 +4237,15 @@
     hence id: "mset ys = add_mset x (mset (ys1 @ ys2))"
       by auto
     show ?thesis unfolding subset_eq_mset_impl.simps
-      unfolding Some option.simps split
-      unfolding id
-      using Cons[of "ys1 @ ys2"]
-      unfolding subset_mset_def subseteq_mset_def by auto
+      by (simp add: Some id Cons)
   qed
 qed
 
 lemma [code]: "mset xs \<subseteq># mset ys \<longleftrightarrow> subset_eq_mset_impl xs ys \<noteq> None"
-  using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto)
+  by (simp add: subset_eq_mset_impl)
 
 lemma [code]: "mset xs \<subset># mset ys \<longleftrightarrow> subset_eq_mset_impl xs ys = Some True"
-  using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto)
+  using subset_eq_mset_impl by blast
 
 instantiation multiset :: (equal) equal
 begin
@@ -4519,14 +4440,10 @@
   rel: rel_mset
   pred: pred_mset
 proof -
-  show "image_mset id = id"
-    by (rule image_mset.id)
   show "image_mset (g \<circ> f) = image_mset g \<circ> image_mset f" for f g
     unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality)
   show "(\<And>z. z \<in> set_mset X \<Longrightarrow> f z = g z) \<Longrightarrow> image_mset f X = image_mset g X" for f g X
     by (induct X) simp_all
-  show "set_mset \<circ> image_mset f = (`) f \<circ> set_mset" for f
-    by auto
   show "card_order natLeq"
     by (rule natLeq_card_order)
   show "BNF_Cardinal_Arithmetic.cinfinite natLeq"
@@ -4538,34 +4455,15 @@
       (auto simp: finite_iff_ordLess_natLeq[symmetric])
   show "rel_mset R OO rel_mset S \<le> rel_mset (R OO S)" for R S
     unfolding rel_mset_def[abs_def] OO_def
-    apply clarify
-    subgoal for X Z Y xs ys' ys zs
-      apply (drule list_all2_reorder_left_invariance [where xs = ys' and ys = zs and xs' = ys])
-      apply (auto intro: list_all2_trans)
-      done
-    done
+    by (smt (verit, ccfv_SIG) list_all2_reorder_left_invariance list_all2_trans predicate2I)
   show "rel_mset R =
     (\<lambda>x y. \<exists>z. set_mset z \<subseteq> {(x, y). R x y} \<and>
     image_mset fst z = x \<and> image_mset snd z = y)" for R
     unfolding rel_mset_def[abs_def]
-    apply (rule ext)+
-    apply safe
-     apply (rule_tac x = "mset (zip xs ys)" in exI;
-       auto simp: in_set_zip list_all2_iff simp flip: mset_map)
-    apply (rename_tac XY)
-    apply (cut_tac X = XY in ex_mset)
-    apply (erule exE)
-    apply (rename_tac xys)
-    apply (rule_tac x = "map fst xys" in exI)
-    apply (auto simp: mset_map)
-    apply (rule_tac x = "map snd xys" in exI)
-    apply (auto simp: mset_map list_all2I subset_eq zip_map_fst_snd)
-    done
-  show "z \<in> set_mset {#} \<Longrightarrow> False" for z
-    by auto
+    by (metis (no_types, lifting) ex_mset list.in_rel mem_Collect_eq mset_map set_mset_mset)
   show "pred_mset P = (\<lambda>x. Ball (set_mset x) P)" for P
     by (simp add: fun_eq_iff pred_mset_iff)
-qed
+qed auto
 
 inductive rel_mset' :: \<open>('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset \<Rightarrow> bool\<close>
 where
@@ -4609,10 +4507,7 @@
     and addL: "\<And>a M N. P M N \<Longrightarrow> P (add_mset a M) N"
     and addR: "\<And>a M N. P M N \<Longrightarrow> P M (add_mset a N)"
   shows "P M N"
-apply(induct N rule: multiset_induct)
-  apply(induct M rule: multiset_induct, rule empty, erule addL)
-  apply(induct M rule: multiset_induct, erule addR, erule addR)
-done
+  by (induct N rule: multiset_induct; induct M rule: multiset_induct) (auto simp: assms)
 
 lemma multiset_induct2_size[consumes 1, case_names empty add]:
   assumes c: "size M = size N"
--- a/src/HOL/Library/Poly_Mapping.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/Library/Poly_Mapping.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -459,7 +459,7 @@
     {b. \<exists>a. (f a * g b when k = a + b) \<noteq> 0} \<subseteq> {a. f a \<noteq> 0} \<times> {b. g b \<noteq> 0}"
     by auto
   ultimately show ?thesis using fin_g
-    by (auto simp add: prod_fun_def
+    by (auto simp: prod_fun_def
       Sum_any.cartesian_product [of "{a. f a \<noteq> 0} \<times> {b. g b \<noteq> 0}"] Sum_any_right_distrib mult_when)
 qed
 
@@ -526,14 +526,13 @@
       have "?lhs k = (\<Sum>(ab, c). (\<Sum>(a, b). f a * g b when ab = a + b) * h c when k = ab + c)"
         by (simp add: prod_fun_unfold_prod)
       also have "\<dots> = (\<Sum>(ab, c). (\<Sum>(a, b). f a * g b * h c when k = ab + c when ab = a + b))"
-        apply (subst Sum_any_left_distrib)
-        using fin_fg apply (simp add: split_def)
-        apply (subst Sum_any_when_independent [symmetric])
-        apply (simp add: when_when when_mult mult_when split_def conj_commute)
+        using fin_fg 
+        apply (simp add: Sum_any_left_distrib split_def flip: Sum_any_when_independent)
+        apply (simp add: when_when when_mult mult_when conj_commute)
         done
       also have "\<dots> = (\<Sum>(ab, c, a, b). f a * g b * h c when k = ab + c when ab = a + b)"
         apply (subst Sum_any.cartesian_product2 [of "(?FG' \<times> ?H) \<times> ?FG"])
-        apply (auto simp add: finite_cartesian_product_iff fin_fg fin_fg' fin_h dest: mult_not_zero)
+        apply (auto simp: finite_cartesian_product_iff fin_fg fin_fg' fin_h dest: mult_not_zero)
         done
       also have "\<dots> = (\<Sum>(ab, c, a, b). f a * g b * h c when k = a + b + c when ab = a + b)"
         by (rule Sum_any.cong) (simp add: split_def when_def)
@@ -558,7 +557,7 @@
       qed
       also have "\<dots> = (\<Sum>(a, bc). (\<Sum>(b, c). f a * g b * h c when bc = b + c when k = a + bc))"
         apply (subst Sum_any.cartesian_product2 [of "(?F \<times> ?GH') \<times> ?GH"])
-        apply (auto simp add: finite_cartesian_product_iff fin_f fin_gh fin_gh' ac_simps dest: mult_not_zero)
+        apply (auto simp: finite_cartesian_product_iff fin_f fin_gh fin_gh' ac_simps dest: mult_not_zero)
         done
       also have "\<dots> = (\<Sum>(a, bc). f a * (\<Sum>(b, c). g b * h c when bc = b + c) when k = a + bc)"
         apply (subst Sum_any_right_distrib)
@@ -580,10 +579,8 @@
     assume fin_h: "finite {k. h k \<noteq> 0}"
     show "prod_fun (\<lambda>k. f k + g k) h = (\<lambda>k. prod_fun f h k + prod_fun g h k)"
       apply (rule ext)
-      apply (auto simp add: prod_fun_def algebra_simps)
-      apply (subst Sum_any.distrib)
-      using fin_f fin_g apply (auto intro: finite_mult_not_eq_zero_rightI)
-      done
+      apply (simp add: prod_fun_def algebra_simps)
+      by (simp add: Sum_any.distrib fin_f fin_g finite_mult_not_eq_zero_rightI)
   qed
   show "a * (b + c) = a * b + a * c"
   proof transfer
@@ -593,15 +590,8 @@
     assume fin_h: "finite {k. h k \<noteq> 0}"
     show "prod_fun f (\<lambda>k. g k + h k) = (\<lambda>k. prod_fun f g k + prod_fun f h k)"
       apply (rule ext)
-      apply (auto simp add: prod_fun_def Sum_any.distrib algebra_simps when_add_distrib)
-      apply (subst Sum_any.distrib)
-      apply (simp_all add: algebra_simps)
-      apply (auto intro: fin_g fin_h)
-      apply (subst Sum_any.distrib)
-      apply (simp_all add: algebra_simps)
-      using fin_f apply (rule finite_mult_not_eq_zero_rightI)
-      using fin_f apply (rule finite_mult_not_eq_zero_rightI)
-      done
+      apply (auto simp: prod_fun_def Sum_any.distrib algebra_simps when_add_distrib fin_g fin_h)
+      by (simp add: Sum_any.distrib fin_f finite_mult_not_eq_zero_rightI)
   qed
   show "0 * a = 0"
     by transfer (simp add: prod_fun_def [abs_def])
@@ -632,16 +622,10 @@
         using fin_g by auto
       from fin_f fin_g have "finite {(a, b). f a \<noteq> 0 \<and> g b \<noteq> 0}" (is "finite ?AB")
         by simp
-      show "prod_fun f g k = prod_fun g f k"
-        apply (simp add: prod_fun_def)
-        apply (subst Sum_any_right_distrib)
-        apply (rule fin2)
-        apply (subst Sum_any_right_distrib)
-         apply (rule fin1)
-        apply (subst Sum_any.swap [of ?AB])
-        apply (fact \<open>finite ?AB\<close>)
-        apply (auto simp add: mult_when ac_simps)
-        done
+      have "(\<Sum>a. \<Sum>n. f a * (g n when k = a + n)) = (\<Sum>a. \<Sum>n. g a * (f n when k = a + n))"
+        by (subst Sum_any.swap [OF \<open>finite ?AB\<close>]) (auto simp: mult_when ac_simps)
+      then show "prod_fun f g k = prod_fun g f k"
+        by (simp add: prod_fun_def Sum_any_right_distrib [OF fin2] Sum_any_right_distrib [OF fin1])
     qed
   qed
   show "(a + b) * c = a * c + b * c"
@@ -651,12 +635,8 @@
     assume fin_g: "finite {k. g k \<noteq> 0}"
     assume fin_h: "finite {k. h k \<noteq> 0}"
     show "prod_fun (\<lambda>k. f k + g k) h = (\<lambda>k. prod_fun f h k + prod_fun g h k)"
-      apply (auto simp add: prod_fun_def fun_eq_iff algebra_simps)
-      apply (subst Sum_any.distrib)
-      using fin_f apply (rule finite_mult_not_eq_zero_rightI)
-      using fin_g apply (rule finite_mult_not_eq_zero_rightI)
-      apply simp_all
-      done
+      by (auto simp: prod_fun_def fun_eq_iff algebra_simps 
+            Sum_any.distrib fin_f fin_g finite_mult_not_eq_zero_rightI)
   qed
 qed
 
@@ -720,7 +700,7 @@
 
 lemma lookup_single:
   "lookup (single k v) k' = (v when k = k')"
-  by transfer rule
+  by (simp add: single.rep_eq)
 
 lemma lookup_single_eq [simp]:
   "lookup (single k v) k = v"
@@ -810,11 +790,7 @@
 
 lemma lookup_of_int:
   "lookup (of_int l) k = (of_int l when k = 0)"
-proof -
-  have "lookup (of_int l) k = lookup (single 0 (of_int l)) k"
-    by simp
-  then show ?thesis unfolding lookup_single by simp
-qed
+  by (metis lookup_single_not_eq single.rep_eq single_of_int)
 
 
 subsection \<open>Integral domains\<close>
@@ -841,10 +817,10 @@
       by simp
     assume "f \<noteq> (\<lambda>a. 0)"
     then obtain a where "f a \<noteq> 0"
-      by (auto simp add: fun_eq_iff)
+      by (auto simp: fun_eq_iff)
     assume "g \<noteq> (\<lambda>b. 0)"
     then obtain b where "g b \<noteq> 0"
-      by (auto simp add: fun_eq_iff)
+      by (auto simp: fun_eq_iff)
     from \<open>f a \<noteq> 0\<close> and \<open>g b \<noteq> 0\<close> have "F \<noteq> {}" and "G \<noteq> {}"
       by auto
     note Max_F = \<open>finite F\<close> \<open>F \<noteq> {}\<close>
@@ -860,7 +836,7 @@
     define q where "q = Max F + Max G"
     have "(\<Sum>(a, b). f a * g b when q = a + b) =
       (\<Sum>(a, b). f a * g b when q = a + b when a \<in> F \<and> b \<in> G)"
-      by (rule Sum_any.cong) (auto simp add: split_def when_def q_def intro: ccontr)
+      by (rule Sum_any.cong) (auto simp: split_def when_def q_def intro: ccontr)
     also have "\<dots> =
       (\<Sum>(a, b). f a * g b when (Max F, Max G) = (a, b))"
     proof (rule Sum_any.cong)
@@ -873,7 +849,7 @@
         by auto
       show "(case ab of (a, b) \<Rightarrow> f a * g b when q = a + b when a \<in> F \<and> b \<in> G) =
          (case ab of (a, b) \<Rightarrow> f a * g b when (Max F, Max G) = (a, b))"
-        by (auto simp add: split_def when_def q_def dest: add_strict_mono [of a "Max F" b "Max G"])
+        by (auto simp: split_def when_def q_def dest: add_strict_mono [of a "Max F" b "Max G"])
     qed
     also have "\<dots> = (\<Sum>ab. (case ab of (a, b) \<Rightarrow> f a * g b) when
       (Max F, Max G) = ab)"
@@ -883,7 +859,7 @@
     finally have "prod_fun f g q \<noteq> 0"
       by (simp add: prod_fun_unfold_prod)
     then show "prod_fun f g \<noteq> (\<lambda>k. 0)"
-      by (auto simp add: fun_eq_iff)
+      by (auto simp: fun_eq_iff)
   qed
 qed
 
@@ -954,7 +930,7 @@
       by (blast intro: less_funI)
   }
   with * show "less_fun (\<lambda>k. h k + f k) (\<lambda>k. h k + g k) \<or> (\<lambda>k. h k + f k) = (\<lambda>k. h k + g k)"
-    by (auto simp add: fun_eq_iff)
+    by (auto simp: fun_eq_iff)
 qed
 
 instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") linordered_cancel_ab_semigroup_add
@@ -1036,7 +1012,7 @@
 
 lemma range_one [simp]:
   "range 1 = {1}"
-  by transfer (auto simp add: when_def)
+  by transfer (auto simp: when_def)
 
 lemma keys_single [simp]:
   "keys (single k v) = (if v = 0 then {} else {k})"
@@ -1044,13 +1020,12 @@
 
 lemma range_single [simp]:
   "range (single k v) = (if v = 0 then {} else {v})"
-  by transfer (auto simp add: when_def)
+  by transfer (auto simp: when_def)
 
 lemma keys_mult:
   "keys (f * g) \<subseteq> {a + b | a b. a \<in> keys f \<and> b \<in> keys g}"
   apply transfer
-  apply (auto simp add: prod_fun_def dest!: mult_not_zero elim!: Sum_any.not_neutral_obtains_not_neutral)
-  apply blast
+  apply (force simp: prod_fun_def dest!: mult_not_zero elim!: Sum_any.not_neutral_obtains_not_neutral)
   done
 
 lemma setsum_keys_plus_distrib:
@@ -1064,13 +1039,7 @@
 proof -
   let ?A = "Poly_Mapping.keys p \<union> Poly_Mapping.keys q"
   have "?lhs = (\<Sum>k\<in>?A. f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k))"
-    apply(rule sum.mono_neutral_cong_left)
-       apply(simp_all add: Poly_Mapping.keys_add)
-     apply(transfer fixing: f)
-     apply(auto simp add: hom_0)[1]
-    apply(transfer fixing: f)
-    apply(auto simp add: hom_0)[1]
-    done
+    by(intro sum.mono_neutral_cong_left) (auto simp: sum.mono_neutral_cong_left hom_0 in_keys_iff lookup_add)
   also have "\<dots> = (\<Sum>k\<in>?A. f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k))"
     by(rule sum.cong)(simp_all add: hom_plus)
   also have "\<dots> = (\<Sum>k\<in>?A. f k (Poly_Mapping.lookup p k)) + (\<Sum>k\<in>?A. f k (Poly_Mapping.lookup q k))"
@@ -1138,14 +1107,14 @@
   shows "degree f - 1 \<in> keys f"
 proof -
   from assms have "keys f \<noteq> {}"
-    by (auto simp add: degree_def)
+    by (auto simp: degree_def)
   then show ?thesis unfolding degree_def
     by (simp add: mono_Max_commute [symmetric] mono_Suc)
 qed
 
 lemma in_keys_less_degree:
   "n \<in> keys f \<Longrightarrow> n < degree f"
-unfolding degree_def by transfer (auto simp add: Max_gr_iff)
+unfolding degree_def by transfer (auto simp: Max_gr_iff)
 
 lemma beyond_degree_lookup_zero:
   "degree f \<le> n \<Longrightarrow> lookup f n = 0"
@@ -1307,7 +1276,7 @@
   { fix x
     have "prod_fun (\<lambda>k'. s when 0 = k') p x =
           (\<Sum>l :: 'a. if l = 0 then s * (\<Sum>q. p q when x = q) else 0)"
-      by(auto simp add: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta)
+      by(auto simp: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta)
     also have "\<dots> = (\<lambda>k. s * p k when p k \<noteq> 0) x" by(simp add: when_def)
     also note calculation }
   then show "(\<lambda>k. s * p k when p k \<noteq> 0) = prod_fun (\<lambda>k'. s when 0 = k') p"
@@ -1316,10 +1285,10 @@
 
 lemma map_single [simp]:
   "(c = 0 \<Longrightarrow> f 0 = 0) \<Longrightarrow> map f (single x c) = single x (f c)"
-by transfer(auto simp add: fun_eq_iff when_def)
+  by transfer(auto simp: fun_eq_iff when_def)
 
 lemma map_eq_zero_iff: "map f p = 0 \<longleftrightarrow> (\<forall>k \<in> keys p. f (lookup p k) = 0)"
-by transfer(auto simp add: fun_eq_iff when_def)
+  by transfer(auto simp: fun_eq_iff when_def)
 
 subsection \<open>Canonical dense representation of @{typ "nat \<Rightarrow>\<^sub>0 'a"}\<close>
 
@@ -1345,7 +1314,7 @@
 proof (transfer, rule ext)
   fix n :: nat and v :: 'a
   show "nth_default 0 [v] n = (v when 0 = n)"
-    by (auto simp add: nth_default_def nth_append)
+    by (auto simp: nth_default_def nth_append)
 qed
 
 lemma nth_replicate [simp]:
@@ -1353,7 +1322,7 @@
 proof (transfer, rule ext)
   fix m n :: nat and v :: 'a
   show "nth_default 0 (replicate n 0 @ [v]) m = (v when n = m)"
-    by (auto simp add: nth_default_def nth_append)
+    by (auto simp: nth_default_def nth_append)
 qed
 
 lemma nth_strip_while [simp]:
@@ -1379,16 +1348,16 @@
   { fix n
     assume "nth_default 0 xs n \<noteq> 0"
     then have "n < length xs" and "xs ! n \<noteq> 0"
-      by (auto simp add: nth_default_def split: if_splits)
+      by (auto simp: nth_default_def split: if_splits)
     then have "(n, xs ! n) \<in> {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}" (is "?x \<in> ?A")
-      by (auto simp add: in_set_conv_nth enumerate_eq_zip)
+      by (auto simp: in_set_conv_nth enumerate_eq_zip)
     then have "fst ?x \<in> fst ` ?A"
       by blast
     then have "n \<in> fst ` {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}"
       by simp
   }
   then show "{k. nth_default 0 xs k \<noteq> 0} = fst ` {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}"
-    by (auto simp add: in_enumerate_iff_nth_default_eq)
+    by (auto simp: in_enumerate_iff_nth_default_eq)
 qed
 
 lemma range_nth [simp]:
@@ -1409,16 +1378,16 @@
     with * obtain n where n: "n < length xs" "xs ! n \<noteq> 0"
       by (fastforce simp add: no_trailing_unfold last_conv_nth neq_Nil_conv)
     then have "?bound = Max (Suc ` {k. (k < length xs \<longrightarrow> xs ! k \<noteq> (0::'a)) \<and> k < length xs})"
-      by (subst Max_insert) (auto simp add: nth_default_def)
+      by (subst Max_insert) (auto simp: nth_default_def)
     also let ?A = "{k. k < length xs \<and> xs ! k \<noteq> 0}"
     have "{k. (k < length xs \<longrightarrow> xs ! k \<noteq> (0::'a)) \<and> k < length xs} = ?A" by auto
     also have "Max (Suc ` ?A) = Suc (Max ?A)" using n
-      by (subst mono_Max_commute [where f = Suc, symmetric]) (auto simp add: mono_Suc)
+      by (subst mono_Max_commute [where f = Suc, symmetric]) (auto simp: mono_Suc)
     also {
       have "Max ?A \<in> ?A" using n Max_in [of ?A] by fastforce
       hence "Suc (Max ?A) \<le> length xs" by simp
       moreover from * False have "length xs - 1 \<in> ?A"
-        by(auto simp add: no_trailing_unfold last_conv_nth)
+        by(auto simp: no_trailing_unfold last_conv_nth)
       hence "length xs - 1 \<le> Max ?A" using Max_ge[of ?A "length xs - 1"] by auto
       hence "length xs \<le> Suc (Max ?A)" by simp
       ultimately have "Suc (Max ?A) = length xs" by simp }
@@ -1433,7 +1402,7 @@
 lemma nth_idem:
   "nth (List.map (lookup f) [0..<degree f]) = f"
   unfolding degree_def by transfer
-    (auto simp add: nth_default_def fun_eq_iff not_less)
+    (auto simp: nth_default_def fun_eq_iff not_less)
 
 lemma nth_idem_bound:
   assumes "degree f \<le> n"
@@ -1444,7 +1413,7 @@
   then have "[0..<n] = [0..<degree f] @ [degree f..<degree f + m]"
     by (simp add: upt_add_eq_append [of 0])
   moreover have "List.map (lookup f) [degree f..<degree f + m] = replicate m 0"
-    by (rule replicate_eqI) (auto simp add: beyond_degree_lookup_zero)
+    by (rule replicate_eqI) (auto simp: beyond_degree_lookup_zero)
   ultimately show ?thesis by (simp add: nth_idem)
 qed
 
@@ -1489,10 +1458,10 @@
   from assms have "sorted_list_of_set (set (List.map fst xs)) = List.map fst xs"
     unfolding sorted_list_of_set_sort_remdups by (simp add: distinct_remdups_id sort_key_id_if_sorted)
   moreover from assms have "keys (the_value xs) = fst ` set xs"
-    by transfer (auto simp add: image_def split: option.split dest: set_map_of_compr)
+    by transfer (auto simp: image_def split: option.split dest: set_map_of_compr)
   ultimately show ?thesis
     unfolding items_def using assms
-    by (auto simp add: lookup_the_value intro: map_idI)
+    by (auto simp: lookup_the_value intro: map_idI)
 qed
 
 lemma the_value_Nil [simp]:
@@ -1549,7 +1518,7 @@
   then have "f k + g (m k) = (\<Sum>k' \<in> ?keys. f k' + g (m k') when k' = k)"
     by (simp add: sum.delta when_def)
   also have "\<dots> < (\<Sum>k' \<in> ?keys. Suc (f k' + g (m k')))" using *
-    by (intro sum_strict_mono) (auto simp add: when_def)
+    by (intro sum_strict_mono) (auto simp: when_def)
   also have "\<dots> \<le> g 0 + \<dots>" by simp
   finally have "f k + g (m k) < \<dots>" .
   then show "f k + g (m k) < g 0 + (\<Sum>k | m k \<noteq> 0. Suc (f k + g (m k)))"
@@ -1599,7 +1568,7 @@
   "m = m' \<Longrightarrow> g 0 = g' 0 \<Longrightarrow> (\<And>k. k \<in> keys m' \<Longrightarrow> f k = f' k)
     \<Longrightarrow> (\<And>v. v \<in> range m' \<Longrightarrow> g v = g' v)
     \<Longrightarrow> poly_mapping_size f g m = poly_mapping_size f' g' m'"
-  by (auto simp add: poly_mapping_size_def intro!: sum.cong)
+  by (auto simp: poly_mapping_size_def intro!: sum.cong)
 
 instantiation poly_mapping :: (type, zero) size
 begin
@@ -1623,7 +1592,7 @@
 lemma mapp_cong [fundef_cong]:
   "\<lbrakk> m = m'; \<And>k. k \<in> keys m' \<Longrightarrow> f k (lookup m' k) = f' k (lookup m' k) \<rbrakk>
   \<Longrightarrow> mapp f m = mapp f' m'"
-  by transfer (auto simp add: fun_eq_iff)
+  by transfer (auto simp: fun_eq_iff)
 
 lemma lookup_mapp:
   "lookup (mapp f p) k = (f k (lookup p k) when k \<in> keys p)"
@@ -1688,8 +1657,6 @@
   
 
 lemma keys_cmul_iff [iff]: "i \<in> Poly_Mapping.keys (frag_cmul c x) \<longleftrightarrow> i \<in> Poly_Mapping.keys x \<and> c \<noteq> 0"
-  apply auto
-  apply (meson subsetD keys_cmul)
   by (metis in_keys_iff lookup_frag_cmul mult_eq_0_iff)
 
 lemma keys_minus [simp]: "Poly_Mapping.keys(-a) = Poly_Mapping.keys a"
@@ -1697,7 +1664,7 @@
 
 lemma keys_diff: 
   "Poly_Mapping.keys(a - b) \<subseteq> Poly_Mapping.keys a \<union> Poly_Mapping.keys b"
-  by (auto simp add: in_keys_iff lookup_minus)
+  by (auto simp: in_keys_iff lookup_minus)
 
 lemma keys_eq_empty [simp]: "Poly_Mapping.keys c = {} \<longleftrightarrow> c = 0"
   by (metis in_keys_iff keys_zero lookup_zero poly_mapping_eqI)
@@ -1769,10 +1736,11 @@
     by (auto simp: frag_extend_def Poly_Mapping.lookup_add frag_cmul_distrib)
   also have "... = (\<Sum>i \<in> Poly_Mapping.keys a \<union> Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i) 
                          + frag_cmul (poly_mapping.lookup b i) (f i))"
-    apply (rule sum.mono_neutral_cong_left)
-    using keys_add [of a b]
-     apply (auto simp add: in_keys_iff plus_poly_mapping.rep_eq frag_cmul_distrib [symmetric])
-    done
+  proof (rule sum.mono_neutral_cong_left)
+    show "\<forall>i\<in>keys a \<union> keys b - keys (a + b).
+       frag_cmul (lookup a i) (f i) + frag_cmul (lookup b i) (f i) = 0"
+      by (metis DiffD2 frag_cmul_distrib frag_cmul_zero in_keys_iff lookup_add)
+  qed (auto simp: keys_add)
   also have "... = (frag_extend f a) + (frag_extend f b)"
     by (auto simp: * sum.distrib frag_extend_def)
   finally show ?thesis .
@@ -1814,9 +1782,15 @@
   shows "P(frag_cmul k c)"
 proof -
   have "P (frag_cmul (int n) c)" for n
-    apply (induction n)
-     apply (simp_all add: assms frag_cmul_distrib)
-    by (metis add.left_neutral add_diff_cancel_right' add_uminus_conv_diff P)
+  proof (induction n)
+    case 0
+    then show ?case
+      by (simp add: assms)
+  next
+    case (Suc n)
+    then show ?case
+      by (metis assms diff_0 diff_minus_eq_add frag_cmul_distrib frag_cmul_one of_nat_Suc)
+  qed
   then show ?thesis
     by (metis (no_types, opaque_lifting) add_diff_eq assms(2) diff_add_cancel frag_cmul_distrib int_diff_cases)
 qed
@@ -1840,14 +1814,11 @@
       by simp
     have Pfrag: "P (frag_cmul (poly_mapping.lookup c i) (frag_of i))"
       by (metis "0" diff frag_closure_minus_cmul insert.prems insert_subset sing subset_iff)
-    show ?case
-      apply (simp add: insert.hyps)
-      apply (subst ab)
-      using insert apply (blast intro: assms Pfrag)
-      done
+    with insert show ?case
+      by (metis (mono_tags, lifting) "0" ab diff insert_subset sum.insert)
   qed
   then show ?thesis
-    by (subst frag_expansion) (auto simp add: frag_extend_def)
+    by (subst frag_expansion) (auto simp: frag_extend_def)
 qed
 
 lemma frag_extend_compose:
--- a/src/HOL/ex/Sketch_and_Explore.thy	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/HOL/ex/Sketch_and_Explore.thy	Tue Apr 16 17:28:58 2024 +0200
@@ -6,7 +6,7 @@
 
 theory Sketch_and_Explore
   imports Main \<comment> \<open>TODO: generalize existing sledgehammer functions to Pure\<close>
-  keywords "sketch" "explore" :: diag
+  keywords "sketch" "explore" "sketch_subgoals" :: diag
 begin
 
 ML \<open>
@@ -82,6 +82,27 @@
       :: map (print_isar_skeleton ctxt 2 "moreover have") clauses
       @ ["  ultimately show ?thesis", "    by" ^ method_text, "qed"];
 
+fun print_subgoal apply_line_opt (is_prems, is_for, is_sh) ctxt indent stmt =
+  let
+    val ((fixes, _, _), ctxt') = eigen_context_for_statement stmt ctxt;
+    val prefix = replicate_string indent " ";
+    val fixes_s =
+      if not is_for orelse null fixes then NONE
+      else SOME ("for " ^ space_implode " "
+        (map (fn (v, _) => Proof_Context.print_name ctxt' v) fixes));
+    val premises_s = if is_prems then SOME "premises prems" else NONE;
+    val sh_s = if is_sh then SOME "sledgehammer" else NONE;
+    val subgoal_s = map_filter I [SOME "subgoal", premises_s, fixes_s]
+      |> space_implode " ";
+    val using_s = if is_prems then SOME "using prems" else NONE;
+    val s = cat_lines (
+      [subgoal_s]
+      @ map_filter (Option.map (fn s => prefix ^ s)) [using_s, apply_line_opt, sh_s]
+      @ [prefix ^ "sorry"]);
+  in
+    s
+  end;
+
 fun coalesce_method_txt [] = ""
   | coalesce_method_txt [s] = s
   | coalesce_method_txt (s1 :: s2 :: ss) =
@@ -90,6 +111,9 @@
       then s1 ^ coalesce_method_txt (s2 :: ss)
       else s1 ^ " " ^ coalesce_method_txt (s2 :: ss);
 
+fun print_subgoals options apply_line_opt ctxt _ clauses =
+  separate "" (map (print_subgoal apply_line_opt options ctxt 2) clauses) @ ["done"];
+
 fun print_proof_text_from_state print (some_method_ref : ((Method.text * Position.range) * Token.T list) option) state =
   let
     val state' = state
@@ -120,12 +144,31 @@
 
 fun explore method_ref = print_proof_text_from_state print_exploration (SOME method_ref);
 
+fun subgoals options method_ref =
+  let
+    val apply_line_opt = case method_ref of
+        NONE => NONE
+      | SOME (_, toks) => SOME ("apply " ^ coalesce_method_txt (map Token.unparse toks))
+    val succeed_method_ref = SOME ((Method.succeed_text, Position.no_range), [])
+  in
+    print_proof_text_from_state (print_subgoals options apply_line_opt) succeed_method_ref
+  end;
+
 fun sketch_cmd some_method_text =
   Toplevel.keep_proof (K () o sketch some_method_text o Toplevel.proof_of)
 
 fun explore_cmd method_text =
   Toplevel.keep_proof (K () o explore method_text o Toplevel.proof_of)
 
+fun subgoals_cmd (modes, method_ref) =
+  let
+    val is_prems = not (member (op =) modes "noprems")
+    val is_for = not (member (op =) modes "nofor")
+    val is_sh = member (op =) modes "sh"
+  in
+    Toplevel.keep_proof (K () o subgoals (is_prems, is_for, is_sh) method_ref o Toplevel.proof_of)
+  end
+
 val _ =
   Outer_Syntax.command \<^command_keyword>\<open>sketch\<close>
     "print sketch of Isar proof text after method application"
@@ -135,6 +178,14 @@
   Outer_Syntax.command \<^command_keyword>\<open>explore\<close>
     "explore proof obligations after method application as Isar proof text"
     (Scan.trace Method.parse >> explore_cmd);
+
+val opt_modes =
+  Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) [];
+
+val _ =
+  Outer_Syntax.command \<^command_keyword>\<open>sketch_subgoals\<close>
+    "sketch proof obligations as subgoals, applying a method and/or sledgehammer to each"
+    (opt_modes -- Scan.option (Scan.trace Method.parse) >> subgoals_cmd);
 \<close>
 
 end
--- a/src/Pure/Admin/ci_build.scala	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/Pure/Admin/ci_build.scala	Tue Apr 16 17:28:58 2024 +0200
@@ -136,7 +136,7 @@
     Mercurial.repository(path).id()
 
   def print_section(title: String): Unit =
-    println(s"\n=== $title ===\n")
+    println("\n=== " + title + " ===\n")
 
   def ci_build(options: Options, job: Job): Unit = {
     val profile = job.profile
@@ -156,11 +156,12 @@
       with_documents(options, config).int.update("threads", profile.threads) +
         "parallel_proofs=1" + "system_heaps"
 
-    println(s"jobs = ${profile.jobs}, threads = ${profile.threads}, numa = ${profile.numa}")
+    println(
+      "jobs = " + profile.jobs + ", threads = " + profile.threads + ", numa = " + profile.numa)
 
     print_section("BUILD")
-    println(s"Build started at $formatted_time")
-    println(s"Isabelle id $isabelle_id")
+    println("Build started at " + formatted_time)
+    println("Isabelle id " + isabelle_id)
     val pre_result = config.pre_hook(options)
 
     print_section("LOG")
@@ -186,7 +187,7 @@
 
     val groups = results.sessions.map(results.info).flatMap(_.groups)
     for (group <- groups)
-      println(s"Group $group: " + compute_timing(results, Some(group)).message_resources)
+      println("Group " + group + ": " + compute_timing(results, Some(group)).message_resources)
 
     val total_timing = compute_timing(results, None).copy(elapsed = elapsed_time)
     println("Overall: " + total_timing.message_resources)
@@ -195,10 +196,10 @@
       print_section("FAILED SESSIONS")
 
       for (name <- results.sessions) {
-        if (results.cancelled(name)) println(s"Session $name: CANCELLED")
+        if (results.cancelled(name)) println("Session " + name + ": CANCELLED")
         else {
           val result = results(name)
-          if (!result.ok) println(s"Session $name: FAILED ${ result.rc }")
+          if (!result.ok) println("Session " + name + ": FAILED " + result.rc)
         }
       }
     }
--- a/src/Pure/System/web_app.scala	Tue Apr 16 17:06:05 2024 +0200
+++ b/src/Pure/System/web_app.scala	Tue Apr 16 17:28:58 2024 +0200
@@ -179,12 +179,7 @@
   /* request parameters as structured data */
 
   object Params {
-    def key(elems: (String | Int)*): Key =
-      Key(
-        elems.toList.map {
-          case s: String => Key.elem(s)
-          case i: Int => Key.elem(i)
-        })
+    def key(s: String): Key = Key(List(Key.elem(s)))
 
     object Key {
       sealed trait Elem { def print: String }
@@ -213,21 +208,25 @@
         override def toString: String = print
       }
 
-      def elem(s: String): Nested_Elem =
-        if (!s.forall(Symbol.is_ascii_letter)) error("Illegal element in: " + quote(s))
-        else new Nested_Elem(s)
-
-      def elem(i: Int): List_Elem =
-        if (i < 0) error("Illegal list element") else new List_Elem(i)
+      def elem(s: String): Elem =
+        if (s.contains('/')) error("Illegal separator in " + quote(s))
+        else {
+          for {
+            c <- s
+            if Symbol.is_ascii_blank(c)
+          } error("Illegal blank character " + c.toInt + " in " + quote(s))
+          s match {
+            case Value.Int(i) => new List_Elem(i)
+            case s => new Nested_Elem(s)
+          }
+        }
 
-      def empty: Key = Key(Nil)
-      def apply(s: String): Key = new Key(List(elem(s)))
-      def explode(s: String): Key =
-        new Key(
-          space_explode('_', s).map {
-            case Value.Int(i) => elem(i)
-            case s => elem(s)
-          })
+      def elem(i: Int): List_Elem = if (i < 0) error("Illegal list element") else new List_Elem(i)
+
+      def is_blank(es: List[Elem]): Boolean = es.length <= 1 && es.forall(_.print.isBlank)
+      def apply(es: List[Elem]): Key = if (is_blank(es)) error("Empty key") else new Key(es)
+
+      def explode(s: String): Key = Key(space_explode('/', s).map(elem))
 
 
       /* ordering */
@@ -237,7 +236,7 @@
       }
     }
 
-    case class Key(rep: List[Key.Elem]) {
+    class Key(private val rep: List[Key.Elem]) {
       def +(elem: Key.Elem): Key = Key(rep :+ elem)
       def +(i: Int): Key = this + Key.elem(i)
       def +(s: String): Key = this + Key.elem(s)
@@ -250,10 +249,22 @@
         }
 
       def get(key: Key): Option[Key] =
-        if (rep.startsWith(key.rep)) Some(Key(rep.drop(key.rep.length))) else None
+        if (!rep.startsWith(key.rep)) None
+        else {
+          val rest = rep.drop(key.rep.length)
+          if (Key.is_blank(rest)) None
+          else Some(Key(rest))
+        }
 
-      def print = rep.map(_.print).mkString("_")
+      def print = rep.map(_.print).mkString("/")
+
       override def toString: String = print
+      override def equals(that: Any): Boolean =
+        that match {
+          case other: Key => rep == other.rep
+          case _ => false
+        }
+      override def hashCode(): Int = rep.hashCode()
     }
 
     def indexed[A, B](key: Key, xs: List[A], f: (Key, A) => B): List[B] =