redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
authorimmler
Thu, 17 Jan 2019 16:28:07 -0500
changeset 69681 689997a8a582
parent 69680 96a43caa4902
child 69682 500a7acdccd4
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
src/HOL/Analysis/Cross3.thy
src/HOL/Analysis/Extended_Real_Limits.thy
src/HOL/Analysis/Fashoda_Theorem.thy
src/HOL/Analysis/Finite_Cartesian_Product.thy
src/HOL/Analysis/Finite_Product_Measure.thy
src/HOL/Analysis/Function_Topology.thy
src/HOL/Analysis/Further_Topology.thy
src/HOL/Analysis/Great_Picard.thy
src/HOL/Analysis/Homeomorphism.thy
src/HOL/Analysis/Improper_Integral.thy
src/HOL/Analysis/Interval_Integral.thy
src/HOL/Analysis/Tagged_Division.thy
--- a/src/HOL/Analysis/Cross3.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Cross3.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -79,14 +79,14 @@
 
 hide_fact (open) left_diff_distrib right_diff_distrib
 
-lemma%important  Jacobi: "x \<times> (y \<times> z) + y \<times> (z \<times> x) + z \<times> (x \<times> y) = 0" for x::"real^3"
+proposition Jacobi: "x \<times> (y \<times> z) + y \<times> (z \<times> x) + z \<times> (x \<times> y) = 0" for x::"real^3"
   by%unimportant (simp add: cross3_simps)
 
-lemma%important  Lagrange: "x \<times> (y \<times> z) = (x \<bullet> z) *\<^sub>R y - (x \<bullet> y) *\<^sub>R z"
+proposition  Lagrange: "x \<times> (y \<times> z) = (x \<bullet> z) *\<^sub>R y - (x \<bullet> y) *\<^sub>R z"
   by%unimportant (simp add: cross3_simps) (metis (full_types) exhaust_3)
 
-lemma  cross_triple: "(x \<times> y) \<bullet> z = (y \<times> z) \<bullet> x"
-  by (simp add: cross3_def inner_vec_def sum_3 vec_eq_iff algebra_simps)
+proposition  cross_triple: "(x \<times> y) \<bullet> z = (y \<times> z) \<bullet> x"
+  by%unimportant (simp add: cross3_def inner_vec_def sum_3 vec_eq_iff algebra_simps)
 
 lemma  cross_components:
    "(x \<times> y)$1 = x$2 * y$3 - y$2 * x$3" "(x \<times> y)$2 = x$3 * y$1 - y$3 * x$1" "(x \<times> y)$3 = x$1 * y$2 - y$1 * x$2"
@@ -126,15 +126,15 @@
 lemma  cross_cross_det: "(w \<times> x) \<times> (y \<times> z) = det(vector[w,x,z]) *\<^sub>R y - det(vector[w,x,y]) *\<^sub>R z"
   using exhaust_3 by (force simp add: cross3_simps) 
 
-lemma%important  dot_cross: "(w \<times> x) \<bullet> (y \<times> z) = (w \<bullet> y) * (x \<bullet> z) - (w \<bullet> z) * (x \<bullet> y)"
+proposition  dot_cross: "(w \<times> x) \<bullet> (y \<times> z) = (w \<bullet> y) * (x \<bullet> z) - (w \<bullet> z) * (x \<bullet> y)"
   by%unimportant (force simp add: cross3_simps) 
 
-lemma  norm_cross: "(norm (x \<times> y))\<^sup>2 = (norm x)\<^sup>2 * (norm y)\<^sup>2 - (x \<bullet> y)\<^sup>2"
+proposition  norm_cross: "(norm (x \<times> y))\<^sup>2 = (norm x)\<^sup>2 * (norm y)\<^sup>2 - (x \<bullet> y)\<^sup>2"
   unfolding power2_norm_eq_inner power_mult_distrib
   by (simp add: cross3_simps power2_eq_square)
 
-lemma%important  cross_eq_0: "x \<times> y = 0 \<longleftrightarrow> collinear{0,x,y}"
-proof%unimportant -
+lemma  cross_eq_0: "x \<times> y = 0 \<longleftrightarrow> collinear{0,x,y}"
+proof -
   have "x \<times> y = 0 \<longleftrightarrow> norm (x \<times> y) = 0"
     by simp
   also have "... \<longleftrightarrow> (norm x * norm y)\<^sup>2 = (x \<bullet> y)\<^sup>2"
@@ -175,10 +175,10 @@
   apply (simp add: vector_matrix_mult_def matrix_vector_mult_def forall_3 cross3_simps)
   done
 
-lemma%important  cross_orthogonal_matrix:
+lemma  cross_orthogonal_matrix:
   assumes "orthogonal_matrix A"
   shows "(A *v x) \<times> (A *v y) = det A *\<^sub>R (A *v (x \<times> y))"
-proof%unimportant -
+proof -
   have "mat 1 = transpose (A ** transpose A)"
     by (metis (no_types) assms orthogonal_matrix_def transpose_mat)
   then show ?thesis
@@ -191,10 +191,10 @@
 lemma  cross_rotoinversion_matrix: "rotoinversion_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) = - A *v (x \<times> y)"
   by (simp add: rotoinversion_matrix_def cross_orthogonal_matrix scaleR_matrix_vector_assoc)
 
-lemma%important  cross_orthogonal_transformation:
+lemma  cross_orthogonal_transformation:
   assumes "orthogonal_transformation f"
   shows   "(f x) \<times> (f y) = det(matrix f) *\<^sub>R f(x \<times> y)"
-proof%unimportant -
+proof -
   have orth: "orthogonal_matrix (matrix f)"
     using assms orthogonal_transformation_matrix by blast
   have "matrix f *v z = f z" for z
@@ -208,7 +208,7 @@
            \<Longrightarrow> (f x) \<times> (f y) = f(x \<times> y)"
   by (simp add: cross_orthogonal_transformation orthogonal_transformation)
 
-subsection%unimportant \<open>Continuity\<close>
+subsection%important \<open>Continuity\<close>
 
 lemma  continuous_cross: "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. (f x) \<times> (g x))"
   apply (subst continuous_componentwise)
--- a/src/HOL/Analysis/Extended_Real_Limits.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Extended_Real_Limits.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -16,23 +16,23 @@
   "HOL-Library.Indicator_Function"
 begin
 
-lemma%important compact_UNIV:
+lemma compact_UNIV:
   "compact (UNIV :: 'a::{complete_linorder,linorder_topology,second_countable_topology} set)"
-  using%unimportant compact_complete_linorder
+  using compact_complete_linorder
   by (auto simp: seq_compact_eq_compact[symmetric] seq_compact_def)
 
-lemma%important compact_eq_closed:
+lemma compact_eq_closed:
   fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set"
   shows "compact S \<longleftrightarrow> closed S"
   using%unimportant closed_Int_compact[of S, OF _ compact_UNIV] compact_imp_closed
   by auto
 
-lemma%important closed_contains_Sup_cl:
+lemma closed_contains_Sup_cl:
   fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set"
   assumes "closed S"
     and "S \<noteq> {}"
   shows "Sup S \<in> S"
-proof%unimportant -
+proof -
   from compact_eq_closed[of S] compact_attains_sup[of S] assms
   obtain s where S: "s \<in> S" "\<forall>t\<in>S. t \<le> s"
     by auto
@@ -42,12 +42,12 @@
     by simp
 qed
 
-lemma%important closed_contains_Inf_cl:
+lemma closed_contains_Inf_cl:
   fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set"
   assumes "closed S"
     and "S \<noteq> {}"
   shows "Inf S \<in> S"
-proof%unimportant -
+proof -
   from compact_eq_closed[of S] compact_attains_inf[of S] assms
   obtain s where S: "s \<in> S" "\<forall>t\<in>S. s \<le> t"
     by auto
@@ -57,7 +57,7 @@
     by simp
 qed
 
-instance enat :: second_countable_topology
+instance%unimportant enat :: second_countable_topology
 proof
   show "\<exists>B::enat set set. countable B \<and> open = generate_topology B"
   proof (intro exI conjI)
@@ -66,8 +66,8 @@
   qed (simp add: open_enat_def)
 qed
 
-instance%important ereal :: second_countable_topology
-proof%unimportant (standard, intro exI conjI)
+instance%unimportant ereal :: second_countable_topology
+proof (standard, intro exI conjI)
   let ?B = "(\<Union>r\<in>\<rat>. {{..< r}, {r <..}} :: ereal set set)"
   show "countable ?B"
     by (auto intro: countable_rat)
@@ -98,8 +98,8 @@
 
 text \<open>This is a copy from \<open>ereal :: second_countable_topology\<close>. Maybe find a common super class of
 topological spaces where the rational numbers are densely embedded ?\<close>
-instance%important ennreal :: second_countable_topology
-proof%unimportant (standard, intro exI conjI)
+instance ennreal :: second_countable_topology
+proof (standard, intro exI conjI)
   let ?B = "(\<Union>r\<in>\<rat>. {{..< r}, {r <..}} :: ennreal set set)"
   show "countable ?B"
     by (auto intro: countable_rat)
@@ -128,13 +128,13 @@
   qed
 qed
 
-lemma%important ereal_open_closed_aux:
+lemma ereal_open_closed_aux:
   fixes S :: "ereal set"
   assumes "open S"
     and "closed S"
     and S: "(-\<infinity>) \<notin> S"
   shows "S = {}"
-proof%unimportant (rule ccontr)
+proof (rule ccontr)
   assume "\<not> ?thesis"
   then have *: "Inf S \<in> S"
 
@@ -172,10 +172,10 @@
     by auto
 qed
 
-lemma%important ereal_open_closed:
+lemma ereal_open_closed:
   fixes S :: "ereal set"
   shows "open S \<and> closed S \<longleftrightarrow> S = {} \<or> S = UNIV"
-proof%unimportant -
+proof -
   {
     assume lhs: "open S \<and> closed S"
     {
@@ -196,10 +196,10 @@
     by auto
 qed
 
-lemma%important ereal_open_atLeast:
+lemma ereal_open_atLeast:
   fixes x :: ereal
   shows "open {x..} \<longleftrightarrow> x = -\<infinity>"
-proof%unimportant
+proof
   assume "x = -\<infinity>"
   then have "{x..} = UNIV"
     by auto
@@ -215,12 +215,12 @@
     by (simp add: bot_ereal_def atLeast_eq_UNIV_iff)
 qed
 
-lemma%important mono_closed_real:
+lemma mono_closed_real:
   fixes S :: "real set"
   assumes mono: "\<forall>y z. y \<in> S \<and> y \<le> z \<longrightarrow> z \<in> S"
     and "closed S"
   shows "S = {} \<or> S = UNIV \<or> (\<exists>a. S = {a..})"
-proof%unimportant -
+proof -
   {
     assume "S \<noteq> {}"
     { assume ex: "\<exists>B. \<forall>x\<in>S. B \<le> x"
@@ -261,12 +261,12 @@
     by blast
 qed
 
-lemma%important mono_closed_ereal:
+lemma mono_closed_ereal:
   fixes S :: "real set"
   assumes mono: "\<forall>y z. y \<in> S \<and> y \<le> z \<longrightarrow> z \<in> S"
     and "closed S"
   shows "\<exists>a. S = {x. a \<le> ereal x}"
-proof%unimportant -
+proof -
   {
     assume "S = {}"
     then have ?thesis
@@ -296,11 +296,11 @@
     using mono_closed_real[of S] assms by auto
 qed
 
-lemma%important Liminf_within:
+lemma Liminf_within:
   fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
   shows "Liminf (at x within S) f = (SUP e\<in>{0<..}. INF y\<in>(S \<inter> ball x e - {x}). f y)"
   unfolding Liminf_def eventually_at
-proof%unimportant (rule SUP_eq, simp_all add: Ball_def Bex_def, safe)
+proof (rule SUP_eq, simp_all add: Ball_def Bex_def, safe)
   fix P d
   assume "0 < d" and "\<forall>y. y \<in> S \<longrightarrow> y \<noteq> x \<and> dist y x < d \<longrightarrow> P y"
   then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
@@ -316,11 +316,11 @@
        (auto intro!: INF_mono exI[of _ d] simp: dist_commute)
 qed
 
-lemma%important Limsup_within:
+lemma Limsup_within:
   fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
   shows "Limsup (at x within S) f = (INF e\<in>{0<..}. SUP y\<in>(S \<inter> ball x e - {x}). f y)"
   unfolding Limsup_def eventually_at
-proof%unimportant (rule INF_eq, simp_all add: Ball_def Bex_def, safe)
+proof (rule INF_eq, simp_all add: Ball_def Bex_def, safe)
   fix P d
   assume "0 < d" and "\<forall>y. y \<in> S \<longrightarrow> y \<noteq> x \<and> dist y x < d \<longrightarrow> P y"
   then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
@@ -357,7 +357,7 @@
   done
 
 
-subsection%important \<open>Extended-Real.thy\<close> (*FIX title *)
+subsection%important \<open>Extended-Real.thy\<close> (*FIX ME change title *)
 
 lemma sum_constant_ereal:
   fixes a::ereal
@@ -377,10 +377,10 @@
   ultimately show ?thesis by (metis (mono_tags, lifting) eventually_mono)
 qed
 
-lemma%important ereal_Inf_cmult:
+lemma ereal_Inf_cmult:
   assumes "c>(0::real)"
   shows "Inf {ereal c * x |x. P x} = ereal c * Inf {x. P x}"
-proof%unimportant -
+proof -
   have "(\<lambda>x::ereal. c * x) (Inf {x::ereal. P x}) = Inf ((\<lambda>x::ereal. c * x)`{x::ereal. P x})"
     apply (rule mono_bij_Inf)
     apply (simp add: assms ereal_mult_left_mono less_imp_le mono_def)
@@ -399,12 +399,12 @@
 It is much more convenient in many situations, see for instance the proof of
 \<open>tendsto_sum_ereal\<close> below.\<close>
 
-lemma%important tendsto_add_ereal_PInf:
+lemma tendsto_add_ereal_PInf:
   fixes y :: ereal
   assumes y: "y \<noteq> -\<infinity>"
   assumes f: "(f \<longlongrightarrow> \<infinity>) F" and g: "(g \<longlongrightarrow> y) F"
   shows "((\<lambda>x. f x + g x) \<longlongrightarrow> \<infinity>) F"
-proof%unimportant -
+proof -
   have "\<exists>C. eventually (\<lambda>x. g x > ereal C) F"
   proof (cases y)
     case (real r)
@@ -437,12 +437,12 @@
 that \<open>- (x + y)\<close> is in general different from \<open>(- x) + (- y)\<close> in ereal creates difficulties,
 so it is more efficient to copy the previous proof.\<close>
 
-lemma%important tendsto_add_ereal_MInf:
+lemma tendsto_add_ereal_MInf:
   fixes y :: ereal
   assumes y: "y \<noteq> \<infinity>"
   assumes f: "(f \<longlongrightarrow> -\<infinity>) F" and g: "(g \<longlongrightarrow> y) F"
   shows "((\<lambda>x. f x + g x) \<longlongrightarrow> -\<infinity>) F"
-proof%unimportant -
+proof -
   have "\<exists>C. eventually (\<lambda>x. g x < ereal C) F"
   proof (cases y)
     case (real r)
@@ -470,12 +470,12 @@
   then show ?thesis by (simp add: tendsto_MInfty)
 qed
 
-lemma%important tendsto_add_ereal_general1:
+lemma tendsto_add_ereal_general1:
   fixes x y :: ereal
   assumes y: "\<bar>y\<bar> \<noteq> \<infinity>"
   assumes f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
   shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
-proof%unimportant (cases x)
+proof (cases x)
   case (real r)
   have a: "\<bar>x\<bar> \<noteq> \<infinity>" by (simp add: real)
   show ?thesis by (rule tendsto_add_ereal[OF a, OF y, OF f, OF g])
@@ -488,12 +488,12 @@
     by (metis abs_ereal.simps(3) ereal_MInfty_eq_plus)
 qed
 
-lemma%important tendsto_add_ereal_general2:
+lemma tendsto_add_ereal_general2:
   fixes x y :: ereal
   assumes x: "\<bar>x\<bar> \<noteq> \<infinity>"
       and f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
   shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
-proof%unimportant -
+proof -
   have "((\<lambda>x. g x + f x) \<longlongrightarrow> x + y) F"
     using tendsto_add_ereal_general1[OF x, OF g, OF f] add.commute[of "y", of "x"] by simp
   moreover have "\<And>x. g x + f x = f x + g x" using add.commute by auto
@@ -503,12 +503,12 @@
 text \<open>The next lemma says that the addition is continuous on \<open>ereal\<close>, except at
 the pairs \<open>(-\<infinity>, \<infinity>)\<close> and \<open>(\<infinity>, -\<infinity>)\<close>.\<close>
 
-lemma%important tendsto_add_ereal_general [tendsto_intros]:
+lemma tendsto_add_ereal_general [tendsto_intros]:
   fixes x y :: ereal
   assumes "\<not>((x=\<infinity> \<and> y=-\<infinity>) \<or> (x=-\<infinity> \<and> y=\<infinity>))"
       and f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
   shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
-proof%unimportant (cases x)
+proof (cases x)
   case (real r)
   show ?thesis
     apply (rule tendsto_add_ereal_general2) using real assms by auto
@@ -528,10 +528,10 @@
 ereal times ereal, except at \<open>(\<infinity>, 0)\<close> and \<open>(-\<infinity>, 0)\<close> and \<open>(0, \<infinity>)\<close> and \<open>(0, -\<infinity>)\<close>,
 starting with specific situations.\<close>
 
-lemma%important tendsto_mult_real_ereal:
+lemma tendsto_mult_real_ereal:
   assumes "(u \<longlongrightarrow> ereal l) F" "(v \<longlongrightarrow> ereal m) F"
   shows "((\<lambda>n. u n * v n) \<longlongrightarrow> ereal l * ereal m) F"
-proof%unimportant -
+proof -
   have ureal: "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) F" by (rule real_lim_then_eventually_real[OF assms(1)])
   then have "((\<lambda>n. ereal(real_of_ereal(u n))) \<longlongrightarrow> ereal l) F" using assms by auto
   then have limu: "((\<lambda>n. real_of_ereal(u n)) \<longlongrightarrow> l) F" by auto
@@ -551,11 +551,11 @@
   then show ?thesis using * filterlim_cong by fastforce
 qed
 
-lemma%important tendsto_mult_ereal_PInf:
+lemma tendsto_mult_ereal_PInf:
   fixes f g::"_ \<Rightarrow> ereal"
   assumes "(f \<longlongrightarrow> l) F" "l>0" "(g \<longlongrightarrow> \<infinity>) F"
   shows "((\<lambda>x. f x * g x) \<longlongrightarrow> \<infinity>) F"
-proof%unimportant -
+proof -
   obtain a::real where "0 < ereal a" "a < l" using assms(2) using ereal_dense2 by blast
   have *: "eventually (\<lambda>x. f x > a) F" using \<open>a < l\<close> assms(1) by (simp add: order_tendsto_iff)
   {
@@ -579,11 +579,11 @@
   then show ?thesis by (auto simp add: tendsto_PInfty)
 qed
 
-lemma%important tendsto_mult_ereal_pos:
+lemma tendsto_mult_ereal_pos:
   fixes f g::"_ \<Rightarrow> ereal"
   assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "l>0" "m>0"
   shows "((\<lambda>x. f x * g x) \<longlongrightarrow> l * m) F"
-proof%unimportant (cases)
+proof (cases)
   assume *: "l = \<infinity> \<or> m = \<infinity>"
   then show ?thesis
   proof (cases)
@@ -618,11 +618,11 @@
   shows "sgn(l) * sgn(l) = 1"
 apply (cases l) using assms by (auto simp add: one_ereal_def sgn_if)
 
-lemma%important tendsto_mult_ereal [tendsto_intros]:
+lemma tendsto_mult_ereal [tendsto_intros]:
   fixes f g::"_ \<Rightarrow> ereal"
   assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "\<not>((l=0 \<and> abs(m) = \<infinity>) \<or> (m=0 \<and> abs(l) = \<infinity>))"
   shows "((\<lambda>x. f x * g x) \<longlongrightarrow> l * m) F"
-proof%unimportant (cases)
+proof (cases)
   assume "l=0 \<or> m=0"
   then have "abs(l) \<noteq> \<infinity>" "abs(m) \<noteq> \<infinity>" using assms(3) by auto
   then obtain lr mr where "l = ereal lr" "m = ereal mr" by auto
@@ -661,11 +661,11 @@
 
 subsubsection%important \<open>Continuity of division\<close>
 
-lemma%important tendsto_inverse_ereal_PInf:
+lemma tendsto_inverse_ereal_PInf:
   fixes u::"_ \<Rightarrow> ereal"
   assumes "(u \<longlongrightarrow> \<infinity>) F"
   shows "((\<lambda>x. 1/ u x) \<longlongrightarrow> 0) F"
-proof%unimportant -
+proof -
   {
     fix e::real assume "e>0"
     have "1/e < \<infinity>" by auto
@@ -702,11 +702,11 @@
   shows "(u \<longlongrightarrow> l) F \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> ((\<lambda>x. 1/ u x) \<longlongrightarrow> 1/l) F"
   using tendsto_inverse unfolding inverse_eq_divide .
 
-lemma%important tendsto_inverse_ereal [tendsto_intros]:
+lemma tendsto_inverse_ereal [tendsto_intros]:
   fixes u::"_ \<Rightarrow> ereal"
   assumes "(u \<longlongrightarrow> l) F" "l \<noteq> 0"
   shows "((\<lambda>x. 1/ u x) \<longlongrightarrow> 1/l) F"
-proof%unimportant (cases l)
+proof (cases l)
   case (real r)
   then have "r \<noteq> 0" using assms(2) by auto
   then have "1/l = ereal(1/r)" using real by (simp add: one_ereal_def)
@@ -747,11 +747,11 @@
   then show ?thesis unfolding v_def using Lim_transform_eventually[OF *] \<open> 1/l = 0 \<close> by auto
 qed
 
-lemma%important tendsto_divide_ereal [tendsto_intros]:
+lemma tendsto_divide_ereal [tendsto_intros]:
   fixes f g::"_ \<Rightarrow> ereal"
   assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "m \<noteq> 0" "\<not>(abs(l) = \<infinity> \<and> abs(m) = \<infinity>)"
   shows "((\<lambda>x. f x / g x) \<longlongrightarrow> l / m) F"
-proof%unimportant -
+proof -
   define h where "h = (\<lambda>x. 1/ g x)"
   have *: "(h \<longlongrightarrow> 1/m) F" unfolding h_def using assms(2) assms(3) tendsto_inverse_ereal by auto
   have "((\<lambda>x. f x * h x) \<longlongrightarrow> l * (1/m)) F"
@@ -766,11 +766,11 @@
 
 text \<open>The assumptions of @{thm tendsto_diff_ereal} are too strong, we weaken them here.\<close>
 
-lemma%important tendsto_diff_ereal_general [tendsto_intros]:
+lemma tendsto_diff_ereal_general [tendsto_intros]:
   fixes u v::"'a \<Rightarrow> ereal"
   assumes "(u \<longlongrightarrow> l) F" "(v \<longlongrightarrow> m) F" "\<not>((l = \<infinity> \<and> m = \<infinity>) \<or> (l = -\<infinity> \<and> m = -\<infinity>))"
   shows "((\<lambda>n. u n - v n) \<longlongrightarrow> l - m) F"
-proof%unimportant -
+proof -
   have "((\<lambda>n. u n + (-v n)) \<longlongrightarrow> l + (-m)) F"
     apply (intro tendsto_intros assms) using assms by (auto simp add: ereal_uminus_eq_reorder)
   then show ?thesis by (simp add: minus_ereal_def)
@@ -780,11 +780,11 @@
   "(\<lambda> n::nat. real n) \<longlonglongrightarrow> \<infinity>"
 by (simp add: filterlim_real_sequentially tendsto_PInfty_eq_at_top)
 
-lemma%important tendsto_at_top_pseudo_inverse [tendsto_intros]:
+lemma tendsto_at_top_pseudo_inverse [tendsto_intros]:
   fixes u::"nat \<Rightarrow> nat"
   assumes "LIM n sequentially. u n :> at_top"
   shows "LIM n sequentially. Inf {N. u N \<ge> n} :> at_top"
-proof%unimportant -
+proof -
   {
     fix C::nat
     define M where "M = Max {u n| n. n \<le> C}+1"
@@ -811,11 +811,11 @@
   then show ?thesis using filterlim_at_top by auto
 qed
 
-lemma%important pseudo_inverse_finite_set:
+lemma pseudo_inverse_finite_set:
   fixes u::"nat \<Rightarrow> nat"
   assumes "LIM n sequentially. u n :> at_top"
   shows "finite {N. u N \<le> n}"
-proof%unimportant -
+proof -
   fix n
   have "eventually (\<lambda>N. u N \<ge> n+1) sequentially" using assms
     by (simp add: filterlim_at_top)
@@ -860,11 +860,11 @@
   then show ?thesis by auto
 qed
 
-lemma%important ereal_truncation_real_top [tendsto_intros]:
+lemma ereal_truncation_real_top [tendsto_intros]:
   fixes x::ereal
   assumes "x \<noteq> - \<infinity>"
   shows "(\<lambda>n::nat. real_of_ereal(min x n)) \<longlonglongrightarrow> x"
-proof%unimportant (cases x)
+proof (cases x)
   case (real r)
   then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
   then have "min x n = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
@@ -878,10 +878,10 @@
   then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
 qed (simp add: assms)
 
-lemma%important ereal_truncation_bottom [tendsto_intros]:
+lemma ereal_truncation_bottom [tendsto_intros]:
   fixes x::ereal
   shows "(\<lambda>n::nat. max x (- real n)) \<longlonglongrightarrow> x"
-proof%unimportant (cases x)
+proof (cases x)
   case (real r)
   then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
   then have "max x (-real n) = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
@@ -899,11 +899,11 @@
   then show ?thesis by auto
 qed
 
-lemma%important ereal_truncation_real_bottom [tendsto_intros]:
+lemma ereal_truncation_real_bottom [tendsto_intros]:
   fixes x::ereal
   assumes "x \<noteq> \<infinity>"
   shows "(\<lambda>n::nat. real_of_ereal(max x (- real n))) \<longlonglongrightarrow> x"
-proof%unimportant (cases x)
+proof (cases x)
   case (real r)
   then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
   then have "max x (-real n) = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
@@ -931,9 +931,9 @@
 qed(simp)
 
 
-lemma%important continuous_ereal_abs:
+lemma continuous_ereal_abs:
   "continuous_on (UNIV::ereal set) abs"
-proof%unimportant -
+proof -
   have "continuous_on ({..0} \<union> {(0::ereal)..}) abs"
     apply (rule continuous_on_closed_Un, auto)
     apply (rule iffD1[OF continuous_on_cong, of "{..0}" _ "\<lambda>x. -x"])
@@ -970,11 +970,11 @@
   then show ?thesis by auto
 qed
 
-lemma%important tendsto_mult_ennreal [tendsto_intros]:
+lemma tendsto_mult_ennreal [tendsto_intros]:
   fixes l m::ennreal
   assumes "(u \<longlongrightarrow> l) F" "(v \<longlongrightarrow> m) F" "\<not>((l = 0 \<and> m = \<infinity>) \<or> (l = \<infinity> \<and> m = 0))"
   shows "((\<lambda>n. u n * v n) \<longlongrightarrow> l * m) F"
-proof%unimportant -
+proof -
   have "((\<lambda>n. e2ennreal(enn2ereal (u n) * enn2ereal (v n))) \<longlongrightarrow> e2ennreal(enn2ereal l * enn2ereal m)) F"
     apply (intro tendsto_intros) using assms apply auto
     using enn2ereal_inject zero_ennreal.rep_eq by fastforce+
@@ -987,9 +987,9 @@
 qed
 
 
-subsection%important \<open>monoset\<close>
+subsection%important \<open>monoset\<close> (*FIX ME title *)
 
-definition%important (in order) mono_set:
+definition (in order) mono_set:
   "mono_set S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
 
 lemma (in order) mono_greaterThan [intro, simp]: "mono_set {B<..}" unfolding mono_set by auto
@@ -997,11 +997,11 @@
 lemma (in order) mono_UNIV [intro, simp]: "mono_set UNIV" unfolding mono_set by auto
 lemma (in order) mono_empty [intro, simp]: "mono_set {}" unfolding mono_set by auto
 
-lemma%important (in complete_linorder) mono_set_iff:
+lemma (in complete_linorder) mono_set_iff:
   fixes S :: "'a set"
   defines "a \<equiv> Inf S"
   shows "mono_set S \<longleftrightarrow> S = {a <..} \<or> S = {a..}" (is "_ = ?c")
-proof%unimportant
+proof
   assume "mono_set S"
   then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S"
     by (auto simp: mono_set)
@@ -1043,12 +1043,12 @@
   by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_ereal_atLeast
     ereal_open_closed mono_empty mono_set_iff open_ereal_greaterThan)
 
-lemma%important ereal_Liminf_Sup_monoset:
+lemma ereal_Liminf_Sup_monoset:
   fixes f :: "'a \<Rightarrow> ereal"
   shows "Liminf net f =
     Sup {l. \<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
     (is "_ = Sup ?A")
-proof%unimportant (safe intro!: Liminf_eqI complete_lattice_class.Sup_upper complete_lattice_class.Sup_least)
+proof (safe intro!: Liminf_eqI complete_lattice_class.Sup_upper complete_lattice_class.Sup_least)
   fix P
   assume P: "eventually P net"
   fix S
@@ -1082,12 +1082,12 @@
   qed
 qed
 
-lemma%important ereal_Limsup_Inf_monoset:
+lemma ereal_Limsup_Inf_monoset:
   fixes f :: "'a \<Rightarrow> ereal"
   shows "Limsup net f =
     Inf {l. \<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
     (is "_ = Inf ?A")
-proof%unimportant (safe intro!: Limsup_eqI complete_lattice_class.Inf_lower complete_lattice_class.Inf_greatest)
+proof (safe intro!: Limsup_eqI complete_lattice_class.Inf_lower complete_lattice_class.Inf_greatest)
   fix P
   assume P: "eventually P net"
   fix S
@@ -1121,11 +1121,11 @@
   qed
 qed
 
-lemma%important liminf_bounded_open:
+lemma liminf_bounded_open:
   fixes x :: "nat \<Rightarrow> ereal"
   shows "x0 \<le> liminf x \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> x0 \<in> S \<longrightarrow> (\<exists>N. \<forall>n\<ge>N. x n \<in> S))"
   (is "_ \<longleftrightarrow> ?P x0")
-proof%unimportant
+proof
   assume "?P x0"
   then show "x0 \<le> liminf x"
     unfolding ereal_Liminf_Sup_monoset eventually_sequentially
@@ -1159,11 +1159,11 @@
     by auto
 qed
 
-lemma%important limsup_finite_then_bounded:
+lemma limsup_finite_then_bounded:
   fixes u::"nat \<Rightarrow> real"
   assumes "limsup u < \<infinity>"
   shows "\<exists>C. \<forall>n. u n \<le> C"
-proof%unimportant -
+proof -
   obtain C where C: "limsup u < C" "C < \<infinity>" using assms ereal_dense2 by blast
   then have "C = ereal(real_of_ereal C)" using ereal_real by force
   have "eventually (\<lambda>n. u n < C) sequentially" using C(1) unfolding Limsup_def
@@ -1273,11 +1273,11 @@
   then show ?case using Suc.IH by simp
 qed (auto)
 
-lemma%important Limsup_obtain:
+lemma Limsup_obtain:
   fixes u::"_ \<Rightarrow> 'a :: complete_linorder"
   assumes "Limsup F u > c"
   shows "\<exists>i. u i > c"
-proof%unimportant -
+proof -
   have "(INF P\<in>{P. eventually P F}. SUP x\<in>{x. P x}. u x) > c" using assms by (simp add: Limsup_def)
   then show ?thesis by (metis eventually_True mem_Collect_eq less_INF_D less_SUP_iff)
 qed
@@ -1285,10 +1285,10 @@
 text \<open>The next lemma is extremely useful, as it often makes it possible to reduce statements
 about limsups to statements about limits.\<close>
 
-lemma%important limsup_subseq_lim:
+lemma limsup_subseq_lim:
   fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
   shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (u o r) \<longlonglongrightarrow> limsup u"
-proof%unimportant (cases)
+proof (cases)
   assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. u m \<le> u p"
   then have "\<exists>r. \<forall>n. (\<forall>m\<ge>r n. u m \<le> u (r n)) \<and> r n < r (Suc n)"
     by (intro dependent_nat_choice) (auto simp: conj_commute)
@@ -1378,10 +1378,10 @@
   then show ?thesis using \<open>strict_mono r\<close> by auto
 qed
 
-lemma%important liminf_subseq_lim:
+lemma liminf_subseq_lim:
   fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
   shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (u o r) \<longlonglongrightarrow> liminf u"
-proof%unimportant (cases)
+proof (cases)
   assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. u m \<ge> u p"
   then have "\<exists>r. \<forall>n. (\<forall>m\<ge>r n. u m \<ge> u (r n)) \<and> r n < r (Suc n)"
     by (intro dependent_nat_choice) (auto simp: conj_commute)
@@ -1476,10 +1476,10 @@
 subsequences thanks to \<open>limsup_subseq_lim\<close>. The statement for limits follows for instance from
 \<open>tendsto_add_ereal_general\<close>.\<close>
 
-lemma%important ereal_limsup_add_mono:
+lemma ereal_limsup_add_mono:
   fixes u v::"nat \<Rightarrow> ereal"
   shows "limsup (\<lambda>n. u n + v n) \<le> limsup u + limsup v"
-proof%unimportant (cases)
+proof (cases)
   assume "(limsup u = \<infinity>) \<or> (limsup v = \<infinity>)"
   then have "limsup u + limsup v = \<infinity>" by simp
   then show ?thesis by auto
@@ -1522,11 +1522,11 @@
 This explains why there are more assumptions in the next lemma dealing with liminfs that in the
 previous one about limsups.\<close>
 
-lemma%important ereal_liminf_add_mono:
+lemma ereal_liminf_add_mono:
   fixes u v::"nat \<Rightarrow> ereal"
   assumes "\<not>((liminf u = \<infinity> \<and> liminf v = -\<infinity>) \<or> (liminf u = -\<infinity> \<and> liminf v = \<infinity>))"
   shows "liminf (\<lambda>n. u n + v n) \<ge> liminf u + liminf v"
-proof%unimportant (cases)
+proof (cases)
   assume "(liminf u = -\<infinity>) \<or> (liminf v = -\<infinity>)"
   then have *: "liminf u + liminf v = -\<infinity>" using assms by auto
   show ?thesis by (simp add: *)
@@ -1565,11 +1565,11 @@
   then show ?thesis unfolding w_def by simp
 qed
 
-lemma%important ereal_limsup_lim_add:
+lemma ereal_limsup_lim_add:
   fixes u v::"nat \<Rightarrow> ereal"
   assumes "u \<longlonglongrightarrow> a" "abs(a) \<noteq> \<infinity>"
   shows "limsup (\<lambda>n. u n + v n) = a + limsup v"
-proof%unimportant -
+proof -
   have "limsup u = a" using assms(1) using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
   have "(\<lambda>n. -u n) \<longlonglongrightarrow> -a" using assms(1) by auto
   then have "limsup (\<lambda>n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
@@ -1596,11 +1596,11 @@
   then show ?thesis using up by simp
 qed
 
-lemma%important ereal_limsup_lim_mult:
+lemma ereal_limsup_lim_mult:
   fixes u v::"nat \<Rightarrow> ereal"
   assumes "u \<longlonglongrightarrow> a" "a>0" "a \<noteq> \<infinity>"
   shows "limsup (\<lambda>n. u n * v n) = a * limsup v"
-proof%unimportant -
+proof -
   define w where "w = (\<lambda>n. u n * v n)"
   obtain r where r: "strict_mono r" "(v o r) \<longlonglongrightarrow> limsup v" using limsup_subseq_lim by auto
   have "(u o r) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto
@@ -1626,11 +1626,11 @@
   then show ?thesis using I unfolding w_def by auto
 qed
 
-lemma%important ereal_liminf_lim_mult:
+lemma ereal_liminf_lim_mult:
   fixes u v::"nat \<Rightarrow> ereal"
   assumes "u \<longlonglongrightarrow> a" "a>0" "a \<noteq> \<infinity>"
   shows "liminf (\<lambda>n. u n * v n) = a * liminf v"
-proof%unimportant -
+proof -
   define w where "w = (\<lambda>n. u n * v n)"
   obtain r where r: "strict_mono r" "(v o r) \<longlonglongrightarrow> liminf v" using liminf_subseq_lim by auto
   have "(u o r) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto
@@ -1656,11 +1656,11 @@
   then show ?thesis using I unfolding w_def by auto
 qed
 
-lemma%important ereal_liminf_lim_add:
+lemma ereal_liminf_lim_add:
   fixes u v::"nat \<Rightarrow> ereal"
   assumes "u \<longlonglongrightarrow> a" "abs(a) \<noteq> \<infinity>"
   shows "liminf (\<lambda>n. u n + v n) = a + liminf v"
-proof%unimportant -
+proof -
   have "liminf u = a" using assms(1) tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
   then have *: "abs(liminf u) \<noteq> \<infinity>" using assms(2) by auto
   have "(\<lambda>n. -u n) \<longlonglongrightarrow> -a" using assms(1) by auto
@@ -1689,10 +1689,10 @@
   then show ?thesis using up by simp
 qed
 
-lemma%important ereal_liminf_limsup_add:
+lemma ereal_liminf_limsup_add:
   fixes u v::"nat \<Rightarrow> ereal"
   shows "liminf (\<lambda>n. u n + v n) \<le> liminf u + limsup v"
-proof%unimportant (cases)
+proof (cases)
   assume "limsup v = \<infinity> \<or> liminf u = \<infinity>"
   then show ?thesis by auto
 next
@@ -1741,12 +1741,12 @@
   done
 
 
-lemma%important liminf_minus_ennreal:
+lemma liminf_minus_ennreal:
   fixes u v::"nat \<Rightarrow> ennreal"
   shows "(\<And>n. v n \<le> u n) \<Longrightarrow> liminf (\<lambda>n. u n - v n) \<le> limsup u - limsup v"
   unfolding liminf_SUP_INF limsup_INF_SUP
   including ennreal.lifting
-proof%unimportant (transfer, clarsimp)
+proof (transfer, clarsimp)
   fix v u :: "nat \<Rightarrow> ereal" assume *: "\<forall>x. 0 \<le> v x" "\<forall>x. 0 \<le> u x" "\<And>n. v n \<le> u n"
   moreover have "0 \<le> limsup u - limsup v"
     using * by (intro ereal_diff_positive Limsup_mono always_eventually) simp
@@ -1759,7 +1759,7 @@
     by (auto simp: * ereal_diff_positive max.absorb2 liminf_SUP_INF[symmetric] limsup_INF_SUP[symmetric] ereal_liminf_limsup_minus)
 qed
 
-subsection%unimportant "Relate extended reals and the indicator function"
+subsection%important "Relate extended reals and the indicator function"
 
 lemma ereal_indicator_le_0: "(indicator S x::ereal) \<le> 0 \<longleftrightarrow> x \<notin> S"
   by (auto split: split_indicator simp: one_ereal_def)
--- a/src/HOL/Analysis/Fashoda_Theorem.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Fashoda_Theorem.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -30,14 +30,14 @@
   apply (rule, rule continuous_interval_bij)
   done
 
-lemma%important in_interval_interval_bij:
+lemma in_interval_interval_bij:
   fixes a b u v x :: "'a::euclidean_space"
   assumes "x \<in> cbox a b"
     and "cbox u v \<noteq> {}"
   shows "interval_bij (a, b) (u, v) x \<in> cbox u v"
   apply (simp only: interval_bij_def split_conv mem_box inner_sum_left_Basis cong: ball_cong)
   apply safe
-proof%unimportant -
+proof -
   fix i :: 'a
   assume i: "i \<in> Basis"
   have "cbox a b \<noteq> {}"
@@ -89,7 +89,7 @@
   shows "\<bar>x$1\<bar> \<le> 1" and "\<bar>x$2\<bar> \<le> 1"
   using assms unfolding infnorm_eq_1_2 by auto
 
-lemma%important fashoda_unit:
+proposition fashoda_unit:
   fixes f g :: "real \<Rightarrow> real^2"
   assumes "f ` {-1 .. 1} \<subseteq> cbox (-1) 1"
     and "g ` {-1 .. 1} \<subseteq> cbox (-1) 1"
@@ -99,7 +99,7 @@
     and "f 1$1 = 1" "g (- 1) $2 = -1"
     and "g 1 $2 = 1"
   shows "\<exists>s\<in>{-1 .. 1}. \<exists>t\<in>{-1 .. 1}. f s = g t"
-proof%unimportant (rule ccontr)
+proof (rule ccontr)
   assume "\<not> ?thesis"
   note as = this[unfolded bex_simps,rule_format]
   define sqprojection
@@ -256,7 +256,7 @@
   qed 
 qed
 
-lemma%important fashoda_unit_path:
+proposition fashoda_unit_path:
   fixes f g :: "real \<Rightarrow> real^2"
   assumes "path f"
     and "path g"
@@ -267,7 +267,7 @@
     and "(pathstart g)$2 = -1"
     and "(pathfinish g)$2 = 1"
   obtains z where "z \<in> path_image f" and "z \<in> path_image g"
-proof%unimportant -
+proof -
   note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def]
   define iscale where [abs_def]: "iscale z = inverse 2 *\<^sub>R (z + 1)" for z :: real
   have isc: "iscale ` {- 1..1} \<subseteq> {0..1}"
@@ -312,7 +312,7 @@
     done
 qed
 
-lemma%important fashoda:
+theorem fashoda:
   fixes b :: "real^2"
   assumes "path f"
     and "path g"
@@ -323,7 +323,7 @@
     and "(pathstart g)$2 = a$2"
     and "(pathfinish g)$2 = b$2"
   obtains z where "z \<in> path_image f" and "z \<in> path_image g"
-proof%unimportant -
+proof -
   fix P Q S
   presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" and "Q \<Longrightarrow> thesis" and "S \<Longrightarrow> thesis"
   then show thesis
@@ -632,7 +632,7 @@
 
 subsection%important \<open>Useful Fashoda corollary pointed out to me by Tom Hales\<close>(*FIXME change title? *)
 
-lemma%important fashoda_interlace:
+corollary fashoda_interlace:
   fixes a :: "real^2"
   assumes "path f"
     and "path g"
@@ -646,7 +646,7 @@
     and "(pathstart g)$1 < (pathfinish f)$1"
     and "(pathfinish f)$1 < (pathfinish g)$1"
   obtains z where "z \<in> path_image f" and "z \<in> path_image g"
-proof%unimportant -
+proof -
   have "cbox a b \<noteq> {}"
     using path_image_nonempty[of f] using assms(3) by auto
   note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less]
--- a/src/HOL/Analysis/Finite_Cartesian_Product.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -626,7 +626,7 @@
 lemma%unimportant norm_le_l1_cart: "norm x \<le> sum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
   by (simp add: norm_vec_def L2_set_le_sum)
 
-lemma%unimportant bounded_linear_vec_nth[intro]: "bounded_linear (\<lambda>x. x $ i)"
+lemma%unimportant bounded_linear_vec_nth: "bounded_linear (\<lambda>x. x $ i)"
 apply standard
 apply (rule vector_add_component)
 apply (rule vector_scaleR_component)
--- a/src/HOL/Analysis/Finite_Product_Measure.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Finite_Product_Measure.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -8,14 +8,14 @@
 imports Binary_Product_Measure
 begin
 
-lemma%unimportant PiE_choice: "(\<exists>f\<in>Pi\<^sub>E I F. \<forall>i\<in>I. P i (f i)) \<longleftrightarrow> (\<forall>i\<in>I. \<exists>x\<in>F i. P i x)"
+lemma PiE_choice: "(\<exists>f\<in>Pi\<^sub>E I F. \<forall>i\<in>I. P i (f i)) \<longleftrightarrow> (\<forall>i\<in>I. \<exists>x\<in>F i. P i x)"
   by (auto simp: Bex_def PiE_iff Ball_def dest!: choice_iff'[THEN iffD1])
      (force intro: exI[of _ "restrict f I" for f])
 
-lemma%unimportant case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
+lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
   by auto
 
-subsubsection%unimportant \<open>More about Function restricted by \<^const>\<open>extensional\<close>\<close>
+subsubsection \<open>More about Function restricted by \<^const>\<open>extensional\<close>\<close>
 
 definition
   "merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
@@ -116,42 +116,42 @@
 definition%important prod_emb where
   "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (\<Pi>\<^sub>E i\<in>I. space (M i))"
 
-lemma%important prod_emb_iff:
+lemma prod_emb_iff:
   "f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))"
   unfolding%unimportant prod_emb_def PiE_def by auto
 
-lemma%unimportant
+lemma
   shows prod_emb_empty[simp]: "prod_emb M L K {} = {}"
     and prod_emb_Un[simp]: "prod_emb M L K (A \<union> B) = prod_emb M L K A \<union> prod_emb M L K B"
     and prod_emb_Int: "prod_emb M L K (A \<inter> B) = prod_emb M L K A \<inter> prod_emb M L K B"
     and prod_emb_UN[simp]: "prod_emb M L K (\<Union>i\<in>I. F i) = (\<Union>i\<in>I. prod_emb M L K (F i))"
     and prod_emb_INT[simp]: "I \<noteq> {} \<Longrightarrow> prod_emb M L K (\<Inter>i\<in>I. F i) = (\<Inter>i\<in>I. prod_emb M L K (F i))"
     and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B"
-  by%unimportant (auto simp: prod_emb_def)
+  by (auto simp: prod_emb_def)
 
-lemma%unimportant prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
+lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
     prod_emb I M J (\<Pi>\<^sub>E i\<in>J. E i) = (\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i))"
   by (force simp: prod_emb_def PiE_iff if_split_mem2)
 
-lemma%unimportant prod_emb_PiE_same_index[simp]:
+lemma prod_emb_PiE_same_index[simp]:
     "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E"
   by (auto simp: prod_emb_def PiE_iff)
 
-lemma%unimportant prod_emb_trans[simp]:
+lemma prod_emb_trans[simp]:
   "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X"
   by (auto simp add: Int_absorb1 prod_emb_def PiE_def)
 
-lemma%unimportant prod_emb_Pi:
+lemma prod_emb_Pi:
   assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
   shows "prod_emb K M J (Pi\<^sub>E J X) = (\<Pi>\<^sub>E i\<in>K. if i \<in> J then X i else space (M i))"
   using assms sets.space_closed
   by (auto simp: prod_emb_def PiE_iff split: if_split_asm) blast+
 
-lemma%unimportant prod_emb_id:
+lemma prod_emb_id:
   "B \<subseteq> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
   by (auto simp: prod_emb_def subset_eq extensional_restrict)
 
-lemma%unimportant prod_emb_mono:
+lemma prod_emb_mono:
   "F \<subseteq> G \<Longrightarrow> prod_emb A M B F \<subseteq> prod_emb A M B G"
   by (auto simp: prod_emb_def)
 
@@ -173,20 +173,20 @@
 translations
   "\<Pi>\<^sub>M x\<in>I. M" == "CONST PiM I (%x. M)"
 
-lemma%important extend_measure_cong:
+lemma extend_measure_cong:
   assumes "\<Omega> = \<Omega>'" "I = I'" "G = G'" "\<And>i. i \<in> I' \<Longrightarrow> \<mu> i = \<mu>' i"
   shows "extend_measure \<Omega> I G \<mu> = extend_measure \<Omega>' I' G' \<mu>'"
-  unfolding%unimportant extend_measure_def by (auto simp add: assms)
+  unfolding extend_measure_def by (auto simp add: assms)
 
-lemma%unimportant Pi_cong_sets:
+lemma Pi_cong_sets:
     "\<lbrakk>I = J; \<And>x. x \<in> I \<Longrightarrow> M x = N x\<rbrakk> \<Longrightarrow> Pi I M = Pi J N"
   unfolding Pi_def by auto
 
-lemma%important PiM_cong:
+lemma PiM_cong:
   assumes "I = J" "\<And>x. x \<in> I \<Longrightarrow> M x = N x"
   shows "PiM I M = PiM J N"
   unfolding PiM_def
-proof%unimportant (rule extend_measure_cong, goal_cases)
+proof (rule extend_measure_cong, goal_cases)
   case 1
   show ?case using assms
     by (subst assms(1), intro PiE_cong[of J "\<lambda>i. space (M i)" "\<lambda>i. space (N i)"]) simp_all
@@ -206,14 +206,14 @@
 qed
 
 
-lemma%unimportant prod_algebra_sets_into_space:
+lemma prod_algebra_sets_into_space:
   "prod_algebra I M \<subseteq> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
   by (auto simp: prod_emb_def prod_algebra_def)
 
-lemma%important prod_algebra_eq_finite:
+lemma prod_algebra_eq_finite:
   assumes I: "finite I"
   shows "prod_algebra I M = {(\<Pi>\<^sub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
-proof%unimportant (intro iffI set_eqI)
+proof (intro iffI set_eqI)
   fix A assume "A \<in> ?L"
   then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
     and A: "A = prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j)"
@@ -232,32 +232,32 @@
     by (auto simp: prod_algebra_def)
 qed
 
-lemma%unimportant prod_algebraI:
+lemma prod_algebraI:
   "finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i))
     \<Longrightarrow> prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j) \<in> prod_algebra I M"
   by (auto simp: prod_algebra_def)
 
-lemma%unimportant prod_algebraI_finite:
+lemma prod_algebraI_finite:
   "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>E I E) \<in> prod_algebra I M"
   using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp
 
-lemma%unimportant Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
+lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
 proof (safe intro!: Int_stableI)
   fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
   then show "\<exists>G. Pi\<^sub>E J E \<inter> Pi\<^sub>E J F = Pi\<^sub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
     by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"] simp: PiE_Int)
 qed
 
-lemma%unimportant prod_algebraE:
+lemma prod_algebraE:
   assumes A: "A \<in> prod_algebra I M"
   obtains J E where "A = prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j)"
     "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)"
   using A by (auto simp: prod_algebra_def)
 
-lemma%important prod_algebraE_all:
+lemma prod_algebraE_all:
   assumes A: "A \<in> prod_algebra I M"
   obtains E where "A = Pi\<^sub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
-proof%unimportant -
+proof -
   from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)"
     and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))"
     by (auto simp: prod_algebra_def)
@@ -270,7 +270,7 @@
   ultimately show ?thesis using that by auto
 qed
 
-lemma%unimportant Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
+lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
 proof (unfold Int_stable_def, safe)
   fix A assume "A \<in> prod_algebra I M"
   from prod_algebraE[OF this] guess J E . note A = this
@@ -291,11 +291,11 @@
   finally show "A \<inter> B \<in> prod_algebra I M" .
 qed
 
-lemma%unimportant prod_algebra_mono:
+proposition prod_algebra_mono:
   assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)"
   assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)"
   shows "prod_algebra I E \<subseteq> prod_algebra I F"
-proof%unimportant
+proof
   fix A assume "A \<in> prod_algebra I E"
   then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I"
     and A: "A = prod_emb I E J (\<Pi>\<^sub>E i\<in>J. G i)"
@@ -315,18 +315,17 @@
     apply auto
     done
 qed
-
-lemma%unimportant prod_algebra_cong:
+proposition prod_algebra_cong:
   assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
   shows "prod_algebra I M = prod_algebra J N"
-proof%unimportant -
+proof -
   have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)"
     using sets_eq_imp_space_eq[OF sets] by auto
   with sets show ?thesis unfolding \<open>I = J\<close>
     by (intro antisym prod_algebra_mono) auto
 qed
 
-lemma%unimportant space_in_prod_algebra:
+lemma space_in_prod_algebra:
   "(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
 proof cases
   assume "I = {}" then show ?thesis
@@ -341,26 +340,26 @@
   finally show ?thesis .
 qed
 
-lemma%unimportant space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
+lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
   using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp
 
-lemma%unimportant prod_emb_subset_PiM[simp]: "prod_emb I M K X \<subseteq> space (PiM I M)"
+lemma prod_emb_subset_PiM[simp]: "prod_emb I M K X \<subseteq> space (PiM I M)"
   by (auto simp: prod_emb_def space_PiM)
 
-lemma%unimportant space_PiM_empty_iff[simp]: "space (PiM I M) = {} \<longleftrightarrow>  (\<exists>i\<in>I. space (M i) = {})"
+lemma space_PiM_empty_iff[simp]: "space (PiM I M) = {} \<longleftrightarrow>  (\<exists>i\<in>I. space (M i) = {})"
   by (auto simp: space_PiM PiE_eq_empty_iff)
 
-lemma%unimportant undefined_in_PiM_empty[simp]: "(\<lambda>x. undefined) \<in> space (PiM {} M)"
+lemma undefined_in_PiM_empty[simp]: "(\<lambda>x. undefined) \<in> space (PiM {} M)"
   by (auto simp: space_PiM)
 
-lemma%unimportant sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
+lemma sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
   using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp
 
-lemma%important sets_PiM_single: "sets (PiM I M) =
+proposition sets_PiM_single: "sets (PiM I M) =
     sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
     (is "_ = sigma_sets ?\<Omega> ?R")
   unfolding sets_PiM
-proof%unimportant (rule sigma_sets_eqI)
+proof (rule sigma_sets_eqI)
   interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
   fix A assume "A \<in> prod_algebra I M"
   from prod_algebraE[OF this] guess J X . note X = this
@@ -388,7 +387,7 @@
   finally show "A \<in> sigma_sets ?\<Omega> (prod_algebra I M)" .
 qed
 
-lemma%unimportant sets_PiM_eq_proj:
+lemma sets_PiM_eq_proj:
   "I \<noteq> {} \<Longrightarrow> sets (PiM I M) = sets (SUP i\<in>I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i))"
   apply (simp add: sets_PiM_single)
   apply (subst sets_Sup_eq[where X="\<Pi>\<^sub>E i\<in>I. space (M i)"])
@@ -400,18 +399,18 @@
   apply (auto intro!: arg_cong2[where f=sigma_sets])
   done
 
-lemma%unimportant (*FIX ME needs name *)
+lemma (*FIX ME needs name *)
   shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\<lambda>k. undefined}"
     and sets_PiM_empty: "sets (Pi\<^sub>M {} M) = { {}, {\<lambda>k. undefined} }"
   by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
 
-lemma%important sets_PiM_sigma:
+proposition sets_PiM_sigma:
   assumes \<Omega>_cover: "\<And>i. i \<in> I \<Longrightarrow> \<exists>S\<subseteq>E i. countable S \<and> \<Omega> i = \<Union>S"
   assumes E: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (\<Omega> i)"
   assumes J: "\<And>j. j \<in> J \<Longrightarrow> finite j" "\<Union>J = I"
   defines "P \<equiv> {{f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i} | A j. j \<in> J \<and> A \<in> Pi j E}"
   shows "sets (\<Pi>\<^sub>M i\<in>I. sigma (\<Omega> i) (E i)) = sets (sigma (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P)"
-proof%unimportant cases
+proof cases
   assume "I = {}"
   with \<open>\<Union>J = I\<close> have "P = {{\<lambda>_. undefined}} \<or> P = {}"
     by (auto simp: P_def)
@@ -495,21 +494,21 @@
   finally show "?thesis" .
 qed
 
-lemma%unimportant sets_PiM_in_sets:
+lemma sets_PiM_in_sets:
   assumes space: "space N = (\<Pi>\<^sub>E i\<in>I. space (M i))"
   assumes sets: "\<And>i A. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {x\<in>space N. x i \<in> A} \<in> sets N"
   shows "sets (\<Pi>\<^sub>M i \<in> I. M i) \<subseteq> sets N"
   unfolding sets_PiM_single space[symmetric]
   by (intro sets.sigma_sets_subset subsetI) (auto intro: sets)
 
-lemma%unimportant sets_PiM_cong[measurable_cong]:
+lemma sets_PiM_cong[measurable_cong]:
   assumes "I = J" "\<And>i. i \<in> J \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (PiM I M) = sets (PiM J N)"
   using assms sets_eq_imp_space_eq[OF assms(2)] by (simp add: sets_PiM_single cong: PiE_cong conj_cong)
 
-lemma%important sets_PiM_I:
+lemma sets_PiM_I:
   assumes "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
   shows "prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)"
-proof%unimportant cases
+proof cases
   assume "J = {}"
   then have "prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j) = (\<Pi>\<^sub>E j\<in>I. space (M j))"
     by (auto simp: prod_emb_def)
@@ -520,7 +519,7 @@
     by (force simp add: sets_PiM prod_algebra_def)
 qed
 
-lemma%unimportant measurable_PiM:
+proposition measurable_PiM:
   assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
   assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
     f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N"
@@ -532,13 +531,13 @@
   with sets[of J X] show "f -` A \<inter> space N \<in> sets N" by auto
 qed
 
-lemma%important measurable_PiM_Collect:
+lemma measurable_PiM_Collect:
   assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
   assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
     {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N"
   shows "f \<in> measurable N (PiM I M)"
   using sets_PiM prod_algebra_sets_into_space space
-proof%unimportant (rule measurable_sigma_sets)
+proof (rule measurable_sigma_sets)
   fix A assume "A \<in> prod_algebra I M"
   from prod_algebraE[OF this] guess J X . note X = this
   then have "f -` A \<inter> space N = {\<omega> \<in> space N. \<forall>i\<in>J. f \<omega> i \<in> X i}"
@@ -547,7 +546,7 @@
   finally show "f -` A \<inter> space N \<in> sets N" .
 qed
 
-lemma%unimportant measurable_PiM_single:
+lemma measurable_PiM_single:
   assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
   assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N"
   shows "f \<in> measurable N (PiM I M)"
@@ -561,11 +560,11 @@
   finally show "f -` A \<inter> space N \<in> sets N" .
 qed (auto simp: space)
 
-lemma%important measurable_PiM_single':
+lemma measurable_PiM_single':
   assumes f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> measurable N (M i)"
     and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
   shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^sub>M I M)"
-proof%unimportant (rule measurable_PiM_single)
+proof (rule measurable_PiM_single)
   fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
   then have "{\<omega> \<in> space N. f i \<omega> \<in> A} = f i -` A \<inter> space N"
     by auto
@@ -573,12 +572,12 @@
     using A f by (auto intro!: measurable_sets)
 qed fact
 
-lemma%unimportant sets_PiM_I_finite[measurable]:
+lemma sets_PiM_I_finite[measurable]:
   assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))"
   shows "(\<Pi>\<^sub>E j\<in>I. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)"
   using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] \<open>finite I\<close> sets by auto
 
-lemma%unimportant measurable_component_singleton[measurable (raw)]:
+lemma measurable_component_singleton[measurable (raw)]:
   assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^sub>M I M) (M i)"
 proof (unfold measurable_def, intro CollectI conjI ballI)
   fix A assume "A \<in> sets (M i)"
@@ -589,30 +588,30 @@
     using \<open>A \<in> sets (M i)\<close> \<open>i \<in> I\<close> by (auto intro!: sets_PiM_I)
 qed (insert \<open>i \<in> I\<close>, auto simp: space_PiM)
 
-lemma%unimportant measurable_component_singleton'[measurable_dest]:
+lemma measurable_component_singleton'[measurable_dest]:
   assumes f: "f \<in> measurable N (Pi\<^sub>M I M)"
   assumes g: "g \<in> measurable L N"
   assumes i: "i \<in> I"
   shows "(\<lambda>x. (f (g x)) i) \<in> measurable L (M i)"
   using measurable_compose[OF measurable_compose[OF g f] measurable_component_singleton, OF i] .
 
-lemma%unimportant measurable_PiM_component_rev:
+lemma measurable_PiM_component_rev:
   "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<lambda>x. f (x i)) \<in> measurable (PiM I M) N"
   by simp
 
-lemma%unimportant measurable_case_nat[measurable (raw)]:
+lemma measurable_case_nat[measurable (raw)]:
   assumes [measurable (raw)]: "i = 0 \<Longrightarrow> f \<in> measurable M N"
     "\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N"
   shows "(\<lambda>x. case_nat (f x) (g x) i) \<in> measurable M N"
   by (cases i) simp_all
 
-lemma%unimportant measurable_case_nat'[measurable (raw)]:
+lemma measurable_case_nat'[measurable (raw)]:
   assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
   shows "(\<lambda>x. case_nat (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
   using fg[THEN measurable_space]
   by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split)
 
-lemma%unimportant measurable_add_dim[measurable]:
+lemma measurable_add_dim[measurable]:
   "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M)"
     (is "?f \<in> measurable ?P ?I")
 proof (rule measurable_PiM_single)
@@ -626,12 +625,12 @@
   finally show "{\<omega> \<in> space ?P. case_prod (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
 qed (auto simp: space_pair_measure space_PiM PiE_def)
 
-lemma%important measurable_fun_upd:
+proposition measurable_fun_upd:
   assumes I: "I = J \<union> {i}"
   assumes f[measurable]: "f \<in> measurable N (PiM J M)"
   assumes h[measurable]: "h \<in> measurable N (M i)"
   shows "(\<lambda>x. (f x) (i := h x)) \<in> measurable N (PiM I M)"
-proof%unimportant (intro measurable_PiM_single')
+proof (intro measurable_PiM_single')
   fix j assume "j \<in> I" then show "(\<lambda>\<omega>. ((f \<omega>)(i := h \<omega>)) j) \<in> measurable N (M j)"
     unfolding I by (cases "j = i") auto
 next
@@ -640,14 +639,14 @@
     by (auto simp: space_PiM PiE_iff extensional_def)
 qed
 
-lemma%unimportant measurable_component_update:
+lemma measurable_component_update:
   "x \<in> space (Pi\<^sub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^sub>M (insert i I) M)"
   by simp
 
-lemma%important measurable_merge[measurable]:
+lemma measurable_merge[measurable]:
   "merge I J \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M)"
     (is "?f \<in> measurable ?P ?U")
-proof%unimportant (rule measurable_PiM_single)
+proof (rule measurable_PiM_single)
   fix i A assume A: "A \<in> sets (M i)" "i \<in> I \<union> J"
   then have "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} =
     (if i \<in> I then ((\<lambda>x. x i) \<circ> fst) -` A \<inter> space ?P else ((\<lambda>x. x i) \<circ> snd) -` A \<inter> space ?P)"
@@ -658,7 +657,7 @@
   finally show "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} \<in> sets ?P" .
 qed (auto simp: space_pair_measure space_PiM PiE_iff merge_def extensional_def)
 
-lemma%unimportant measurable_restrict[measurable (raw)]:
+lemma measurable_restrict[measurable (raw)]:
   assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)"
   shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^sub>M I M)"
 proof (rule measurable_PiM_single)
@@ -669,14 +668,14 @@
     using A X by (auto intro!: measurable_sets)
 qed (insert X, auto simp add: PiE_def dest: measurable_space)
 
-lemma%unimportant measurable_abs_UNIV:
+lemma measurable_abs_UNIV:
   "(\<And>n. (\<lambda>\<omega>. f n \<omega>) \<in> measurable M (N n)) \<Longrightarrow> (\<lambda>\<omega> n. f n \<omega>) \<in> measurable M (PiM UNIV N)"
   by (intro measurable_PiM_single) (auto dest: measurable_space)
 
-lemma%unimportant measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
+lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
   by (intro measurable_restrict measurable_component_singleton) auto
 
-lemma%unimportant measurable_restrict_subset':
+lemma measurable_restrict_subset':
   assumes "J \<subseteq> L" "\<And>x. x \<in> J \<Longrightarrow> sets (M x) = sets (N x)"
   shows "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)"
 proof-
@@ -687,12 +686,12 @@
   finally show ?thesis .
 qed
 
-lemma%unimportant measurable_prod_emb[intro, simp]:
+lemma measurable_prod_emb[intro, simp]:
   "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^sub>M L M)"
   unfolding prod_emb_def space_PiM[symmetric]
   by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
 
-lemma%unimportant merge_in_prod_emb:
+lemma merge_in_prod_emb:
   assumes "y \<in> space (PiM I M)" "x \<in> X" and X: "X \<in> sets (Pi\<^sub>M J M)" and "J \<subseteq> I"
   shows "merge J I (x, y) \<in> prod_emb I M J X"
   using assms sets.sets_into_space[OF X]
@@ -700,7 +699,7 @@
            cong: if_cong restrict_cong)
      (simp add: extensional_def)
 
-lemma%unimportant prod_emb_eq_emptyD:
+lemma prod_emb_eq_emptyD:
   assumes J: "J \<subseteq> I" and ne: "space (PiM I M) \<noteq> {}" and X: "X \<in> sets (Pi\<^sub>M J M)"
     and *: "prod_emb I M J X = {}"
   shows "X = {}"
@@ -711,19 +710,19 @@
   from merge_in_prod_emb[OF this \<open>x\<in>X\<close> X J] * show "x \<in> {}" by auto
 qed
 
-lemma%unimportant sets_in_Pi_aux:
+lemma sets_in_Pi_aux:
   "finite I \<Longrightarrow> (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
   {x\<in>space (PiM I M). x \<in> Pi I F} \<in> sets (PiM I M)"
   by (simp add: subset_eq Pi_iff)
 
-lemma%unimportant sets_in_Pi[measurable (raw)]:
+lemma sets_in_Pi[measurable (raw)]:
   "finite I \<Longrightarrow> f \<in> measurable N (PiM I M) \<Longrightarrow>
   (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
   Measurable.pred N (\<lambda>x. f x \<in> Pi I F)"
   unfolding pred_def
   by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_Pi_aux]) auto
 
-lemma%unimportant sets_in_extensional_aux:
+lemma sets_in_extensional_aux:
   "{x\<in>space (PiM I M). x \<in> extensional I} \<in> sets (PiM I M)"
 proof -
   have "{x\<in>space (PiM I M). x \<in> extensional I} = space (PiM I M)"
@@ -731,12 +730,12 @@
   then show ?thesis by simp
 qed
 
-lemma%unimportant sets_in_extensional[measurable (raw)]:
+lemma sets_in_extensional[measurable (raw)]:
   "f \<in> measurable N (PiM I M) \<Longrightarrow> Measurable.pred N (\<lambda>x. f x \<in> extensional I)"
   unfolding pred_def
   by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_extensional_aux]) auto
 
-lemma%unimportant sets_PiM_I_countable:
+lemma sets_PiM_I_countable:
   assumes I: "countable I" and E: "\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i)" shows "Pi\<^sub>E I E \<in> sets (Pi\<^sub>M I M)"
 proof cases
   assume "I \<noteq> {}"
@@ -747,11 +746,11 @@
   finally show ?thesis .
 qed (simp add: sets_PiM_empty)
 
-lemma%important sets_PiM_D_countable:
+lemma sets_PiM_D_countable:
   assumes A: "A \<in> PiM I M"
   shows "\<exists>J\<subseteq>I. \<exists>X\<in>PiM J M. countable J \<and> A = prod_emb I M J X"
   using A[unfolded sets_PiM_single]
-proof%unimportant induction
+proof induction
   case (Basic A)
   then obtain i X where *: "i \<in> I" "X \<in> sets (M i)" and "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> X}"
     by auto
@@ -783,12 +782,12 @@
   qed(auto intro: X)
 qed
 
-lemma%important measure_eqI_PiM_finite:
+proposition measure_eqI_PiM_finite:
   assumes [simp]: "finite I" "sets P = PiM I M" "sets Q = PiM I M"
   assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = Q (Pi\<^sub>E I A)"
   assumes A: "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = space (PiM I M)" "\<And>i::nat. P (A i) \<noteq> \<infinity>"
   shows "P = Q"
-proof%unimportant (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
+proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
   show "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. P (A i) \<noteq> \<infinity>"
     unfolding space_PiM[symmetric] by fact+
   fix X assume "X \<in> prod_algebra I M"
@@ -799,13 +798,13 @@
     unfolding X by (subst (1 2) prod_emb_Pi) (auto simp: eq)
 qed (simp_all add: sets_PiM)
 
-lemma%important measure_eqI_PiM_infinite:
+proposition measure_eqI_PiM_infinite:
   assumes [simp]: "sets P = PiM I M" "sets Q = PiM I M"
   assumes eq: "\<And>A J. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow>
     P (prod_emb I M J (Pi\<^sub>E J A)) = Q (prod_emb I M J (Pi\<^sub>E J A))"
   assumes A: "finite_measure P"
   shows "P = Q"
-proof%unimportant (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
+proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
   interpret finite_measure P by fact
   define i where "i = (SOME i. i \<in> I)"
   have i: "I \<noteq> {} \<Longrightarrow> i \<in> I"
@@ -839,12 +838,12 @@
   fixes I :: "'i set"
   assumes finite_index: "finite I"
 
-lemma%important (in finite_product_sigma_finite) sigma_finite_pairs:
+proposition (in finite_product_sigma_finite) sigma_finite_pairs:
   "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
     (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
     (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k) \<and>
     (\<Union>k. \<Pi>\<^sub>E i\<in>I. F i k) = space (PiM I M)"
-proof%unimportant -
+proof -
   have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. emeasure (M i) (F k) \<noteq> \<infinity>)"
     using M.sigma_finite_incseq by metis
   from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
@@ -870,7 +869,7 @@
   qed
 qed
 
-lemma%unimportant emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
+lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
 proof -
   let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ennreal)"
   have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = 1"
@@ -887,12 +886,12 @@
     by simp
 qed
 
-lemma%unimportant PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}"
+lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}"
   by (rule measure_eqI) (auto simp add: sets_PiM_empty)
 
-lemma%important (in product_sigma_finite) emeasure_PiM:
+lemma (in product_sigma_finite) emeasure_PiM:
   "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
-proof%unimportant (induct I arbitrary: A rule: finite_induct)
+proof (induct I arbitrary: A rule: finite_induct)
   case (insert i I)
   interpret finite_product_sigma_finite M I by standard fact
   have "finite (insert i I)" using \<open>finite I\<close> by auto
@@ -947,7 +946,7 @@
     by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space)
 qed simp
 
-lemma%unimportant (in product_sigma_finite) PiM_eqI:
+lemma (in product_sigma_finite) PiM_eqI:
   assumes I[simp]: "finite I" and P: "sets P = PiM I M"
   assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
   shows "P = PiM I M"
@@ -965,7 +964,7 @@
   qed
 qed
 
-lemma%unimportant (in product_sigma_finite) sigma_finite:
+lemma (in product_sigma_finite) sigma_finite:
   assumes "finite I"
   shows "sigma_finite_measure (PiM I M)"
 proof
@@ -986,11 +985,11 @@
 sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^sub>M I M"
   using sigma_finite[OF finite_index] .
 
-lemma%unimportant (in finite_product_sigma_finite) measure_times:
+lemma (in finite_product_sigma_finite) measure_times:
   "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
   using emeasure_PiM[OF finite_index] by auto
 
-lemma%unimportant (in product_sigma_finite) nn_integral_empty:
+lemma (in product_sigma_finite) nn_integral_empty:
   "0 \<le> f (\<lambda>k. undefined) \<Longrightarrow> integral\<^sup>N (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)"
   by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2)
 
@@ -998,7 +997,7 @@
   assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
   shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M"
    (is "?D = ?P")
-proof%unimportant (rule PiM_eqI)
+proof (rule PiM_eqI)
   interpret I: finite_product_sigma_finite M I by standard fact
   interpret J: finite_product_sigma_finite M J by standard fact
   fix A assume A: "\<And>i. i \<in> I \<union> J \<Longrightarrow> A i \<in> sets (M i)"
@@ -1010,12 +1009,12 @@
        (auto simp: * J.emeasure_pair_measure_Times I.measure_times J.measure_times prod.union_disjoint)
 qed (insert fin, simp_all)
 
-lemma%important (in product_sigma_finite) product_nn_integral_fold:
+proposition (in product_sigma_finite) product_nn_integral_fold:
   assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
   and f[measurable]: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
   shows "integral\<^sup>N (Pi\<^sub>M (I \<union> J) M) f =
     (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^sub>M J M)) \<partial>(Pi\<^sub>M I M))"
-proof%unimportant -
+proof -
   interpret I: finite_product_sigma_finite M I by standard fact
   interpret J: finite_product_sigma_finite M J by standard fact
   interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by standard
@@ -1030,7 +1029,7 @@
     done
 qed
 
-lemma%unimportant (in product_sigma_finite) distr_singleton:
+lemma (in product_sigma_finite) distr_singleton:
   "distr (Pi\<^sub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _")
 proof (intro measure_eqI[symmetric])
   interpret I: finite_product_sigma_finite M "{i}" by standard simp
@@ -1042,7 +1041,7 @@
     by (simp add: emeasure_distr measurable_component_singleton)
 qed simp
 
-lemma%unimportant (in product_sigma_finite) product_nn_integral_singleton:
+lemma (in product_sigma_finite) product_nn_integral_singleton:
   assumes f: "f \<in> borel_measurable (M i)"
   shows "integral\<^sup>N (Pi\<^sub>M {i} M) (\<lambda>x. f (x i)) = integral\<^sup>N (M i) f"
 proof -
@@ -1054,11 +1053,11 @@
     done
 qed
 
-lemma%important (in product_sigma_finite) product_nn_integral_insert:
+proposition (in product_sigma_finite) product_nn_integral_insert:
   assumes I[simp]: "finite I" "i \<notin> I"
     and f: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
   shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^sub>M I M))"
-proof%unimportant -
+proof -
   interpret I: finite_product_sigma_finite M I by standard auto
   interpret i: finite_product_sigma_finite M "{i}" by standard auto
   have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
@@ -1078,7 +1077,7 @@
   qed
 qed
 
-lemma%unimportant (in product_sigma_finite) product_nn_integral_insert_rev:
+lemma (in product_sigma_finite) product_nn_integral_insert_rev:
   assumes I[simp]: "finite I" "i \<notin> I"
     and [measurable]: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
   shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x(i := y)) \<partial>(Pi\<^sub>M I M)) \<partial>(M i))"
@@ -1089,7 +1088,7 @@
   apply measurable
   done
 
-lemma%unimportant (in product_sigma_finite) product_nn_integral_prod:
+lemma (in product_sigma_finite) product_nn_integral_prod:
   assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
   shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>N (M i) (f i))"
 using assms proof (induction I)
@@ -1112,11 +1111,11 @@
     done
 qed (simp add: space_PiM)
 
-lemma%important (in product_sigma_finite) product_nn_integral_pair:
+proposition (in product_sigma_finite) product_nn_integral_pair:
   assumes [measurable]: "case_prod f \<in> borel_measurable (M x \<Otimes>\<^sub>M M y)"
   assumes xy: "x \<noteq> y"
   shows "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {x, y} M) = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))"
-proof%unimportant -
+proof -
   interpret psm: pair_sigma_finite "M x" "M y"
     unfolding pair_sigma_finite_def using sigma_finite_measures by simp_all
   have "{x, y} = {y, x}" by auto
@@ -1129,7 +1128,7 @@
   finally show ?thesis .
 qed
 
-lemma%unimportant (in product_sigma_finite) distr_component:
+lemma (in product_sigma_finite) distr_component:
   "distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P")
 proof (intro PiM_eqI)
   fix A assume A: "\<And>ia. ia \<in> {i} \<Longrightarrow> A ia \<in> sets (M ia)"
@@ -1139,7 +1138,7 @@
     by (subst emeasure_distr) (auto intro!: sets_PiM_I_finite measurable_restrict)
 qed simp_all
 
-lemma%unimportant (in product_sigma_finite)
+lemma (in product_sigma_finite)
   assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)"
   shows emeasure_fold_integral:
     "emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" (is ?I)
@@ -1164,11 +1163,11 @@
     by (simp add: vimage_comp comp_def space_pair_measure cong: measurable_cong)
 qed
 
-lemma%unimportant sets_Collect_single:
+lemma sets_Collect_single:
   "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^sub>M I M). x i \<in> A } \<in> sets (Pi\<^sub>M I M)"
   by simp
 
-lemma%unimportant pair_measure_eq_distr_PiM:
+lemma pair_measure_eq_distr_PiM:
   fixes M1 :: "'a measure" and M2 :: "'a measure"
   assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
   shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (case_bool M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))"
@@ -1195,7 +1194,7 @@
     by (subst emeasure_distr) (auto simp: measurable_pair_iff)
 qed simp
 
-lemma%unimportant infprod_in_sets[intro]:
+lemma infprod_in_sets[intro]:
   fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets (M i)"
   shows "Pi UNIV E \<in> sets (\<Pi>\<^sub>M i\<in>UNIV::nat set. M i)"
 proof -
--- a/src/HOL/Analysis/Function_Topology.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Function_Topology.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -78,14 +78,14 @@
 | UN: "(\<And>k. k \<in> K \<Longrightarrow> generate_topology_on S k) \<Longrightarrow> generate_topology_on S (\<Union>K)"
 | Basis: "s \<in> S \<Longrightarrow> generate_topology_on S s"
 
-lemma%unimportant istopology_generate_topology_on:
+lemma istopology_generate_topology_on:
   "istopology (generate_topology_on S)"
 unfolding istopology_def by (auto intro: generate_topology_on.intros)
 
 text \<open>The basic property of the topology generated by a set \<open>S\<close> is that it is the
 smallest topology containing all the elements of \<open>S\<close>:\<close>
 
-lemma%unimportant generate_topology_on_coarsest:
+lemma generate_topology_on_coarsest:
   assumes "istopology T"
           "\<And>s. s \<in> S \<Longrightarrow> T s"
           "generate_topology_on S s0"
@@ -96,17 +96,17 @@
 abbreviation%unimportant topology_generated_by::"('a set set) \<Rightarrow> ('a topology)"
   where "topology_generated_by S \<equiv> topology (generate_topology_on S)"
 
-lemma%unimportant openin_topology_generated_by_iff:
+lemma openin_topology_generated_by_iff:
   "openin (topology_generated_by S) s \<longleftrightarrow> generate_topology_on S s"
   using topology_inverse'[OF istopology_generate_topology_on[of S]] by simp
 
-lemma%unimportant openin_topology_generated_by:
+lemma openin_topology_generated_by:
   "openin (topology_generated_by S) s \<Longrightarrow> generate_topology_on S s"
 using openin_topology_generated_by_iff by auto
 
-lemma%important topology_generated_by_topspace:
+lemma topology_generated_by_topspace:
   "topspace (topology_generated_by S) = (\<Union>S)"
-proof%unimportant
+proof
   {
     fix s assume "openin (topology_generated_by S) s"
     then have "generate_topology_on S s" by (rule openin_topology_generated_by)
@@ -121,9 +121,9 @@
     unfolding topspace_def using openin_topology_generated_by_iff by auto
 qed
 
-lemma%important topology_generated_by_Basis:
+lemma topology_generated_by_Basis:
   "s \<in> S \<Longrightarrow> openin (topology_generated_by S) s"
-by%unimportant (simp only: openin_topology_generated_by_iff, auto simp: generate_topology_on.Basis)
+  by (simp only: openin_topology_generated_by_iff, auto simp: generate_topology_on.Basis)
 
 lemma generate_topology_on_Inter:
   "\<lbrakk>finite \<F>; \<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on \<S> K; \<F> \<noteq> {}\<rbrakk> \<Longrightarrow> generate_topology_on \<S> (\<Inter>\<F>)"
@@ -266,28 +266,28 @@
   where "continuous_on_topo T1 T2 f = ((\<forall> U. openin T2 U \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1)))
                                       \<and> (f`(topspace T1) \<subseteq> (topspace T2)))"
 
-lemma%important continuous_on_continuous_on_topo:
+lemma continuous_on_continuous_on_topo:
   "continuous_on s f \<longleftrightarrow> continuous_on_topo (subtopology euclidean s) euclidean f"
-unfolding%unimportant continuous_on_open_invariant openin_open vimage_def continuous_on_topo_def
+  unfolding continuous_on_open_invariant openin_open vimage_def continuous_on_topo_def
 topspace_euclidean_subtopology open_openin topspace_euclidean by fast
 
-lemma%unimportant continuous_on_topo_UNIV:
+lemma continuous_on_topo_UNIV:
   "continuous_on UNIV f \<longleftrightarrow> continuous_on_topo euclidean euclidean f"
 using continuous_on_continuous_on_topo[of UNIV f] subtopology_UNIV[of euclidean] by auto
 
-lemma%important continuous_on_topo_open [intro]:
+lemma continuous_on_topo_open [intro]:
   "continuous_on_topo T1 T2 f \<Longrightarrow> openin T2 U \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
-unfolding%unimportant continuous_on_topo_def by auto
+  unfolding continuous_on_topo_def by auto
 
-lemma%unimportant continuous_on_topo_topspace [intro]:
+lemma continuous_on_topo_topspace [intro]:
   "continuous_on_topo T1 T2 f \<Longrightarrow> f`(topspace T1) \<subseteq> (topspace T2)"
 unfolding continuous_on_topo_def by auto
 
-lemma%important continuous_on_generated_topo_iff:
+lemma continuous_on_generated_topo_iff:
   "continuous_on_topo T1 (topology_generated_by S) f \<longleftrightarrow>
       ((\<forall>U. U \<in> S \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1))) \<and> (f`(topspace T1) \<subseteq> (\<Union> S)))"
 unfolding continuous_on_topo_def topology_generated_by_topspace
-proof%unimportant (auto simp add: topology_generated_by_Basis)
+proof (auto simp add: topology_generated_by_Basis)
   assume H: "\<forall>U. U \<in> S \<longrightarrow> openin T1 (f -` U \<inter> topspace T1)"
   fix U assume "openin (topology_generated_by S) U"
   then have "generate_topology_on S U" by (rule openin_topology_generated_by)
@@ -309,17 +309,17 @@
   qed (auto simp add: H)
 qed
 
-lemma%important continuous_on_generated_topo:
+lemma continuous_on_generated_topo:
   assumes "\<And>U. U \<in>S \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
           "f`(topspace T1) \<subseteq> (\<Union> S)"
   shows "continuous_on_topo T1 (topology_generated_by S) f"
-using%unimportant assms continuous_on_generated_topo_iff by blast
+  using assms continuous_on_generated_topo_iff by blast
 
-lemma%important continuous_on_topo_compose:
+proposition continuous_on_topo_compose:
   assumes "continuous_on_topo T1 T2 f" "continuous_on_topo T2 T3 g"
   shows "continuous_on_topo T1 T3 (g o f)"
-using%unimportant assms unfolding continuous_on_topo_def
-proof%unimportant (auto)
+  using assms unfolding continuous_on_topo_def
+proof (auto)
   fix U :: "'c set"
   assume H: "openin T3 U"
   have "openin T1 (f -` (g -` U \<inter> topspace T2) \<inter> topspace T1)"
@@ -330,7 +330,7 @@
     by simp
 qed (blast)
 
-lemma%unimportant continuous_on_topo_preimage_topspace [intro]:
+lemma continuous_on_topo_preimage_topspace [intro]:
   assumes "continuous_on_topo T1 T2 f"
   shows "f-`(topspace T2) \<inter> topspace T1 = topspace T1"
 using assms unfolding continuous_on_topo_def by auto
@@ -349,9 +349,9 @@
 definition%important pullback_topology::"('a set) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b topology) \<Rightarrow> ('a topology)"
   where "pullback_topology A f T = topology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
 
-lemma%important istopology_pullback_topology:
+lemma istopology_pullback_topology:
   "istopology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
-unfolding%unimportant istopology_def proof (auto)
+  unfolding istopology_def proof (auto)
   fix K assume "\<forall>S\<in>K. \<exists>U. openin T U \<and> S = f -` U \<inter> A"
   then have "\<exists>U. \<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
     by (rule bchoice)
@@ -362,19 +362,19 @@
   then show "\<exists>V. openin T V \<and> \<Union>K = f -` V \<inter> A" by auto
 qed
 
-lemma%unimportant openin_pullback_topology:
+lemma openin_pullback_topology:
   "openin (pullback_topology A f T) S \<longleftrightarrow> (\<exists>U. openin T U \<and> S = f-`U \<inter> A)"
 unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto
 
-lemma%unimportant topspace_pullback_topology:
+lemma topspace_pullback_topology:
   "topspace (pullback_topology A f T) = f-`(topspace T) \<inter> A"
 by (auto simp add: topspace_def openin_pullback_topology)
 
-lemma%important continuous_on_topo_pullback [intro]:
+proposition continuous_on_topo_pullback [intro]:
   assumes "continuous_on_topo T1 T2 g"
   shows "continuous_on_topo (pullback_topology A f T1) T2 (g o f)"
 unfolding continuous_on_topo_def
-proof%unimportant (auto)
+proof (auto)
   fix U::"'b set" assume "openin T2 U"
   then have "openin T1 (g-`U \<inter> topspace T1)"
     using assms unfolding continuous_on_topo_def by auto
@@ -394,11 +394,11 @@
     using assms unfolding continuous_on_topo_def by auto
 qed
 
-lemma%important continuous_on_topo_pullback' [intro]:
+proposition continuous_on_topo_pullback' [intro]:
   assumes "continuous_on_topo T1 T2 (f o g)" "topspace T1 \<subseteq> g-`A"
   shows "continuous_on_topo T1 (pullback_topology A f T2) g"
 unfolding continuous_on_topo_def
-proof%unimportant (auto)
+proof (auto)
   fix U assume "openin (pullback_topology A f T2) U"
   then have "\<exists>V. openin T2 V \<and> U = f-`V \<inter> A"
     unfolding openin_pullback_topology by auto
@@ -547,9 +547,9 @@
     done
 qed
 
-lemma%important topspace_product_topology:
+lemma topspace_product_topology:
   "topspace (product_topology T I) = (\<Pi>\<^sub>E i\<in>I. topspace(T i))"
-proof%unimportant
+proof
   show "topspace (product_topology T I) \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (T i))"
     unfolding product_topology_def topology_generated_by_topspace
     unfolding topspace_def by auto
@@ -559,16 +559,16 @@
     unfolding product_topology_def using PiE_def by (auto simp add: topology_generated_by_topspace)
 qed
 
-lemma%unimportant product_topology_basis:
+lemma product_topology_basis:
   assumes "\<And>i. openin (T i) (X i)" "finite {i. X i \<noteq> topspace (T i)}"
   shows "openin (product_topology T I) (\<Pi>\<^sub>E i\<in>I. X i)"
   unfolding product_topology_def
   by (rule topology_generated_by_Basis) (use assms in auto)
 
-lemma%important product_topology_open_contains_basis:
+proposition product_topology_open_contains_basis:
   assumes "openin (product_topology T I) U" "x \<in> U"
   shows "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
-proof%unimportant -
+proof -
   have "generate_topology_on {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}} U"
     using assms unfolding product_topology_def by (intro openin_topology_generated_by) auto
   then have "\<And>x. x\<in>U \<Longrightarrow> \<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
@@ -717,7 +717,7 @@
 
 text \<open>The basic property of the product topology is the continuity of projections:\<close>
 
-lemma%unimportant continuous_on_topo_product_coordinates [simp]:
+lemma continuous_on_topo_product_coordinates [simp]:
   assumes "i \<in> I"
   shows "continuous_on_topo (product_topology T I) (T i) (\<lambda>x. x i)"
 proof -
@@ -739,12 +739,12 @@
     by (auto simp add: assms topspace_product_topology PiE_iff)
 qed
 
-lemma%important continuous_on_topo_coordinatewise_then_product [intro]:
+lemma continuous_on_topo_coordinatewise_then_product [intro]:
   assumes "\<And>i. i \<in> I \<Longrightarrow> continuous_on_topo T1 (T i) (\<lambda>x. f x i)"
           "\<And>i x. i \<notin> I \<Longrightarrow> x \<in> topspace T1 \<Longrightarrow> f x i = undefined"
   shows "continuous_on_topo T1 (product_topology T I) f"
 unfolding product_topology_def
-proof%unimportant (rule continuous_on_generated_topo)
+proof (rule continuous_on_generated_topo)
   fix U assume "U \<in> {Pi\<^sub>E I X |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
   then obtain X where H: "U = Pi\<^sub>E I X" "\<And>i. openin (T i) (X i)" "finite {i. X i \<noteq> topspace (T i)}"
     by blast
@@ -772,7 +772,7 @@
     using assms unfolding continuous_on_topo_def by auto
 qed
 
-lemma%unimportant continuous_on_topo_product_then_coordinatewise [intro]:
+lemma continuous_on_topo_product_then_coordinatewise [intro]:
   assumes "continuous_on_topo T1 (product_topology T I) f"
   shows "\<And>i. i \<in> I \<Longrightarrow> continuous_on_topo T1 (T i) (\<lambda>x. f x i)"
         "\<And>i x. i \<notin> I \<Longrightarrow> x \<in> topspace T1 \<Longrightarrow> f x i = undefined"
@@ -794,7 +794,7 @@
     using \<open>i \<notin> I\<close> by (auto simp add: PiE_iff extensional_def)
 qed
 
-lemma%unimportant continuous_on_restrict:
+lemma continuous_on_restrict:
   assumes "J \<subseteq> I"
   shows "continuous_on_topo (product_topology T I) (product_topology T J) (\<lambda>x. restrict x J)"
 proof (rule continuous_on_topo_coordinatewise_then_product)
@@ -817,7 +817,7 @@
 definition%important open_fun_def:
   "open U = openin (product_topology (\<lambda>i. euclidean) UNIV) U"
 
-instance proof%unimportant
+instance proof
   have "topspace (product_topology (\<lambda>(i::'a). euclidean::('b topology)) UNIV) = UNIV"
     unfolding topspace_product_topology topspace_euclidean by auto
   then show "open (UNIV::('a \<Rightarrow> 'b) set)"
@@ -826,15 +826,15 @@
 
 end
 
-lemma%unimportant euclidean_product_topology:
+lemma euclidean_product_topology:
   "euclidean = product_topology (\<lambda>i. euclidean::('b::topological_space) topology) UNIV"
 by (metis open_openin topology_eq open_fun_def)
 
-lemma%important product_topology_basis':
+proposition product_topology_basis':
   fixes x::"'i \<Rightarrow> 'a" and U::"'i \<Rightarrow> ('b::topological_space) set"
   assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> open (U i)"
   shows "open {f. \<forall>i\<in>I. f (x i) \<in> U i}"
-proof%unimportant -
+proof -
   define J where "J = x`I"
   define V where "V = (\<lambda>y. if y \<in> J then \<Inter>{U i|i. i\<in>I \<and> x i = y} else UNIV)"
   define X where "X = (\<lambda>y. if y \<in> J then V y else UNIV)"
@@ -864,24 +864,24 @@
 text \<open>The results proved in the general situation of products of possibly different
 spaces have their counterparts in this simpler setting.\<close>
 
-lemma%unimportant continuous_on_product_coordinates [simp]:
+lemma continuous_on_product_coordinates [simp]:
   "continuous_on UNIV (\<lambda>x. x i::('b::topological_space))"
 unfolding continuous_on_topo_UNIV euclidean_product_topology
 by (rule continuous_on_topo_product_coordinates, simp)
 
-lemma%unimportant continuous_on_coordinatewise_then_product [intro, continuous_intros]:
+lemma continuous_on_coordinatewise_then_product [intro, continuous_intros]:
   assumes "\<And>i. continuous_on UNIV (\<lambda>x. f x i)"
   shows "continuous_on UNIV f"
 using assms unfolding continuous_on_topo_UNIV euclidean_product_topology
 by (rule continuous_on_topo_coordinatewise_then_product, simp)
 
-lemma%unimportant continuous_on_product_then_coordinatewise:
+lemma continuous_on_product_then_coordinatewise:
   assumes "continuous_on UNIV f"
   shows "continuous_on UNIV (\<lambda>x. f x i)"
 using assms unfolding continuous_on_topo_UNIV euclidean_product_topology
 by (rule continuous_on_topo_product_then_coordinatewise(1), simp)
 
-lemma%unimportant continuous_on_product_coordinatewise_iff:
+lemma continuous_on_product_coordinatewise_iff:
   "continuous_on UNIV f \<longleftrightarrow> (\<forall>i. continuous_on UNIV (\<lambda>x. f x i))"
 by (auto intro: continuous_on_product_then_coordinatewise)
 
@@ -891,7 +891,7 @@
 of product spaces, but they have more to do with countability and could
 be put in the corresponding theory.\<close>
 
-lemma%unimportant countable_nat_product_event_const:
+lemma countable_nat_product_event_const:
   fixes F::"'a set" and a::'a
   assumes "a \<in> F" "countable F"
   shows "countable {x::(nat \<Rightarrow> 'a). (\<forall>i. x i \<in> F) \<and> finite {i. x i \<noteq> a}}"
@@ -928,7 +928,7 @@
   then show ?thesis using countable_subset[OF *] by auto
 qed
 
-lemma%unimportant countable_product_event_const:
+lemma countable_product_event_const:
   fixes F::"('a::countable) \<Rightarrow> 'b set" and b::'b
   assumes "\<And>i. countable (F i)"
   shows "countable {f::('a \<Rightarrow> 'b). (\<forall>i. f i \<in> F i) \<and> (finite {i. f i \<noteq> b})}"
@@ -961,7 +961,7 @@
 qed
 
 instance "fun" :: (countable, first_countable_topology) first_countable_topology
-proof%unimportant
+proof
   fix x::"'a \<Rightarrow> 'b"
   have "\<exists>A::('b \<Rightarrow> nat \<Rightarrow> 'b set). \<forall>x. (\<forall>i. x \<in> A x i \<and> open (A x i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A x i \<subseteq> S))"
     apply (rule choice) using first_countable_basis by auto
@@ -1033,11 +1033,11 @@
     using \<open>countable K\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> x \<in> k\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> open k\<close> Inc by auto
 qed
 
-lemma%important product_topology_countable_basis:
+proposition product_topology_countable_basis:
   shows "\<exists>K::(('a::countable \<Rightarrow> 'b::second_countable_topology) set set).
           topological_basis K \<and> countable K \<and>
           (\<forall>k\<in>K. \<exists>X. (k = Pi\<^sub>E UNIV X) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV})"
-proof%unimportant -
+proof -
   obtain B::"'b set set" where B: "countable B \<and> topological_basis B"
     using ex_countable_basis by auto
   then have "B \<noteq> {}" by (meson UNIV_I empty_iff open_UNIV topological_basisE)
@@ -1140,15 +1140,15 @@
 definition%important strong_operator_topology::"('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector) topology"
 where "strong_operator_topology = pullback_topology UNIV blinfun_apply euclidean"
 
-lemma%unimportant strong_operator_topology_topspace:
+lemma strong_operator_topology_topspace:
   "topspace strong_operator_topology = UNIV"
 unfolding strong_operator_topology_def topspace_pullback_topology topspace_euclidean by auto
 
-lemma%important strong_operator_topology_basis:
+lemma strong_operator_topology_basis:
   fixes f::"('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector)" and U::"'i \<Rightarrow> 'b set" and x::"'i \<Rightarrow> 'a"
   assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> open (U i)"
   shows "openin strong_operator_topology {f. \<forall>i\<in>I. blinfun_apply f (x i) \<in> U i}"
-proof%unimportant -
+proof -
   have "open {g::('a\<Rightarrow>'b). \<forall>i\<in>I. g (x i) \<in> U i}"
     by (rule product_topology_basis'[OF assms])
   moreover have "{f. \<forall>i\<in>I. blinfun_apply f (x i) \<in> U i}
@@ -1158,16 +1158,16 @@
     unfolding strong_operator_topology_def open_openin apply (subst openin_pullback_topology) by auto
 qed
 
-lemma%important strong_operator_topology_continuous_evaluation:
+lemma strong_operator_topology_continuous_evaluation:
   "continuous_on_topo strong_operator_topology euclidean (\<lambda>f. blinfun_apply f x)"
-proof%unimportant -
+proof -
   have "continuous_on_topo strong_operator_topology euclidean ((\<lambda>f. f x) o blinfun_apply)"
     unfolding strong_operator_topology_def apply (rule continuous_on_topo_pullback)
     using continuous_on_topo_UNIV continuous_on_product_coordinates by fastforce
   then show ?thesis unfolding comp_def by simp
 qed
 
-lemma%unimportant continuous_on_strong_operator_topo_iff_coordinatewise:
+lemma continuous_on_strong_operator_topo_iff_coordinatewise:
   "continuous_on_topo T strong_operator_topology f
     \<longleftrightarrow> (\<forall>x. continuous_on_topo T euclidean (\<lambda>y. blinfun_apply (f y) x))"
 proof (auto)
@@ -1190,9 +1190,9 @@
     by (auto simp add: \<open>continuous_on_topo T euclidean (blinfun_apply o f)\<close>)
 qed
 
-lemma%important strong_operator_topology_weaker_than_euclidean:
+lemma strong_operator_topology_weaker_than_euclidean:
   "continuous_on_topo euclidean strong_operator_topology (\<lambda>f. f)"
-by%unimportant (subst continuous_on_strong_operator_topo_iff_coordinatewise,
+  by (subst continuous_on_strong_operator_topo_iff_coordinatewise,
     auto simp add: continuous_on_topo_UNIV[symmetric] linear_continuous_on)
 
 
@@ -1226,9 +1226,9 @@
 spaces. It would thus be desirable to hide them once the instance is proved, but I do not know how
 to do this.\<close>
 
-lemma%important dist_fun_le_dist_first_terms:
+lemma dist_fun_le_dist_first_terms:
   "dist x y \<le> 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} + (1/2)^N"
-proof%unimportant -
+proof -
   have "(\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)
           = (\<Sum>n. (1 / 2) ^ (Suc N) * ((1/2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1))"
     by (rule suminf_cong, simp add: power_add)
@@ -1276,7 +1276,7 @@
   finally show ?thesis unfolding M_def by simp
 qed
 
-lemma%unimportant open_fun_contains_ball_aux:
+lemma open_fun_contains_ball_aux:
   assumes "open (U::(('a \<Rightarrow> 'b) set))"
           "x \<in> U"
   shows "\<exists>e>0. {y. dist x y < e} \<subseteq> U"
@@ -1339,7 +1339,7 @@
   then show "\<exists>m>0. {y. dist x y < m} \<subseteq> U" using * by blast
 qed
 
-lemma%unimportant ball_fun_contains_open_aux:
+lemma ball_fun_contains_open_aux:
   fixes x::"('a \<Rightarrow> 'b)" and e::real
   assumes "e>0"
   shows "\<exists>U. open U \<and> x \<in> U \<and> U \<subseteq> {y. dist x y < e}"
@@ -1386,7 +1386,7 @@
   ultimately show ?thesis by auto
 qed
 
-lemma%unimportant fun_open_ball_aux:
+lemma fun_open_ball_aux:
   fixes U::"('a \<Rightarrow> 'b) set"
   shows "open U \<longleftrightarrow> (\<forall>x\<in>U. \<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U)"
 proof (auto)
@@ -1575,7 +1575,7 @@
 
 
 
-subsection%important \<open>Measurability\<close> (*FIX ME mv *)
+subsection%important \<open>Measurability\<close> (*FIX ME move? *)
 
 text \<open>There are two natural sigma-algebras on a product space: the borel sigma algebra,
 generated by open sets in the product, and the product sigma algebra, countably generated by
@@ -1593,11 +1593,11 @@
 compare it with the product sigma algebra as explained above.
 \<close>
 
-lemma%unimportant measurable_product_coordinates [measurable (raw)]:
+lemma measurable_product_coordinates [measurable (raw)]:
   "(\<lambda>x. x i) \<in> measurable borel borel"
 by (rule borel_measurable_continuous_on1[OF continuous_on_product_coordinates])
 
-lemma%unimportant measurable_product_then_coordinatewise:
+lemma measurable_product_then_coordinatewise:
   fixes f::"'a \<Rightarrow> 'b \<Rightarrow> ('c::topological_space)"
   assumes [measurable]: "f \<in> borel_measurable M"
   shows "(\<lambda>x. f x i) \<in> borel_measurable M"
@@ -1611,10 +1611,10 @@
 of the product sigma algebra that is more similar to the one we used above for the product
 topology.\<close>
 
-lemma%important sets_PiM_finite:
+lemma sets_PiM_finite:
   "sets (Pi\<^sub>M I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i))
         {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}"
-proof%unimportant
+proof
   have "{(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}} \<subseteq> sets (Pi\<^sub>M I M)"
   proof (auto)
     fix X assume H: "\<forall>i. X i \<in> sets (M i)" "finite {i. X i \<noteq> space (M i)}"
@@ -1654,9 +1654,9 @@
     done
 qed
 
-lemma%important sets_PiM_subset_borel:
+lemma sets_PiM_subset_borel:
   "sets (Pi\<^sub>M UNIV (\<lambda>_. borel)) \<subseteq> sets borel"
-proof%unimportant -
+proof -
   have *: "Pi\<^sub>E UNIV X \<in> sets borel" if [measurable]: "\<And>i. X i \<in> sets borel" "finite {i. X i \<noteq> UNIV}" for X::"'a \<Rightarrow> 'b set"
   proof -
     define I where "I = {i. X i \<noteq> UNIV}"
@@ -1673,9 +1673,9 @@
     by (simp add: * sets.sigma_sets_subset')
 qed
 
-lemma%important sets_PiM_equal_borel:
+proposition sets_PiM_equal_borel:
   "sets (Pi\<^sub>M UNIV (\<lambda>i::('a::countable). borel::('b::second_countable_topology measure))) = sets borel"
-proof%unimportant
+proof
   obtain K::"('a \<Rightarrow> 'b) set set" where K: "topological_basis K" "countable K"
             "\<And>k. k \<in> K \<Longrightarrow> \<exists>X. (k = Pi\<^sub>E UNIV X) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV}"
     using product_topology_countable_basis by fast
@@ -1700,11 +1700,11 @@
     unfolding borel_def by auto
 qed (simp add: sets_PiM_subset_borel)
 
-lemma%important measurable_coordinatewise_then_product:
+lemma measurable_coordinatewise_then_product:
   fixes f::"'a \<Rightarrow> ('b::countable) \<Rightarrow> ('c::second_countable_topology)"
   assumes [measurable]: "\<And>i. (\<lambda>x. f x i) \<in> borel_measurable M"
   shows "f \<in> borel_measurable M"
-proof%unimportant -
+proof -
   have "f \<in> measurable M (Pi\<^sub>M UNIV (\<lambda>_. borel))"
     by (rule measurable_PiM_single', auto simp add: assms)
   then show ?thesis using sets_PiM_equal_borel measurable_cong_sets by blast
--- a/src/HOL/Analysis/Further_Topology.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Further_Topology.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -1,4 +1,4 @@
-section%important \<open>Extending Continous Maps, Invariance of Domain, etc\<close>
+section%important \<open>Extending Continous Maps, Invariance of Domain, etc\<close> (*FIX rename? *)
 
 text\<open>Ported from HOL Light (moretop.ml) by L C Paulson\<close>
 
@@ -8,13 +8,13 @@
 
 subsection%important\<open>A map from a sphere to a higher dimensional sphere is nullhomotopic\<close>
 
-lemma%important spheremap_lemma1:
+lemma spheremap_lemma1:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
   assumes "subspace S" "subspace T" and dimST: "dim S < dim T"
       and "S \<subseteq> T"
       and diff_f: "f differentiable_on sphere 0 1 \<inter> S"
     shows "f ` (sphere 0 1 \<inter> S) \<noteq> sphere 0 1 \<inter> T"
-proof%unimportant
+proof
   assume fim: "f ` (sphere 0 1 \<inter> S) = sphere 0 1 \<inter> T"
   have inS: "\<And>x. \<lbrakk>x \<in> S; x \<noteq> 0\<rbrakk> \<Longrightarrow> (x /\<^sub>R norm x) \<in> S"
     using subspace_mul \<open>subspace S\<close> by blast
@@ -158,14 +158,14 @@
 qed
 
 
-lemma%important spheremap_lemma2:
+lemma spheremap_lemma2:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
   assumes ST: "subspace S" "subspace T" "dim S < dim T"
       and "S \<subseteq> T"
       and contf: "continuous_on (sphere 0 1 \<inter> S) f"
       and fim: "f ` (sphere 0 1 \<inter> S) \<subseteq> sphere 0 1 \<inter> T"
     shows "\<exists>c. homotopic_with (\<lambda>x. True) (sphere 0 1 \<inter> S) (sphere 0 1 \<inter> T) f (\<lambda>x. c)"
-proof%unimportant -
+proof -
   have [simp]: "\<And>x. \<lbrakk>norm x = 1; x \<in> S\<rbrakk> \<Longrightarrow> norm (f x) = 1"
     using fim by (simp add: image_subset_iff)
   have "compact (sphere 0 1 \<inter> S)"
@@ -252,11 +252,11 @@
 qed
 
 
-lemma%important spheremap_lemma3:
+lemma spheremap_lemma3:
   assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \<le> dim U"
   obtains T where "subspace T" "T \<subseteq> U" "S \<noteq> {} \<Longrightarrow> aff_dim T = aff_dim S"
                   "(rel_frontier S) homeomorphic (sphere 0 1 \<inter> T)"
-proof%unimportant (cases "S = {}")
+proof (cases "S = {}")
   case True
   with \<open>subspace U\<close> subspace_0 show ?thesis
     by (rule_tac T = "{0}" in that) auto
@@ -300,14 +300,14 @@
 qed
 
 
-proposition%important inessential_spheremap_lowdim_gen:
+proposition inessential_spheremap_lowdim_gen:
   fixes f :: "'M::euclidean_space \<Rightarrow> 'a::euclidean_space"
   assumes "convex S" "bounded S" "convex T" "bounded T"
       and affST: "aff_dim S < aff_dim T"
       and contf: "continuous_on (rel_frontier S) f"
       and fim: "f ` (rel_frontier S) \<subseteq> rel_frontier T"
   obtains c where "homotopic_with (\<lambda>z. True) (rel_frontier S) (rel_frontier T) f (\<lambda>x. c)"
-proof%unimportant (cases "S = {}")
+proof (cases "S = {}")
   case True
   then show ?thesis
     by (simp add: that)
@@ -350,7 +350,7 @@
   qed
 qed
 
-lemma%unimportant inessential_spheremap_lowdim:
+lemma inessential_spheremap_lowdim:
   fixes f :: "'M::euclidean_space \<Rightarrow> 'a::euclidean_space"
   assumes
    "DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \<subseteq> (sphere b s)"
@@ -379,7 +379,7 @@
 
 subsection%important\<open> Some technical lemmas about extending maps from cell complexes\<close>
 
-lemma%unimportant extending_maps_Union_aux:
+lemma extending_maps_Union_aux:
   assumes fin: "finite \<F>"
       and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
       and "\<And>S T. \<lbrakk>S \<in> \<F>; T \<in> \<F>; S \<noteq> T\<rbrakk> \<Longrightarrow> S \<inter> T \<subseteq> K"
@@ -410,7 +410,7 @@
     done
 qed
 
-lemma%unimportant extending_maps_Union:
+lemma extending_maps_Union:
   assumes fin: "finite \<F>"
       and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g ` S \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
       and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
@@ -422,14 +422,14 @@
 by (metis K psubsetI)
 
 
-lemma%important extend_map_lemma:
+lemma extend_map_lemma:
   assumes "finite \<F>" "\<G> \<subseteq> \<F>" "convex T" "bounded T"
       and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
       and aff: "\<And>X. X \<in> \<F> - \<G> \<Longrightarrow> aff_dim X < aff_dim T"
       and face: "\<And>S T. \<lbrakk>S \<in> \<F>; T \<in> \<F>\<rbrakk> \<Longrightarrow> (S \<inter> T) face_of S \<and> (S \<inter> T) face_of T"
       and contf: "continuous_on (\<Union>\<G>) f" and fim: "f ` (\<Union>\<G>) \<subseteq> rel_frontier T"
   obtains g where "continuous_on (\<Union>\<F>) g" "g ` (\<Union>\<F>) \<subseteq> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"
-proof%unimportant (cases "\<F> - \<G> = {}")
+proof (cases "\<F> - \<G> = {}")
   case True
   then have "\<Union>\<F> \<subseteq> \<Union>\<G>"
     by (simp add: Union_mono)
@@ -608,7 +608,7 @@
     using extendf [of i] unfolding eq by (metis that)
 qed
 
-lemma%unimportant extend_map_lemma_cofinite0:
+lemma extend_map_lemma_cofinite0:
   assumes "finite \<F>"
       and "pairwise (\<lambda>S T. S \<inter> T \<subseteq> K) \<F>"
       and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g ` (S - {a}) \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
@@ -665,7 +665,7 @@
 qed
 
 
-lemma%unimportant extend_map_lemma_cofinite1:
+lemma extend_map_lemma_cofinite1:
 assumes "finite \<F>"
     and \<F>: "\<And>X. X \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (X - {a}) g \<and> g ` (X - {a}) \<subseteq> T \<and> (\<forall>x \<in> X \<inter> K. g x = h x)"
     and clo: "\<And>X. X \<in> \<F> \<Longrightarrow> closed X"
@@ -695,7 +695,7 @@
 qed
 
 
-lemma%important extend_map_lemma_cofinite:
+lemma extend_map_lemma_cofinite:
   assumes "finite \<F>" "\<G> \<subseteq> \<F>" and T: "convex T" "bounded T"
       and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
       and contf: "continuous_on (\<Union>\<G>) f" and fim: "f ` (\<Union>\<G>) \<subseteq> rel_frontier T"
@@ -704,7 +704,7 @@
   obtains C g where
      "finite C" "disjnt C (\<Union>\<G>)" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"
      "g ` (\<Union> \<F> - C) \<subseteq> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"
-proof%unimportant -
+proof -
   define \<H> where "\<H> \<equiv> \<G> \<union> {D. \<exists>C \<in> \<F> - \<G>. D face_of C \<and> aff_dim D < aff_dim T}"
   have "finite \<G>"
     using assms finite_subset by blast
@@ -869,7 +869,7 @@
 qed
 
 text\<open>The next two proofs are similar\<close>
-theorem%important extend_map_cell_complex_to_sphere:
+theorem extend_map_cell_complex_to_sphere:
   assumes "finite \<F>" and S: "S \<subseteq> \<Union>\<F>" "closed S" and T: "convex T" "bounded T"
       and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
       and aff: "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X < aff_dim T"
@@ -877,7 +877,7 @@
       and contf: "continuous_on S f" and fim: "f ` S \<subseteq> rel_frontier T"
   obtains g where "continuous_on (\<Union>\<F>) g"
      "g ` (\<Union>\<F>) \<subseteq> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-proof%unimportant -
+proof -
   obtain V g where "S \<subseteq> V" "open V" "continuous_on V g" and gim: "g ` V \<subseteq> rel_frontier T" and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
     using neighbourhood_extension_into_ANR [OF contf fim _ \<open>closed S\<close>] ANR_rel_frontier_convex T by blast
   have "compact S"
@@ -924,7 +924,7 @@
 qed
 
 
-theorem%important extend_map_cell_complex_to_sphere_cofinite:
+theorem extend_map_cell_complex_to_sphere_cofinite:
   assumes "finite \<F>" and S: "S \<subseteq> \<Union>\<F>" "closed S" and T: "convex T" "bounded T"
       and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
       and aff: "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X \<le> aff_dim T"
@@ -932,7 +932,7 @@
       and contf: "continuous_on S f" and fim: "f ` S \<subseteq> rel_frontier T"
   obtains C g where "finite C" "disjnt C S" "continuous_on (\<Union>\<F> - C) g"
      "g ` (\<Union>\<F> - C) \<subseteq> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-proof%unimportant -
+proof -
   obtain V g where "S \<subseteq> V" "open V" "continuous_on V g" and gim: "g ` V \<subseteq> rel_frontier T" and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
     using neighbourhood_extension_into_ANR [OF contf fim _ \<open>closed S\<close>] ANR_rel_frontier_convex T by blast
   have "compact S"
@@ -993,10 +993,10 @@
 
 subsection%important\<open> Special cases and corollaries involving spheres\<close>
 
-lemma%unimportant disjnt_Diff1: "X \<subseteq> Y' \<Longrightarrow> disjnt (X - Y) (X' - Y')"
+lemma disjnt_Diff1: "X \<subseteq> Y' \<Longrightarrow> disjnt (X - Y) (X' - Y')"
   by (auto simp: disjnt_def)
 
-proposition%important extend_map_affine_to_sphere_cofinite_simple:
+proposition extend_map_affine_to_sphere_cofinite_simple:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "compact S" "convex U" "bounded U"
       and aff: "aff_dim T \<le> aff_dim U"
@@ -1005,7 +1005,7 @@
  obtains K g where "finite K" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"
                    "g ` (T - K) \<subseteq> rel_frontier U"
                    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-proof%unimportant -
+proof -
   have "\<exists>K g. finite K \<and> disjnt K S \<and> continuous_on (T - K) g \<and>
               g ` (T - K) \<subseteq> rel_frontier U \<and> (\<forall>x \<in> S. g x = f x)"
        if "affine T" "S \<subseteq> T" and aff: "aff_dim T \<le> aff_dim U"  for T
@@ -1144,14 +1144,14 @@
 
 (*Up to extend_map_affine_to_sphere_cofinite_gen*)
 
-lemma%important extend_map_affine_to_sphere1:
+lemma extend_map_affine_to_sphere1:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::topological_space"
   assumes "finite K" "affine U" and contf: "continuous_on (U - K) f"
       and fim: "f ` (U - K) \<subseteq> T"
       and comps: "\<And>C. \<lbrakk>C \<in> components(U - S); C \<inter> K \<noteq> {}\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
       and clo: "closedin (subtopology euclidean U) S" and K: "disjnt K S" "K \<subseteq> U"
   obtains g where "continuous_on (U - L) g" "g ` (U - L) \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-proof%unimportant (cases "K = {}")
+proof (cases "K = {}")
   case True
   then show ?thesis
     by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that)
@@ -1436,7 +1436,7 @@
 qed
 
 
-lemma%important extend_map_affine_to_sphere2:
+lemma extend_map_affine_to_sphere2:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "compact S" "convex U" "bounded U" "affine T" "S \<subseteq> T"
       and affTU: "aff_dim T \<le> aff_dim U"
@@ -1446,7 +1446,7 @@
     obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S"
                       "continuous_on (T - K) g" "g ` (T - K) \<subseteq> rel_frontier U"
                       "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-proof%unimportant -
+proof -
   obtain K g where K: "finite K" "K \<subseteq> T" "disjnt K S"
                and contg: "continuous_on (T - K) g"
                and gim: "g ` (T - K) \<subseteq> rel_frontier U"
@@ -1490,7 +1490,7 @@
 qed
 
 
-proposition%important extend_map_affine_to_sphere_cofinite_gen:
+proposition extend_map_affine_to_sphere_cofinite_gen:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \<subseteq> T"
       and aff: "aff_dim T \<le> aff_dim U"
@@ -1500,7 +1500,7 @@
  obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"
                    "g ` (T - K) \<subseteq> rel_frontier U"
                    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-proof%unimportant (cases "S = {}")
+proof (cases "S = {}")
   case True
   show ?thesis
   proof (cases "rel_frontier U = {}")
@@ -1645,7 +1645,7 @@
 qed
 
 
-corollary%important extend_map_affine_to_sphere_cofinite:
+corollary extend_map_affine_to_sphere_cofinite:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes SUT: "compact S" "affine T" "S \<subseteq> T"
       and aff: "aff_dim T \<le> DIM('b)" and "0 \<le> r"
@@ -1654,7 +1654,7 @@
       and dis: "\<And>C. \<lbrakk>C \<in> components(T - S); bounded C\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
   obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"
                     "g ` (T - K) \<subseteq> sphere a r" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-proof%unimportant (cases "r = 0")
+proof (cases "r = 0")
   case True
   with fim show ?thesis
     by (rule_tac K="{}" and g = "\<lambda>x. a" in that) (auto simp: continuous_on_const)
@@ -1670,7 +1670,7 @@
     done
 qed
 
-corollary%important extend_map_UNIV_to_sphere_cofinite:
+corollary extend_map_UNIV_to_sphere_cofinite:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes aff: "DIM('a) \<le> DIM('b)" and "0 \<le> r"
       and SUT: "compact S"
@@ -1684,7 +1684,7 @@
  apply (auto simp: assms that Compl_eq_Diff_UNIV [symmetric])
 done
 
-corollary%important extend_map_UNIV_to_sphere_no_bounded_component:
+corollary extend_map_UNIV_to_sphere_no_bounded_component:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes aff: "DIM('a) \<le> DIM('b)" and "0 \<le> r"
       and SUT: "compact S"
@@ -1696,14 +1696,14 @@
    apply (auto simp: that dest: dis)
 done
 
-theorem%important Borsuk_separation_theorem_gen:
+theorem Borsuk_separation_theorem_gen:
   fixes S :: "'a::euclidean_space set"
   assumes "compact S"
     shows "(\<forall>c \<in> components(- S). \<not>bounded c) \<longleftrightarrow>
            (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::'a) 1
                 \<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)))"
        (is "?lhs = ?rhs")
-proof%unimportant
+proof
   assume L [rule_format]: ?lhs
   show ?rhs
   proof clarify
@@ -1734,14 +1734,14 @@
 qed
 
 
-corollary%important Borsuk_separation_theorem:
+corollary Borsuk_separation_theorem:
   fixes S :: "'a::euclidean_space set"
   assumes "compact S" and 2: "2 \<le> DIM('a)"
     shows "connected(- S) \<longleftrightarrow>
            (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::'a) 1
                 \<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)))"
        (is "?lhs = ?rhs")
-proof%unimportant
+proof
   assume L: ?lhs
   show ?rhs
   proof (cases "S = {}")
@@ -1764,7 +1764,7 @@
 qed
 
 
-lemma%unimportant homotopy_eqv_separation:
+lemma homotopy_eqv_separation:
   fixes S :: "'a::euclidean_space set" and T :: "'a set"
   assumes "S homotopy_eqv T" and "compact S" and "compact T"
   shows "connected(- S) \<longleftrightarrow> connected(- T)"
@@ -1783,11 +1783,11 @@
   qed
 qed
 
-lemma%important Jordan_Brouwer_separation:
+proposition Jordan_Brouwer_separation:
   fixes S :: "'a::euclidean_space set" and a::'a
   assumes hom: "S homeomorphic sphere a r" and "0 < r"
     shows "\<not> connected(- S)"
-proof%unimportant -
+proof -
   have "- sphere a r \<inter> ball a r \<noteq> {}"
     using \<open>0 < r\<close> by (simp add: Int_absorb1 subset_eq)
   moreover
@@ -1817,11 +1817,11 @@
 qed
 
 
-lemma%important Jordan_Brouwer_frontier:
+proposition Jordan_Brouwer_frontier:
   fixes S :: "'a::euclidean_space set" and a::'a
   assumes S: "S homeomorphic sphere a r" and T: "T \<in> components(- S)" and 2: "2 \<le> DIM('a)"
     shows "frontier T = S"
-proof%unimportant (cases r rule: linorder_cases)
+proof (cases r rule: linorder_cases)
   assume "r < 0"
   with S T show ?thesis by auto
 next
@@ -1866,11 +1866,11 @@
   qed (rule T)
 qed
 
-lemma%important Jordan_Brouwer_nonseparation:
+proposition Jordan_Brouwer_nonseparation:
   fixes S :: "'a::euclidean_space set" and a::'a
   assumes S: "S homeomorphic sphere a r" and "T \<subset> S" and 2: "2 \<le> DIM('a)"
     shows "connected(- T)"
-proof%unimportant -
+proof -
   have *: "connected(C \<union> (S - T))" if "C \<in> components(- S)" for C
   proof (rule connected_intermediate_closure)
     show "connected C"
@@ -1894,7 +1894,7 @@
 
 subsection%important\<open> Invariance of domain and corollaries\<close>
 
-lemma%unimportant invariance_of_domain_ball:
+lemma invariance_of_domain_ball:
   fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
   assumes contf: "continuous_on (cball a r) f" and "0 < r"
      and inj: "inj_on f (cball a r)"
@@ -1983,12 +1983,12 @@
 
 
 text\<open>Proved by L. E. J. Brouwer (1912)\<close>
-theorem%important invariance_of_domain:
+theorem invariance_of_domain:
   fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
   assumes "continuous_on S f" "open S" "inj_on f S"
     shows "open(f ` S)"
   unfolding open_subopen [of "f`S"]
-proof%unimportant clarify
+proof clarify
   fix a
   assume "a \<in> S"
   obtain \<delta> where "\<delta> > 0" and \<delta>: "cball a \<delta> \<subseteq> S"
@@ -2004,7 +2004,7 @@
   qed
 qed
 
-lemma%unimportant inv_of_domain_ss0:
+lemma inv_of_domain_ss0:
   fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
   assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S"
       and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)"
@@ -2049,7 +2049,7 @@
     by (metis (no_types, hide_lams) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV)
 qed
 
-lemma%unimportant inv_of_domain_ss1:
+lemma inv_of_domain_ss1:
   fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
   assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S"
       and "subspace S"
@@ -2090,14 +2090,14 @@
 qed
 
 
-corollary%important invariance_of_domain_subspaces:
+corollary invariance_of_domain_subspaces:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes ope: "openin (subtopology euclidean U) S"
       and "subspace U" "subspace V" and VU: "dim V \<le> dim U"
       and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
       and injf: "inj_on f S"
     shows "openin (subtopology euclidean V) (f ` S)"
-proof%unimportant -
+proof -
   obtain V' where "subspace V'" "V' \<subseteq> U" "dim V' = dim V"
     using choose_subspace_of_subspace [OF VU]
     by (metis span_eq_iff \<open>subspace U\<close>)
@@ -2134,14 +2134,14 @@
   qed
 qed
 
-corollary%important invariance_of_dimension_subspaces:
+corollary invariance_of_dimension_subspaces:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes ope: "openin (subtopology euclidean U) S"
       and "subspace U" "subspace V"
       and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
       and injf: "inj_on f S" and "S \<noteq> {}"
     shows "dim U \<le> dim V"
-proof%unimportant -
+proof -
   have "False" if "dim V < dim U"
   proof -
     obtain T where "subspace T" "T \<subseteq> U" "dim T = dim V"
@@ -2174,14 +2174,14 @@
     using not_less by blast
 qed
 
-corollary%important invariance_of_domain_affine_sets:
+corollary invariance_of_domain_affine_sets:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes ope: "openin (subtopology euclidean U) S"
       and aff: "affine U" "affine V" "aff_dim V \<le> aff_dim U"
       and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
       and injf: "inj_on f S"
     shows "openin (subtopology euclidean V) (f ` S)"
-proof%unimportant (cases "S = {}")
+proof (cases "S = {}")
   case True
   then show ?thesis by auto
 next
@@ -2209,14 +2209,14 @@
     by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois)
 qed
 
-corollary%important invariance_of_dimension_affine_sets:
+corollary invariance_of_dimension_affine_sets:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes ope: "openin (subtopology euclidean U) S"
       and aff: "affine U" "affine V"
       and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
       and injf: "inj_on f S" and "S \<noteq> {}"
     shows "aff_dim U \<le> aff_dim V"
-proof%unimportant -
+proof -
   obtain a b where "a \<in> S" "a \<in> U" "b \<in> V"
     using \<open>S \<noteq> {}\<close> fim ope openin_contains_cball by fastforce
   have "dim ((+) (- a) ` U) \<le> dim ((+) (- b) ` V)"
@@ -2238,7 +2238,7 @@
     by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff)
 qed
 
-corollary%important invariance_of_dimension:
+corollary invariance_of_dimension:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes contf: "continuous_on S f" and "open S"
       and injf: "inj_on f S" and "S \<noteq> {}"
@@ -2247,7 +2247,7 @@
   by auto
 
 
-corollary%important continuous_injective_image_subspace_dim_le:
+corollary continuous_injective_image_subspace_dim_le:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "subspace S" "subspace T"
       and contf: "continuous_on S f" and fim: "f ` S \<subseteq> T"
@@ -2256,7 +2256,7 @@
   apply (rule invariance_of_dimension_subspaces [of S S _ f])
   using%unimportant assms by (auto simp: subspace_affine)
 
-lemma%unimportant invariance_of_dimension_convex_domain:
+lemma invariance_of_dimension_convex_domain:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "convex S"
       and contf: "continuous_on S f" and fim: "f ` S \<subseteq> affine hull T"
@@ -2285,7 +2285,7 @@
 qed
 
 
-lemma%unimportant homeomorphic_convex_sets_le:
+lemma homeomorphic_convex_sets_le:
   assumes "convex S" "S homeomorphic T"
   shows "aff_dim S \<le> aff_dim T"
 proof -
@@ -2302,23 +2302,23 @@
   qed
 qed
 
-lemma%unimportant homeomorphic_convex_sets:
+lemma homeomorphic_convex_sets:
   assumes "convex S" "convex T" "S homeomorphic T"
   shows "aff_dim S = aff_dim T"
   by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym)
 
-lemma%unimportant homeomorphic_convex_compact_sets_eq:
+lemma homeomorphic_convex_compact_sets_eq:
   assumes "convex S" "compact S" "convex T" "compact T"
   shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T"
   by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets)
 
-lemma%unimportant invariance_of_domain_gen:
+lemma invariance_of_domain_gen:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
     shows "open(f ` S)"
   using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto
 
-lemma%unimportant injective_into_1d_imp_open_map_UNIV:
+lemma injective_into_1d_imp_open_map_UNIV:
   fixes f :: "'a::euclidean_space \<Rightarrow> real"
   assumes "open T" "continuous_on S f" "inj_on f S" "T \<subseteq> S"
     shows "open (f ` T)"
@@ -2326,7 +2326,7 @@
   using assms apply (auto simp: elim: continuous_on_subset subset_inj_on)
   done
 
-lemma%unimportant continuous_on_inverse_open:
+lemma continuous_on_inverse_open:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" and gf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
     shows "continuous_on (f ` S) g"
@@ -2345,7 +2345,7 @@
     by (metis \<open>open T\<close> continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq)
 qed
 
-lemma%unimportant invariance_of_domain_homeomorphism:
+lemma invariance_of_domain_homeomorphism:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
   obtains g where "homeomorphism S (f ` S) f g"
@@ -2354,14 +2354,14 @@
     by (simp add: assms continuous_on_inverse_open homeomorphism_def)
 qed
 
-corollary%important invariance_of_domain_homeomorphic:
+corollary invariance_of_domain_homeomorphic:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
   shows "S homeomorphic (f ` S)"
   using%unimportant invariance_of_domain_homeomorphism [OF assms]
   by%unimportant (meson homeomorphic_def)
 
-lemma%unimportant continuous_image_subset_interior:
+lemma continuous_image_subset_interior:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
   shows "f ` (interior S) \<subseteq> interior(f ` S)"
@@ -2372,13 +2372,13 @@
      apply (auto simp: subset_inj_on interior_subset continuous_on_subset)
   done
 
-lemma%important homeomorphic_interiors_same_dimension:
+lemma homeomorphic_interiors_same_dimension:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)"
   shows "(interior S) homeomorphic (interior T)"
   using assms [unfolded homeomorphic_minimal]
   unfolding homeomorphic_def
-proof%unimportant (clarify elim!: ex_forward)
+proof (clarify elim!: ex_forward)
   fix f g
   assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
      and contf: "continuous_on S f" and contg: "continuous_on T g"
@@ -2422,7 +2422,7 @@
   qed
 qed
 
-lemma%unimportant homeomorphic_open_imp_same_dimension:
+lemma homeomorphic_open_imp_same_dimension:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" "open S" "S \<noteq> {}" "open T" "T \<noteq> {}"
   shows "DIM('a) = DIM('b)"
@@ -2431,7 +2431,7 @@
     apply (rule order_antisym; metis inj_onI invariance_of_dimension)
     done
 
-lemma%unimportant homeomorphic_interiors:
+proposition homeomorphic_interiors:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" "interior S = {} \<longleftrightarrow> interior T = {}"
     shows "(interior S) homeomorphic (interior T)"
@@ -2449,7 +2449,7 @@
     by (rule homeomorphic_interiors_same_dimension [OF \<open>S homeomorphic T\<close>])
 qed
 
-lemma%unimportant homeomorphic_frontiers_same_dimension:
+lemma homeomorphic_frontiers_same_dimension:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)"
   shows "(frontier S) homeomorphic (frontier T)"
@@ -2505,7 +2505,7 @@
   qed
 qed
 
-lemma%unimportant homeomorphic_frontiers:
+lemma homeomorphic_frontiers:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" "closed S" "closed T"
           "interior S = {} \<longleftrightarrow> interior T = {}"
@@ -2522,7 +2522,7 @@
     using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast
 qed
 
-lemma%unimportant continuous_image_subset_rel_interior:
+lemma continuous_image_subset_rel_interior:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
       and TS: "aff_dim T \<le> aff_dim S"
@@ -2545,7 +2545,7 @@
   qed auto
 qed
 
-lemma%unimportant homeomorphic_rel_interiors_same_dimension:
+lemma homeomorphic_rel_interiors_same_dimension:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T"
   shows "(rel_interior S) homeomorphic (rel_interior T)"
@@ -2597,11 +2597,11 @@
   qed
 qed
 
-lemma%important homeomorphic_rel_interiors:
+lemma homeomorphic_rel_interiors:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" "rel_interior S = {} \<longleftrightarrow> rel_interior T = {}"
     shows "(rel_interior S) homeomorphic (rel_interior T)"
-proof%unimportant (cases "rel_interior T = {}")
+proof (cases "rel_interior T = {}")
   case True
   with assms show ?thesis by auto
 next
@@ -2630,7 +2630,7 @@
 qed
 
 
-lemma%unimportant homeomorphic_rel_boundaries_same_dimension:
+lemma homeomorphic_rel_boundaries_same_dimension:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T"
   shows "(S - rel_interior S) homeomorphic (T - rel_interior T)"
@@ -2664,7 +2664,7 @@
   qed
 qed
 
-lemma%unimportant homeomorphic_rel_boundaries:
+lemma homeomorphic_rel_boundaries:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "S homeomorphic T" "rel_interior S = {} \<longleftrightarrow> rel_interior T = {}"
     shows "(S - rel_interior S) homeomorphic (T - rel_interior T)"
@@ -2696,11 +2696,11 @@
     by (rule homeomorphic_rel_boundaries_same_dimension [OF \<open>S homeomorphic T\<close>])
 qed
 
-proposition%important uniformly_continuous_homeomorphism_UNIV_trivial:
+proposition uniformly_continuous_homeomorphism_UNIV_trivial:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
   assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g"
   shows "S = UNIV"
-proof%unimportant (cases "S = {}")
+proof (cases "S = {}")
   case True
   then show ?thesis
     by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI)
@@ -2744,7 +2744,7 @@
 
 subsection%important\<open>Dimension-based conditions for various homeomorphisms\<close>
 
-lemma%unimportant homeomorphic_subspaces_eq:
+lemma homeomorphic_subspaces_eq:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "subspace S" "subspace T"
   shows "S homeomorphic T \<longleftrightarrow> dim S = dim T"
@@ -2765,7 +2765,7 @@
     by (simp add: assms homeomorphic_subspaces)
 qed
 
-lemma%unimportant homeomorphic_affine_sets_eq:
+lemma homeomorphic_affine_sets_eq:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "affine S" "affine T"
   shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T"
@@ -2783,19 +2783,19 @@
     by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets)
 qed
 
-lemma%unimportant homeomorphic_hyperplanes_eq:
+lemma homeomorphic_hyperplanes_eq:
   fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space"
   assumes "a \<noteq> 0" "c \<noteq> 0"
   shows "({x. a \<bullet> x = b} homeomorphic {x. c \<bullet> x = d} \<longleftrightarrow> DIM('a) = DIM('b))"
   apply (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms)
   by (metis DIM_positive Suc_pred)
 
-lemma%unimportant homeomorphic_UNIV_UNIV:
+lemma homeomorphic_UNIV_UNIV:
   shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \<longleftrightarrow>
     DIM('a::euclidean_space) = DIM('b::euclidean_space)"
   by (simp add: homeomorphic_subspaces_eq)
 
-lemma%unimportant simply_connected_sphere_gen:
+lemma simply_connected_sphere_gen:
    assumes "convex S" "bounded S" and 3: "3 \<le> aff_dim S"
    shows "simply_connected(rel_frontier S)"
 proof -
@@ -2819,16 +2819,16 @@
   qed
 qed
 
-subsection%important\<open>more invariance of domain\<close>
-
-proposition%important invariance_of_domain_sphere_affine_set_gen:
+subsection%important\<open>more invariance of domain\<close>(*FIX ME title? *)
+
+proposition invariance_of_domain_sphere_affine_set_gen:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
       and U: "bounded U" "convex U"
       and "affine T" and affTU: "aff_dim T < aff_dim U"
       and ope: "openin (subtopology euclidean (rel_frontier U)) S"
    shows "openin (subtopology euclidean T) (f ` S)"
-proof%unimportant (cases "rel_frontier U = {}")
+proof (cases "rel_frontier U = {}")
   case True
   then show ?thesis
     using ope openin_subset by force
@@ -2927,7 +2927,7 @@
 qed
 
 
-lemma%unimportant invariance_of_domain_sphere_affine_set:
+lemma invariance_of_domain_sphere_affine_set:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
       and "r \<noteq> 0" "affine T" and affTU: "aff_dim T < DIM('a)"
@@ -2948,7 +2948,7 @@
   qed
 qed
 
-lemma%unimportant no_embedding_sphere_lowdim:
+lemma no_embedding_sphere_lowdim:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0"
    shows "DIM('a) \<le> DIM('b)"
@@ -2969,7 +2969,7 @@
     using not_less by blast
 qed
 
-lemma%unimportant simply_connected_sphere:
+lemma simply_connected_sphere:
   fixes a :: "'a::euclidean_space"
   assumes "3 \<le> DIM('a)"
     shows "simply_connected(sphere a r)"
@@ -2986,7 +2986,7 @@
     by (simp add: aff_dim_cball)
 qed
 
-lemma%unimportant simply_connected_sphere_eq:
+lemma simply_connected_sphere_eq:
   fixes a :: "'a::euclidean_space"
   shows "simply_connected(sphere a r) \<longleftrightarrow> 3 \<le> DIM('a) \<or> r \<le> 0"  (is "?lhs = ?rhs")
 proof (cases "r \<le> 0")
@@ -3029,7 +3029,7 @@
 qed
 
 
-lemma%unimportant simply_connected_punctured_universe_eq:
+lemma simply_connected_punctured_universe_eq:
   fixes a :: "'a::euclidean_space"
   shows "simply_connected(- {a}) \<longleftrightarrow> 3 \<le> DIM('a)"
 proof -
@@ -3047,17 +3047,17 @@
   finally show ?thesis .
 qed
 
-lemma%unimportant not_simply_connected_circle:
+lemma not_simply_connected_circle:
   fixes a :: complex
   shows "0 < r \<Longrightarrow> \<not> simply_connected(sphere a r)"
 by (simp add: simply_connected_sphere_eq)
 
 
-proposition%important simply_connected_punctured_convex:
+proposition simply_connected_punctured_convex:
   fixes a :: "'a::euclidean_space"
   assumes "convex S" and 3: "3 \<le> aff_dim S"
     shows "simply_connected(S - {a})"
-proof%unimportant (cases "a \<in> rel_interior S")
+proof (cases "a \<in> rel_interior S")
   case True
   then obtain e where "a \<in> S" "0 < e" and e: "cball a e \<inter> affine hull S \<subseteq> S"
     by (auto simp: rel_interior_cball)
@@ -3096,11 +3096,11 @@
     by (meson Diff_subset closure_subset subset_trans)
 qed
 
-corollary%important simply_connected_punctured_universe:
+corollary simply_connected_punctured_universe:
   fixes a :: "'a::euclidean_space"
   assumes "3 \<le> DIM('a)"
   shows "simply_connected(- {a})"
-proof%unimportant -
+proof -
   have [simp]: "affine hull cball a 1 = UNIV"
     apply auto
     by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq)
@@ -3116,10 +3116,10 @@
 
 subsection%important\<open>The power, squaring and exponential functions as covering maps\<close>
 
-proposition%important covering_space_power_punctured_plane:
+proposition covering_space_power_punctured_plane:
   assumes "0 < n"
     shows "covering_space (- {0}) (\<lambda>z::complex. z^n) (- {0})"
-proof%unimportant -
+proof -
   consider "n = 1" | "2 \<le> n" using assms by linarith
   then obtain e where "0 < e"
                 and e: "\<And>w z. cmod(w - z) < e * cmod z \<Longrightarrow> (w^n = z^n \<longleftrightarrow> w = z)"
@@ -3367,14 +3367,14 @@
     done
 qed
 
-corollary%important covering_space_square_punctured_plane:
+corollary covering_space_square_punctured_plane:
   "covering_space (- {0}) (\<lambda>z::complex. z^2) (- {0})"
   by%unimportant (simp add: covering_space_power_punctured_plane)
 
 
-proposition%important covering_space_exp_punctured_plane:
+proposition covering_space_exp_punctured_plane:
   "covering_space UNIV (\<lambda>z::complex. exp z) (- {0})"
-proof%unimportant (simp add: covering_space_def, intro conjI ballI)
+proof (simp add: covering_space_def, intro conjI ballI)
   show "continuous_on UNIV (\<lambda>z::complex. exp z)"
     by (rule continuous_on_exp [OF continuous_on_id])
   show "range exp = - {0::complex}"
@@ -3487,9 +3487,9 @@
 qed
 
 
-subsection%important\<open>Hence the Borsukian results about mappings into circles\<close>
-
-lemma%unimportant inessential_eq_continuous_logarithm:
+subsection%important\<open>Hence the Borsukian results about mappings into circles\<close>(*FIX ME title *)
+
+lemma inessential_eq_continuous_logarithm:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   shows "(\<exists>a. homotopic_with (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)) \<longleftrightarrow>
          (\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x)))"
@@ -3513,7 +3513,7 @@
     by (simp add: f homotopic_with_eq)
 qed
 
-corollary%important inessential_imp_continuous_logarithm_circle:
+corollary inessential_imp_continuous_logarithm_circle:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   assumes "homotopic_with (\<lambda>h. True) S (sphere 0 1) f (\<lambda>t. a)"
   obtains g where "continuous_on S g" and "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
@@ -3525,7 +3525,7 @@
 qed
 
 
-lemma%unimportant inessential_eq_continuous_logarithm_circle:
+lemma inessential_eq_continuous_logarithm_circle:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   shows "(\<exists>a. homotopic_with (\<lambda>h. True) S (sphere 0 1) f (\<lambda>t. a)) \<longleftrightarrow>
          (\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(\<i> * of_real(g x))))"
@@ -3565,12 +3565,12 @@
     by (simp add: homotopic_with_eq)
 qed
 
-lemma%important homotopic_with_sphere_times:
+proposition homotopic_with_sphere_times:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   assumes hom: "homotopic_with (\<lambda>x. True) S (sphere 0 1) f g" and conth: "continuous_on S h"
       and hin: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> sphere 0 1"
     shows "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x * h x) (\<lambda>x. g x * h x)"
-proof%unimportant -
+proof -
   obtain k where contk: "continuous_on ({0..1::real} \<times> S) k"
              and kim: "k ` ({0..1} \<times> S) \<subseteq> sphere 0 1"
              and k0:  "\<And>x. k(0, x) = f x"
@@ -3585,14 +3585,13 @@
     done
 qed
 
-
-lemma%important homotopic_circlemaps_divide:
+proposition homotopic_circlemaps_divide:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
     shows "homotopic_with (\<lambda>x. True) S (sphere 0 1) f g \<longleftrightarrow>
            continuous_on S f \<and> f ` S \<subseteq> sphere 0 1 \<and>
            continuous_on S g \<and> g ` S \<subseteq> sphere 0 1 \<and>
            (\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c))"
-proof%unimportant -
+proof -
   have "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"
        if "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c)" for c
   proof -
@@ -3646,7 +3645,7 @@
 
 
 text\<open>Many similar proofs below.\<close>
-lemma%unimportant upper_hemicontinuous:
+lemma upper_hemicontinuous:
   assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
     shows "((\<forall>U. openin (subtopology euclidean T) U
                  \<longrightarrow> openin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
@@ -3677,7 +3676,7 @@
     by (simp add: openin_closedin_eq)
 qed
 
-lemma%unimportant lower_hemicontinuous:
+lemma lower_hemicontinuous:
   assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
     shows "((\<forall>U. closedin (subtopology euclidean T) U
                  \<longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
@@ -3708,7 +3707,7 @@
     by (simp add: openin_closedin_eq)
 qed
 
-lemma%unimportant open_map_iff_lower_hemicontinuous_preimage:
+lemma open_map_iff_lower_hemicontinuous_preimage:
   assumes "f ` S \<subseteq> T"
     shows "((\<forall>U. openin (subtopology euclidean S) U
                  \<longrightarrow> openin (subtopology euclidean T) (f ` U)) \<longleftrightarrow>
@@ -3739,7 +3738,7 @@
     using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq)
 qed
 
-lemma%unimportant closed_map_iff_upper_hemicontinuous_preimage:
+lemma closed_map_iff_upper_hemicontinuous_preimage:
   assumes "f ` S \<subseteq> T"
     shows "((\<forall>U. closedin (subtopology euclidean S) U
                  \<longrightarrow> closedin (subtopology euclidean T) (f ` U)) \<longleftrightarrow>
@@ -3770,7 +3769,7 @@
     by (simp add: openin_closedin_eq)
 qed
 
-proposition%important upper_lower_hemicontinuous_explicit:
+proposition upper_lower_hemicontinuous_explicit:
   fixes T :: "('b::{real_normed_vector,heine_borel}) set"
   assumes fST: "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
       and ope: "\<And>U. openin (subtopology euclidean T) U
@@ -3782,7 +3781,7 @@
              "\<And>x'. \<lbrakk>x' \<in> S; dist x x' < d\<rbrakk>
                            \<Longrightarrow> (\<forall>y \<in> f x. \<exists>y'. y' \<in> f x' \<and> dist y y' < e) \<and>
                                (\<forall>y' \<in> f x'. \<exists>y. y \<in> f x \<and> dist y' y < e)"
-proof%unimportant -
+proof -
   have "openin (subtopology euclidean T) (T \<inter> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b}))"
     by (auto simp: open_sums openin_open_Int)
   with ope have "openin (subtopology euclidean S)
@@ -3843,11 +3842,11 @@
 
 subsection%important\<open>Complex logs exist on various "well-behaved" sets\<close>
 
-lemma%important continuous_logarithm_on_contractible:
+lemma continuous_logarithm_on_contractible:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   assumes "continuous_on S f" "contractible S" "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
   obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
-proof%unimportant -
+proof -
   obtain c where hom: "homotopic_with (\<lambda>h. True) S (-{0}) f (\<lambda>x. c)"
     using nullhomotopic_from_contractible assms
     by (metis imageE subset_Compl_singleton)
@@ -3855,27 +3854,27 @@
     by (metis inessential_eq_continuous_logarithm that)
 qed
 
-lemma%important continuous_logarithm_on_simply_connected:
+lemma continuous_logarithm_on_simply_connected:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S"
       and f: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
   obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
-  using%unimportant covering_space_lift [OF covering_space_exp_punctured_plane S contf]
-  by%unimportant (metis (full_types) f imageE subset_Compl_singleton)
-
-lemma%unimportant continuous_logarithm_on_cball:
+  using covering_space_lift [OF covering_space_exp_punctured_plane S contf]
+  by (metis (full_types) f imageE subset_Compl_singleton)
+
+lemma continuous_logarithm_on_cball:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   assumes "continuous_on (cball a r) f" and "\<And>z. z \<in> cball a r \<Longrightarrow> f z \<noteq> 0"
     obtains h where "continuous_on (cball a r) h" "\<And>z. z \<in> cball a r \<Longrightarrow> f z = exp(h z)"
   using assms continuous_logarithm_on_contractible convex_imp_contractible by blast
 
-lemma%unimportant continuous_logarithm_on_ball:
+lemma continuous_logarithm_on_ball:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   assumes "continuous_on (ball a r) f" and "\<And>z. z \<in> ball a r \<Longrightarrow> f z \<noteq> 0"
   obtains h where "continuous_on (ball a r) h" "\<And>z. z \<in> ball a r \<Longrightarrow> f z = exp(h z)"
   using assms continuous_logarithm_on_contractible convex_imp_contractible by blast
 
-lemma%unimportant continuous_sqrt_on_contractible:
+lemma continuous_sqrt_on_contractible:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   assumes "continuous_on S f" "contractible S"
       and "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
@@ -3892,7 +3891,7 @@
   qed
 qed
 
-lemma%unimportant continuous_sqrt_on_simply_connected:
+lemma continuous_sqrt_on_simply_connected:
   fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
   assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S"
       and f: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
@@ -3912,12 +3911,12 @@
 
 subsection%important\<open>Another simple case where sphere maps are nullhomotopic\<close>
 
-lemma%important inessential_spheremap_2_aux:
+lemma inessential_spheremap_2_aux:
   fixes f :: "'a::euclidean_space \<Rightarrow> complex"
   assumes 2: "2 < DIM('a)" and contf: "continuous_on (sphere a r) f" 
       and fim: "f `(sphere a r) \<subseteq> (sphere 0 1)" 
   obtains c where "homotopic_with (\<lambda>z. True) (sphere a r) (sphere 0 1) f (\<lambda>x. c)"
-proof%unimportant -
+proof -
   obtain g where contg: "continuous_on (sphere a r) g" 
              and feq: "\<And>x. x \<in> sphere a r \<Longrightarrow> f x = exp(g x)"
   proof (rule continuous_logarithm_on_simply_connected [OF contf])
@@ -3939,12 +3938,12 @@
     by metis
 qed
 
-lemma%important inessential_spheremap_2:
+lemma inessential_spheremap_2:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes a2: "2 < DIM('a)" and b2: "DIM('b) = 2" 
       and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \<subseteq> (sphere b s)"
   obtains c where "homotopic_with (\<lambda>z. True) (sphere a r) (sphere b s) f (\<lambda>x. c)"
-proof%unimportant (cases "s \<le> 0")
+proof (cases "s \<le> 0")
   case True
   then show ?thesis
     using contf contractible_sphere fim nullhomotopic_into_contractible that by blast
@@ -3978,12 +3977,12 @@
 
 subsection%important\<open>Holomorphic logarithms and square roots\<close>
 
-lemma%important contractible_imp_holomorphic_log:
+lemma contractible_imp_holomorphic_log:
   assumes holf: "f holomorphic_on S"
       and S: "contractible S"
       and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
   obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"
-proof%unimportant -
+proof -
   have contf: "continuous_on S f"
     by (simp add: holf holomorphic_on_imp_continuous_on)
   obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp (g x)"
@@ -4026,12 +4025,12 @@
 qed
 
 (*Identical proofs*)
-lemma%important simply_connected_imp_holomorphic_log:
+lemma simply_connected_imp_holomorphic_log:
   assumes holf: "f holomorphic_on S"
       and S: "simply_connected S" "locally path_connected S"
       and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
   obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"
-proof%unimportant -
+proof -
   have contf: "continuous_on S f"
     by (simp add: holf holomorphic_on_imp_continuous_on)
   obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp (g x)"
@@ -4074,7 +4073,7 @@
 qed
 
 
-lemma%unimportant contractible_imp_holomorphic_sqrt:
+lemma contractible_imp_holomorphic_sqrt:
   assumes holf: "f holomorphic_on S"
       and S: "contractible S"
       and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
@@ -4092,7 +4091,7 @@
   qed
 qed
 
-lemma%unimportant simply_connected_imp_holomorphic_sqrt:
+lemma simply_connected_imp_holomorphic_sqrt:
   assumes holf: "f holomorphic_on S"
       and S: "simply_connected S" "locally path_connected S"
       and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
@@ -4112,11 +4111,11 @@
 
 text\<open> Related theorems about holomorphic inverse cosines.\<close>
 
-lemma%important contractible_imp_holomorphic_arccos:
+lemma contractible_imp_holomorphic_arccos:
   assumes holf: "f holomorphic_on S" and S: "contractible S"
       and non1: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 1 \<and> f z \<noteq> -1"
   obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = cos(g z)"
-proof%unimportant -
+proof -
   have hol1f: "(\<lambda>z. 1 - f z ^ 2) holomorphic_on S"
     by (intro holomorphic_intros holf)
   obtain g where holg: "g holomorphic_on S" and eq: "\<And>z. z \<in> S \<Longrightarrow> 1 - (f z)\<^sup>2 = (g z)\<^sup>2"
@@ -4150,11 +4149,11 @@
 qed
 
 
-lemma%important contractible_imp_holomorphic_arccos_bounded:
+lemma contractible_imp_holomorphic_arccos_bounded:
   assumes holf: "f holomorphic_on S" and S: "contractible S" and "a \<in> S"
       and non1: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 1 \<and> f z \<noteq> -1"
   obtains g where "g holomorphic_on S" "norm(g a) \<le> pi + norm(f a)" "\<And>z. z \<in> S \<Longrightarrow> f z = cos(g z)"
-proof%unimportant -
+proof -
   obtain g where holg: "g holomorphic_on S" and feq: "\<And>z. z \<in> S \<Longrightarrow> f z = cos (g z)"
     using contractible_imp_holomorphic_arccos [OF holf S non1] by blast
   obtain b where "cos b = f a" "norm b \<le> pi + norm (f a)"
@@ -4200,11 +4199,11 @@
         \<forall>f. continuous_on S f \<and> f ` S \<subseteq> (- {0::complex})
             \<longrightarrow> (\<exists>a. homotopic_with (\<lambda>h. True) S (- {0}) f (\<lambda>x. a))"
 
-lemma%important Borsukian_retraction_gen:
+lemma Borsukian_retraction_gen:
   assumes "Borsukian S" "continuous_on S h" "h ` S = T"
           "continuous_on T k"  "k ` T \<subseteq> S"  "\<And>y. y \<in> T \<Longrightarrow> h(k y) = y"
     shows "Borsukian T"
-proof%unimportant -
+proof -
   interpret R: Retracts S h T k
     using assms by (simp add: Retracts.intro)
   show ?thesis
@@ -4214,42 +4213,42 @@
     done
 qed
 
-lemma%unimportant retract_of_Borsukian: "\<lbrakk>Borsukian T; S retract_of T\<rbrakk> \<Longrightarrow> Borsukian S"
+lemma retract_of_Borsukian: "\<lbrakk>Borsukian T; S retract_of T\<rbrakk> \<Longrightarrow> Borsukian S"
   apply (auto simp: retract_of_def retraction_def)
   apply (erule (1) Borsukian_retraction_gen)
   apply (meson retraction retraction_def)
     apply (auto simp: continuous_on_id)
     done
 
-lemma%unimportant homeomorphic_Borsukian: "\<lbrakk>Borsukian S; S homeomorphic T\<rbrakk> \<Longrightarrow> Borsukian T"
+lemma homeomorphic_Borsukian: "\<lbrakk>Borsukian S; S homeomorphic T\<rbrakk> \<Longrightarrow> Borsukian T"
   using Borsukian_retraction_gen order_refl
   by (fastforce simp add: homeomorphism_def homeomorphic_def)
 
-lemma%unimportant homeomorphic_Borsukian_eq:
+lemma homeomorphic_Borsukian_eq:
    "S homeomorphic T \<Longrightarrow> Borsukian S \<longleftrightarrow> Borsukian T"
   by (meson homeomorphic_Borsukian homeomorphic_sym)
 
-lemma%unimportant Borsukian_translation:
+lemma Borsukian_translation:
   fixes S :: "'a::real_normed_vector set"
   shows "Borsukian (image (\<lambda>x. a + x) S) \<longleftrightarrow> Borsukian S"
   apply (rule homeomorphic_Borsukian_eq)
     using homeomorphic_translation homeomorphic_sym by blast
 
-lemma%unimportant Borsukian_injective_linear_image:
+lemma Borsukian_injective_linear_image:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "linear f" "inj f"
     shows "Borsukian(f ` S) \<longleftrightarrow> Borsukian S"
   apply (rule homeomorphic_Borsukian_eq)
   using assms homeomorphic_sym linear_homeomorphic_image by blast
 
-lemma%unimportant homotopy_eqv_Borsukianness:
+lemma homotopy_eqv_Borsukianness:
   fixes S :: "'a::real_normed_vector set"
     and T :: "'b::real_normed_vector set"
    assumes "S homotopy_eqv T"
      shows "(Borsukian S \<longleftrightarrow> Borsukian T)"
   by (meson Borsukian_def assms homotopy_eqv_cohomotopic_triviality_null)
 
-lemma%unimportant Borsukian_alt:
+lemma Borsukian_alt:
   fixes S :: "'a::real_normed_vector set"
   shows
    "Borsukian S \<longleftrightarrow>
@@ -4259,20 +4258,20 @@
   unfolding Borsukian_def homotopic_triviality
   by (simp add: path_connected_punctured_universe)
 
-lemma%unimportant Borsukian_continuous_logarithm:
+lemma Borsukian_continuous_logarithm:
   fixes S :: "'a::real_normed_vector set"
   shows "Borsukian S \<longleftrightarrow>
             (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> (- {0::complex})
                  \<longrightarrow> (\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))))"
   by (simp add: Borsukian_def inessential_eq_continuous_logarithm)
 
-lemma%important Borsukian_continuous_logarithm_circle:
+lemma Borsukian_continuous_logarithm_circle:
   fixes S :: "'a::real_normed_vector set"
   shows "Borsukian S \<longleftrightarrow>
              (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1
                   \<longrightarrow> (\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))))"
    (is "?lhs = ?rhs")
-proof%unimportant
+proof
   assume ?lhs then show ?rhs
     by (force simp: Borsukian_continuous_logarithm)
 next
@@ -4302,13 +4301,13 @@
 qed
 
 
-lemma%important Borsukian_continuous_logarithm_circle_real:
+lemma Borsukian_continuous_logarithm_circle_real:
   fixes S :: "'a::real_normed_vector set"
   shows "Borsukian S \<longleftrightarrow>
          (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1
               \<longrightarrow> (\<exists>g. continuous_on S (complex_of_real \<circ> g) \<and> (\<forall>x \<in> S. f x = exp(\<i> * of_real(g x)))))"
    (is "?lhs = ?rhs")
-proof%unimportant
+proof
   assume LHS: ?lhs
   show ?rhs
   proof (clarify)
@@ -4335,52 +4334,52 @@
   qed
 qed
 
-lemma%unimportant Borsukian_circle:
+lemma Borsukian_circle:
   fixes S :: "'a::real_normed_vector set"
   shows "Borsukian S \<longleftrightarrow>
          (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1
               \<longrightarrow> (\<exists>a. homotopic_with (\<lambda>h. True) S (sphere (0::complex) 1) f (\<lambda>x. a)))"
 by (simp add: inessential_eq_continuous_logarithm_circle Borsukian_continuous_logarithm_circle_real)
 
-lemma%unimportant contractible_imp_Borsukian: "contractible S \<Longrightarrow> Borsukian S"
+lemma contractible_imp_Borsukian: "contractible S \<Longrightarrow> Borsukian S"
   by (meson Borsukian_def nullhomotopic_from_contractible)
 
-lemma%unimportant simply_connected_imp_Borsukian:
+lemma simply_connected_imp_Borsukian:
   fixes S :: "'a::real_normed_vector set"
   shows  "\<lbrakk>simply_connected S; locally path_connected S\<rbrakk> \<Longrightarrow> Borsukian S"
   apply (simp add: Borsukian_continuous_logarithm)
   by (metis (no_types, lifting) continuous_logarithm_on_simply_connected image_iff)
 
-lemma%unimportant starlike_imp_Borsukian:
+lemma starlike_imp_Borsukian:
   fixes S :: "'a::real_normed_vector set"
   shows "starlike S \<Longrightarrow> Borsukian S"
   by (simp add: contractible_imp_Borsukian starlike_imp_contractible)
 
-lemma%unimportant Borsukian_empty: "Borsukian {}"
+lemma Borsukian_empty: "Borsukian {}"
   by (auto simp: contractible_imp_Borsukian)
 
-lemma%unimportant Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)"
+lemma Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)"
   by (auto simp: contractible_imp_Borsukian)
 
-lemma%unimportant convex_imp_Borsukian:
+lemma convex_imp_Borsukian:
   fixes S :: "'a::real_normed_vector set"
   shows "convex S \<Longrightarrow> Borsukian S"
   by (meson Borsukian_def convex_imp_contractible nullhomotopic_from_contractible)
 
-lemma%unimportant Borsukian_sphere:
+proposition Borsukian_sphere:
   fixes a :: "'a::euclidean_space"
   shows "3 \<le> DIM('a) \<Longrightarrow> Borsukian (sphere a r)"
   apply (rule simply_connected_imp_Borsukian)
   using simply_connected_sphere apply blast
   using ENR_imp_locally_path_connected ENR_sphere by blast
 
-lemma%important Borsukian_open_Un:
+proposition Borsukian_open_Un:
   fixes S :: "'a::real_normed_vector set"
   assumes opeS: "openin (subtopology euclidean (S \<union> T)) S"
       and opeT: "openin (subtopology euclidean (S \<union> T)) T"
       and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \<inter> T)"
     shows "Borsukian(S \<union> T)"
-proof%unimportant (clarsimp simp add: Borsukian_continuous_logarithm)
+proof (clarsimp simp add: Borsukian_continuous_logarithm)
   fix f :: "'a \<Rightarrow> complex"
   assume contf: "continuous_on (S \<union> T) f" and 0: "0 \<notin> f ` (S \<union> T)"
   then have contfS: "continuous_on S f" and contfT: "continuous_on T f"
@@ -4444,13 +4443,13 @@
 qed
 
 text\<open>The proof is a duplicate of that of \<open>Borsukian_open_Un\<close>.\<close>
-lemma%important Borsukian_closed_Un:
+lemma Borsukian_closed_Un:
   fixes S :: "'a::real_normed_vector set"
   assumes cloS: "closedin (subtopology euclidean (S \<union> T)) S"
       and cloT: "closedin (subtopology euclidean (S \<union> T)) T"
       and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \<inter> T)"
     shows "Borsukian(S \<union> T)"
-proof%unimportant (clarsimp simp add: Borsukian_continuous_logarithm)
+proof (clarsimp simp add: Borsukian_continuous_logarithm)
   fix f :: "'a \<Rightarrow> complex"
   assume contf: "continuous_on (S \<union> T) f" and 0: "0 \<notin> f ` (S \<union> T)"
   then have contfS: "continuous_on S f" and contfT: "continuous_on T f"
@@ -4513,18 +4512,18 @@
   qed
 qed
 
-lemma%unimportant Borsukian_separation_compact:
+lemma Borsukian_separation_compact:
   fixes S :: "complex set"
   assumes "compact S"
     shows "Borsukian S \<longleftrightarrow> connected(- S)"
   by (simp add: Borsuk_separation_theorem Borsukian_circle assms)
 
-lemma%important Borsukian_monotone_image_compact:
+lemma Borsukian_monotone_image_compact:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T"
       and "compact S" and conn: "\<And>y. y \<in> T \<Longrightarrow> connected {x. x \<in> S \<and> f x = y}"
     shows "Borsukian T"
-proof%unimportant (clarsimp simp add: Borsukian_continuous_logarithm)
+proof (clarsimp simp add: Borsukian_continuous_logarithm)
   fix g :: "'b \<Rightarrow> complex"
   assume contg: "continuous_on T g" and 0: "0 \<notin> g ` T"
   have "continuous_on S (g \<circ> f)"
@@ -4584,13 +4583,13 @@
 qed
 
 
-lemma%important Borsukian_open_map_image_compact:
+lemma Borsukian_open_map_image_compact:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S"
       and ope: "\<And>U. openin (subtopology euclidean S) U
                      \<Longrightarrow> openin (subtopology euclidean T) (f ` U)"
     shows "Borsukian T"
-proof%unimportant (clarsimp simp add: Borsukian_continuous_logarithm_circle_real)
+proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real)
   fix g :: "'b \<Rightarrow> complex"
   assume contg: "continuous_on T g" and gim: "g ` T \<subseteq> sphere 0 1"
   have "continuous_on S (g \<circ> f)"
@@ -4671,12 +4670,12 @@
 
 
 text\<open>If two points are separated by a closed set, there's a minimal one.\<close>
-proposition%important closed_irreducible_separator:
+proposition closed_irreducible_separator:
   fixes a :: "'a::real_normed_vector"
   assumes "closed S" and ab: "\<not> connected_component (- S) a b"
   obtains T where "T \<subseteq> S" "closed T" "T \<noteq> {}" "\<not> connected_component (- T) a b"
                   "\<And>U. U \<subset> T \<Longrightarrow> connected_component (- U) a b"
-proof%unimportant (cases "a \<in> S \<or> b \<in> S")
+proof (cases "a \<in> S \<or> b \<in> S")
   case True
   then show ?thesis
   proof
@@ -4779,7 +4778,7 @@
   qed
 qed
 
-lemma%unimportant frontier_minimal_separating_closed_pointwise:
+lemma frontier_minimal_separating_closed_pointwise:
   fixes S :: "'a::real_normed_vector set"
   assumes S: "closed S" "a \<notin> S" and nconn: "\<not> connected_component (- S) a b"
       and conn: "\<And>T. \<lbrakk>closed T; T \<subset> S\<rbrakk> \<Longrightarrow> connected_component (- T) a b"
@@ -4816,21 +4815,21 @@
         closedin (subtopology euclidean U) S \<and> closedin (subtopology euclidean U) T
         \<longrightarrow> connected (S \<inter> T)"
 
-lemma%unimportant unicoherentI [intro?]:
+lemma unicoherentI [intro?]:
   assumes "\<And>S T. \<lbrakk>connected S; connected T; U = S \<union> T; closedin (subtopology euclidean U) S; closedin (subtopology euclidean U) T\<rbrakk>
           \<Longrightarrow> connected (S \<inter> T)"
   shows "unicoherent U"
   using assms unfolding unicoherent_def by blast
 
-lemma%unimportant unicoherentD:
+lemma unicoherentD:
   assumes "unicoherent U" "connected S" "connected T" "U = S \<union> T" "closedin (subtopology euclidean U) S" "closedin (subtopology euclidean U) T"
   shows "connected (S \<inter> T)"
   using assms unfolding unicoherent_def by blast
 
-lemma%important homeomorphic_unicoherent:
+proposition homeomorphic_unicoherent:
   assumes ST: "S homeomorphic T" and S: "unicoherent S"
   shows "unicoherent T"
-proof%unimportant -
+proof -
   obtain f g where gf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" and fim: "T = f ` S" and gfim: "g ` f ` S = S"
     and contf: "continuous_on S f" and contg: "continuous_on (f ` S) g"
     using ST by (auto simp: homeomorphic_def homeomorphism_def)
@@ -4871,24 +4870,24 @@
 qed
 
 
-lemma%unimportant homeomorphic_unicoherent_eq:
+lemma homeomorphic_unicoherent_eq:
    "S homeomorphic T \<Longrightarrow> (unicoherent S \<longleftrightarrow> unicoherent T)"
   by (meson homeomorphic_sym homeomorphic_unicoherent)
 
-lemma%unimportant unicoherent_translation:
+lemma unicoherent_translation:
   fixes S :: "'a::real_normed_vector set"
   shows
    "unicoherent (image (\<lambda>x. a + x) S) \<longleftrightarrow> unicoherent S"
   using homeomorphic_translation homeomorphic_unicoherent_eq by blast
 
-lemma%unimportant unicoherent_injective_linear_image:
+lemma unicoherent_injective_linear_image:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "linear f" "inj f"
   shows "(unicoherent(f ` S) \<longleftrightarrow> unicoherent S)"
   using assms homeomorphic_unicoherent_eq linear_homeomorphic_image by blast
 
 
-lemma%unimportant Borsukian_imp_unicoherent:
+lemma Borsukian_imp_unicoherent:
   fixes U :: "'a::euclidean_space set"
   assumes "Borsukian U"  shows "unicoherent U"
   unfolding unicoherent_def
@@ -4998,27 +4997,27 @@
 qed
 
 
-corollary%important contractible_imp_unicoherent:
+corollary contractible_imp_unicoherent:
   fixes U :: "'a::euclidean_space set"
   assumes "contractible U"  shows "unicoherent U"
   by%unimportant (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian)
 
-corollary%important convex_imp_unicoherent:
+corollary convex_imp_unicoherent:
   fixes U :: "'a::euclidean_space set"
   assumes "convex U"  shows "unicoherent U"
   by%unimportant (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian)
 
 text\<open>If the type class constraint can be relaxed, I don't know how!\<close>
-corollary%important unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)"
+corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)"
   by%unimportant (simp add: convex_imp_unicoherent)
 
 
-lemma%important unicoherent_monotone_image_compact:
+lemma unicoherent_monotone_image_compact:
   fixes T :: "'b :: t2_space set"
   assumes S: "unicoherent S" "compact S" and contf: "continuous_on S f" and fim: "f ` S = T"
   and conn: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
   shows "unicoherent T"
-proof%unimportant
+proof
   fix U V
   assume UV: "connected U" "connected V" "T = U \<union> V"
      and cloU: "closedin (subtopology euclidean T) U"
@@ -5054,7 +5053,7 @@
 
 subsection%important\<open>Several common variants of unicoherence\<close>
 
-lemma%unimportant connected_frontier_simple:
+lemma connected_frontier_simple:
   fixes S :: "'a :: euclidean_space set"
   assumes "connected S" "connected(- S)" shows "connected(frontier S)"
   unfolding frontier_closures
@@ -5062,18 +5061,18 @@
       apply (simp_all add: assms connected_imp_connected_closure)
   by (simp add: closure_def)
 
-lemma%unimportant connected_frontier_component_complement:
+lemma connected_frontier_component_complement:
   fixes S :: "'a :: euclidean_space set"
   assumes "connected S" and C: "C \<in> components(- S)" shows "connected(frontier C)"
   apply (rule connected_frontier_simple)
   using C in_components_connected apply blast
   by (metis Compl_eq_Diff_UNIV connected_UNIV assms top_greatest component_complement_connected)
 
-lemma%important connected_frontier_disjoint:
+lemma connected_frontier_disjoint:
   fixes S :: "'a :: euclidean_space set"
   assumes "connected S" "connected T" "disjnt S T" and ST: "frontier S \<subseteq> frontier T"
   shows "connected(frontier S)"
-proof%unimportant (cases "S = UNIV")
+proof (cases "S = UNIV")
   case True then show ?thesis
     by simp
 next
@@ -5107,11 +5106,11 @@
 
 subsection%important\<open>Some separation results\<close>
 
-lemma%important separation_by_component_closed_pointwise:
+lemma separation_by_component_closed_pointwise:
   fixes S :: "'a :: euclidean_space set"
   assumes "closed S" "\<not> connected_component (- S) a b"
   obtains C where "C \<in> components S" "\<not> connected_component(- C) a b"
-proof%unimportant (cases "a \<in> S \<or> b \<in> S")
+proof (cases "a \<in> S \<or> b \<in> S")
   case True
   then show ?thesis
     using connected_component_in componentsI that by fastforce
@@ -5144,11 +5143,11 @@
 qed
 
 
-lemma%important separation_by_component_closed:
+lemma separation_by_component_closed:
   fixes S :: "'a :: euclidean_space set"
   assumes "closed S" "\<not> connected(- S)"
   obtains C where "C \<in> components S" "\<not> connected(- C)"
-proof%unimportant -
+proof -
   obtain x y where "closed S" "x \<notin> S" "y \<notin> S" and "\<not> connected_component (- S) x y"
     using assms by (auto simp: connected_iff_connected_component)
   then obtain C where "C \<in> components S" "\<not> connected_component(- C) x y"
@@ -5158,12 +5157,12 @@
     by (metis Compl_iff \<open>C \<in> components S\<close> \<open>x \<notin> S\<close> \<open>y \<notin> S\<close> connected_component_eq connected_component_eq_eq connected_iff_connected_component that)
 qed
 
-lemma%important separation_by_Un_closed_pointwise:
+lemma separation_by_Un_closed_pointwise:
   fixes S :: "'a :: euclidean_space set"
   assumes ST: "closed S" "closed T" "S \<inter> T = {}"
       and conS: "connected_component (- S) a b" and conT: "connected_component (- T) a b"
     shows "connected_component (- (S \<union> T)) a b"
-proof%unimportant (rule ccontr)
+proof (rule ccontr)
   have "a \<notin> S" "b \<notin> S" "a \<notin> T" "b \<notin> T"
     using conS conT connected_component_in by auto
   assume "\<not> connected_component (- (S \<union> T)) a b"
@@ -5182,14 +5181,14 @@
     by (meson Compl_anti_mono C conS conT connected_component_of_subset)
 qed
 
-lemma%unimportant separation_by_Un_closed:
+lemma separation_by_Un_closed:
   fixes S :: "'a :: euclidean_space set"
   assumes ST: "closed S" "closed T" "S \<inter> T = {}" and conS: "connected(- S)" and conT: "connected(- T)"
   shows "connected(- (S \<union> T))"
   using assms separation_by_Un_closed_pointwise
   by (fastforce simp add: connected_iff_connected_component)
 
-lemma%unimportant open_unicoherent_UNIV:
+lemma open_unicoherent_UNIV:
   fixes S :: "'a :: euclidean_space set"
   assumes "open S" "open T" "connected S" "connected T" "S \<union> T = UNIV"
   shows "connected(S \<inter> T)"
@@ -5200,7 +5199,7 @@
     by simp
 qed
 
-lemma%unimportant separation_by_component_open_aux:
+lemma separation_by_component_open_aux:
   fixes S :: "'a :: euclidean_space set"
   assumes ST: "closed S" "closed T" "S \<inter> T = {}"
       and "S \<noteq> {}" "T \<noteq> {}"
@@ -5280,11 +5279,11 @@
 qed
 
 
-lemma%important separation_by_component_open:
+proposition separation_by_component_open:
   fixes S :: "'a :: euclidean_space set"
   assumes "open S" and non: "\<not> connected(- S)"
   obtains C where "C \<in> components S" "\<not> connected(- C)"
-proof%unimportant -
+proof -
   obtain T U
     where "closed T" "closed U" and TU: "T \<union> U = - S" "T \<inter> U = {}" "T \<noteq> {}" "U \<noteq> {}"
     using assms by (auto simp: connected_closed_set closed_def)
@@ -5307,18 +5306,18 @@
   qed
 qed
 
-lemma%unimportant separation_by_Un_open:
+lemma separation_by_Un_open:
   fixes S :: "'a :: euclidean_space set"
   assumes "open S" "open T" "S \<inter> T = {}" and cS: "connected(-S)" and cT: "connected(-T)"
     shows "connected(- (S \<union> T))"
   using assms unicoherent_UNIV unfolding unicoherent_def by force
 
 
-lemma%important nonseparation_by_component_eq:
+lemma nonseparation_by_component_eq:
   fixes S :: "'a :: euclidean_space set"
   assumes "open S \<or> closed S"
   shows "((\<forall>C \<in> components S. connected(-C)) \<longleftrightarrow> connected(- S))" (is "?lhs = ?rhs")
-proof%unimportant
+proof
   assume ?lhs with assms show ?rhs
     by (meson separation_by_component_closed separation_by_component_open)
 next
@@ -5328,13 +5327,13 @@
 
 
 text\<open>Another interesting equivalent of an inessential mapping into C-{0}\<close>
-proposition%important inessential_eq_extensible:
+proposition inessential_eq_extensible:
   fixes f :: "'a::euclidean_space \<Rightarrow> complex"
   assumes "closed S"
   shows "(\<exists>a. homotopic_with (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)) \<longleftrightarrow>
          (\<exists>g. continuous_on UNIV g \<and> (\<forall>x \<in> S. g x = f x) \<and> (\<forall>x. g x \<noteq> 0))"
      (is "?lhs = ?rhs")
-proof%unimportant
+proof
   assume ?lhs
   then obtain a where a: "homotopic_with (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)" ..
   show ?rhs
@@ -5394,14 +5393,14 @@
     by (simp add: inessential_eq_continuous_logarithm)
 qed
 
-lemma%important inessential_on_clopen_Union:
+lemma inessential_on_clopen_Union:
   fixes \<F> :: "'a::euclidean_space set set"
   assumes T: "path_connected T"
       and "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (subtopology euclidean (\<Union>\<F>)) S"
       and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean (\<Union>\<F>)) S"
       and hom: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. a)"
   obtains a where "homotopic_with (\<lambda>x. True) (\<Union>\<F>) T f (\<lambda>x. a)"
-proof%unimportant (cases "\<Union>\<F> = {}")
+proof (cases "\<Union>\<F> = {}")
   case True
   with that show ?thesis
     by force
@@ -5442,12 +5441,12 @@
   then show ?thesis ..
 qed
 
-lemma%important Janiszewski_dual:
+proposition Janiszewski_dual:
   fixes S :: "complex set"
   assumes
    "compact S" "compact T" "connected S" "connected T" "connected(- (S \<union> T))"
  shows "connected(S \<inter> T)"
-proof%unimportant -
+proof -
   have ST: "compact (S \<union> T)"
     by (simp add: assms compact_Un)
   with Borsukian_imp_unicoherent [of "S \<union> T"] ST assms
--- a/src/HOL/Analysis/Great_Picard.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Great_Picard.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -9,13 +9,13 @@
   
 subsection\<open>Schottky's theorem\<close>
 
-lemma%important Schottky_lemma0:
+lemma Schottky_lemma0:
   assumes holf: "f holomorphic_on S" and cons: "contractible S" and "a \<in> S"
       and f: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 1 \<and> f z \<noteq> -1"
   obtains g where "g holomorphic_on S"
                   "norm(g a) \<le> 1 + norm(f a) / 3"
                   "\<And>z. z \<in> S \<Longrightarrow> f z = cos(of_real pi * g z)"
-proof%unimportant -
+proof -
   obtain g where holg: "g holomorphic_on S" and g: "norm(g a) \<le> pi + norm(f a)"
              and f_eq_cos: "\<And>z. z \<in> S \<Longrightarrow> f z = cos(g z)"
     using contractible_imp_holomorphic_arccos_bounded [OF assms]
@@ -38,7 +38,7 @@
 qed
 
 
-lemma%unimportant Schottky_lemma1:
+lemma Schottky_lemma1:
   fixes n::nat
   assumes "0 < n"
   shows "0 < n + sqrt(real n ^ 2 - 1)"
@@ -54,7 +54,7 @@
 qed
 
 
-lemma%unimportant Schottky_lemma2:
+lemma Schottky_lemma2:
   fixes x::real
   assumes "0 \<le> x"
   obtains n where "0 < n" "\<bar>x - ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi\<bar> < 1/2"
@@ -158,12 +158,12 @@
 qed
 
 
-lemma%important Schottky_lemma3:
+lemma Schottky_lemma3:
   fixes z::complex
   assumes "z \<in> (\<Union>m \<in> Ints. \<Union>n \<in> {0<..}. {Complex m (ln(n + sqrt(real n ^ 2 - 1)) / pi)})
              \<union> (\<Union>m \<in> Ints. \<Union>n \<in> {0<..}. {Complex m (-ln(n + sqrt(real n ^ 2 - 1)) / pi)})"
   shows "cos(pi * cos(pi * z)) = 1 \<or> cos(pi * cos(pi * z)) = -1"
-proof%unimportant -
+proof -
   have sqrt2 [simp]: "complex_of_real (sqrt x) * complex_of_real (sqrt x) = x" if "x \<ge> 0" for x::real
     by (metis abs_of_nonneg of_real_mult real_sqrt_mult_self that)
   have 1: "\<exists>k. exp (\<i> * (of_int m * complex_of_real pi) -
@@ -223,13 +223,13 @@
 qed
 
 
-theorem%important Schottky:
+theorem Schottky:
   assumes holf: "f holomorphic_on cball 0 1"
       and nof0: "norm(f 0) \<le> r"
       and not01: "\<And>z. z \<in> cball 0 1 \<Longrightarrow> \<not>(f z = 0 \<or> f z = 1)"
       and "0 < t" "t < 1" "norm z \<le> t"
     shows "norm(f z) \<le> exp(pi * exp(pi * (2 + 2 * r + 12 * t / (1 - t))))"
-proof%unimportant -
+proof -
   obtain h where holf: "h holomorphic_on cball 0 1"
              and nh0: "norm (h 0) \<le> 1 + norm(2 * f 0 - 1) / 3"
              and h:   "\<And>z. z \<in> cball 0 1 \<Longrightarrow> 2 * f z - 1 = cos(of_real pi * h z)"
@@ -376,12 +376,12 @@
   
 subsection%important\<open>The Little Picard Theorem\<close>
 
-lemma%important Landau_Picard:
+proposition Landau_Picard:
   obtains R
     where "\<And>z. 0 < R z"
           "\<And>f. \<lbrakk>f holomorphic_on cball 0 (R(f 0));
                  \<And>z. norm z \<le> R(f 0) \<Longrightarrow> f z \<noteq> 0 \<and> f z \<noteq> 1\<rbrakk> \<Longrightarrow> norm(deriv f 0) < 1"
-proof%unimportant -
+proof -
   define R where "R \<equiv> \<lambda>z. 3 * exp(pi * exp(pi*(2 + 2 * cmod z + 12)))"
   show ?thesis
   proof
@@ -436,10 +436,10 @@
   qed
 qed
 
-lemma%important little_Picard_01:
+lemma little_Picard_01:
   assumes holf: "f holomorphic_on UNIV" and f01: "\<And>z. f z \<noteq> 0 \<and> f z \<noteq> 1"
   obtains c where "f = (\<lambda>x. c)"
-proof%unimportant -
+proof -
   obtain R
     where Rpos: "\<And>z. 0 < R z"
       and R:    "\<And>h. \<lbrakk>h holomorphic_on cball 0 (R(h 0));
@@ -481,11 +481,11 @@
 qed
 
 
-theorem%important little_Picard:
+theorem little_Picard:
   assumes holf: "f holomorphic_on UNIV"
       and "a \<noteq> b" "range f \<inter> {a,b} = {}"
     obtains c where "f = (\<lambda>x. c)"
-proof%unimportant -
+proof -
   let ?g = "\<lambda>x. 1/(b - a)*(f x - b) + 1"
   obtain c where "?g = (\<lambda>x. c)"
   proof (rule little_Picard_01)
@@ -505,7 +505,7 @@
 
 text\<open>A couple of little applications of Little Picard\<close>
 
-lemma%unimportant holomorphic_periodic_fixpoint:
+lemma holomorphic_periodic_fixpoint:
   assumes holf: "f holomorphic_on UNIV"
       and "p \<noteq> 0" and per: "\<And>z. f(z + p) = f z"
   obtains x where "f x = x"
@@ -527,7 +527,7 @@
 qed
 
 
-lemma%unimportant holomorphic_involution_point:
+lemma holomorphic_involution_point:
   assumes holfU: "f holomorphic_on UNIV" and non: "\<And>a. f \<noteq> (\<lambda>x. a + x)"
   obtains x where "f(f x) = x"
 proof -
@@ -616,7 +616,7 @@
 
 subsection%important\<open>The Arzelà--Ascoli theorem\<close>
 
-lemma%unimportant subsequence_diagonalization_lemma:
+lemma subsequence_diagonalization_lemma:
   fixes P :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> bool"
   assumes sub: "\<And>i r. \<exists>k. strict_mono (k :: nat \<Rightarrow> nat) \<and> P i (r \<circ> k)"
       and P_P:  "\<And>i r::nat \<Rightarrow> 'a. \<And>k1 k2 N.
@@ -660,7 +660,7 @@
   qed
 qed
 
-lemma%unimportant function_convergent_subsequence:
+lemma function_convergent_subsequence:
   fixes f :: "[nat,'a] \<Rightarrow> 'b::{real_normed_vector,heine_borel}"
   assumes "countable S" and M: "\<And>n::nat. \<And>x. x \<in> S \<Longrightarrow> norm(f n x) \<le> M"
    obtains k where "strict_mono (k::nat\<Rightarrow>nat)" "\<And>x. x \<in> S \<Longrightarrow> \<exists>l. (\<lambda>n. f (k n) x) \<longlonglongrightarrow> l"
@@ -698,7 +698,7 @@
 qed
 
 
-theorem%important Arzela_Ascoli:
+theorem Arzela_Ascoli:
   fixes \<F> :: "[nat,'a::euclidean_space] \<Rightarrow> 'b::{real_normed_vector,heine_borel}"
   assumes "compact S"
       and M: "\<And>n x. x \<in> S \<Longrightarrow> norm(\<F> n x) \<le> M"
@@ -707,7 +707,7 @@
                  \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>n y. y \<in> S \<and> norm(x - y) < d \<longrightarrow> norm(\<F> n x - \<F> n y) < e)"
   obtains g k where "continuous_on S g" "strict_mono (k :: nat \<Rightarrow> nat)"
                     "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<and> x \<in> S \<longrightarrow> norm(\<F>(k n) x - g x) < e"
-proof%unimportant -
+proof -
   have UEQ: "\<And>e. 0 < e \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>n. \<forall>x \<in> S. \<forall>x' \<in> S. dist x' x < d \<longrightarrow> dist (\<F> n x') (\<F> n x) < e)"
     apply (rule compact_uniformly_equicontinuous [OF \<open>compact S\<close>, of "range \<F>"])
     using equicont by (force simp: dist_commute dist_norm)+
@@ -795,7 +795,7 @@
 holomorphic function, and converges \emph{uniformly} on compact subsets of S.\<close>
 
 
-theorem%important Montel:
+theorem Montel:
   fixes \<F> :: "[nat,complex] \<Rightarrow> complex"
   assumes "open S"
       and \<H>: "\<And>h. h \<in> \<H> \<Longrightarrow> h holomorphic_on S"
@@ -805,7 +805,7 @@
     where "g holomorphic_on S" "strict_mono (r :: nat \<Rightarrow> nat)"
           "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>n. \<F> (r n) x) \<longlongrightarrow> g x) sequentially"
           "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> uniform_limit K (\<F> \<circ> r) g sequentially"        
-proof%unimportant -
+proof -
   obtain K where comK: "\<And>n. compact(K n)" and KS: "\<And>n::nat. K n \<subseteq> S"
              and subK: "\<And>X. \<lbrakk>compact X; X \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. X \<subseteq> K n"
     using open_Union_compact_subsets [OF \<open>open S\<close>] by metis
@@ -995,7 +995,7 @@
 
 subsection%important\<open>Some simple but useful cases of Hurwitz's theorem\<close>
 
-proposition%important Hurwitz_no_zeros:
+proposition Hurwitz_no_zeros:
   assumes S: "open S" "connected S"
       and holf: "\<And>n::nat. \<F> n holomorphic_on S"
       and holg: "g holomorphic_on S"
@@ -1004,7 +1004,7 @@
       and nz: "\<And>n z. z \<in> S \<Longrightarrow> \<F> n z \<noteq> 0"
       and "z0 \<in> S"
       shows "g z0 \<noteq> 0"
-proof%unimportant
+proof
   assume g0: "g z0 = 0"
   obtain h r m
     where "0 < m" "0 < r" and subS: "ball z0 r \<subseteq> S"
@@ -1159,7 +1159,7 @@
   ultimately show False using \<open>0 < m\<close> by auto
 qed
 
-corollary%important Hurwitz_injective:
+corollary Hurwitz_injective:
   assumes S: "open S" "connected S"
       and holf: "\<And>n::nat. \<F> n holomorphic_on S"
       and holg: "g holomorphic_on S"
@@ -1167,7 +1167,7 @@
       and nonconst: "\<not> g constant_on S"
       and inj: "\<And>n. inj_on (\<F> n) S"
     shows "inj_on g S"
-proof%unimportant -
+proof -
   have False if z12: "z1 \<in> S" "z2 \<in> S" "z1 \<noteq> z2" "g z2 = g z1" for z1 z2
   proof -
     obtain z0 where "z0 \<in> S" and z0: "g z0 \<noteq> g z2"
@@ -1231,13 +1231,13 @@
 
 subsection%important\<open>The Great Picard theorem\<close>
 
-lemma%important GPicard1:
+lemma GPicard1:
   assumes S: "open S" "connected S" and "w \<in> S" "0 < r" "Y \<subseteq> X"
       and holX: "\<And>h. h \<in> X \<Longrightarrow> h holomorphic_on S"
       and X01:  "\<And>h z. \<lbrakk>h \<in> X; z \<in> S\<rbrakk> \<Longrightarrow> h z \<noteq> 0 \<and> h z \<noteq> 1"
       and r:    "\<And>h. h \<in> Y \<Longrightarrow> norm(h w) \<le> r"
   obtains B Z where "0 < B" "open Z" "w \<in> Z" "Z \<subseteq> S" "\<And>h z. \<lbrakk>h \<in> Y; z \<in> Z\<rbrakk> \<Longrightarrow> norm(h z) \<le> B"
-proof%unimportant -
+proof -
   obtain e where "e > 0" and e: "cball w e \<subseteq> S"
     using assms open_contains_cball_eq by blast
   show ?thesis
@@ -1277,20 +1277,20 @@
   qed (use \<open>e > 0\<close> in auto)
 qed 
 
-lemma%important GPicard2:
+lemma GPicard2:
   assumes "S \<subseteq> T" "connected T" "S \<noteq> {}" "open S" "\<And>x. \<lbrakk>x islimpt S; x \<in> T\<rbrakk> \<Longrightarrow> x \<in> S"
     shows "S = T"
-  by%unimportant (metis assms open_subset connected_clopen closedin_limpt)
+  by (metis assms open_subset connected_clopen closedin_limpt)
 
     
-lemma%important GPicard3:
+lemma GPicard3:
   assumes S: "open S" "connected S" "w \<in> S" and "Y \<subseteq> X"
       and holX: "\<And>h. h \<in> X \<Longrightarrow> h holomorphic_on S"
       and X01:  "\<And>h z. \<lbrakk>h \<in> X; z \<in> S\<rbrakk> \<Longrightarrow> h z \<noteq> 0 \<and> h z \<noteq> 1"
       and no_hw_le1: "\<And>h. h \<in> Y \<Longrightarrow> norm(h w) \<le> 1"
       and "compact K" "K \<subseteq> S"
   obtains B where "\<And>h z. \<lbrakk>h \<in> Y; z \<in> K\<rbrakk> \<Longrightarrow> norm(h z) \<le> B"
-proof%unimportant -
+proof -
   define U where "U \<equiv> {z \<in> S. \<exists>B Z. 0 < B \<and> open Z \<and> z \<in> Z \<and> Z \<subseteq> S \<and>
                                (\<forall>h z'. h \<in> Y \<and> z' \<in> Z \<longrightarrow> norm(h z') \<le> B)}"
   then have "U \<subseteq> S" by blast
@@ -1432,11 +1432,11 @@
 qed
 
 
-lemma%important GPicard4:
+lemma GPicard4:
   assumes "0 < k" and holf: "f holomorphic_on (ball 0 k - {0})" 
       and AE: "\<And>e. \<lbrakk>0 < e; e < k\<rbrakk> \<Longrightarrow> \<exists>d. 0 < d \<and> d < e \<and> (\<forall>z \<in> sphere 0 d. norm(f z) \<le> B)"
   obtains \<epsilon> where "0 < \<epsilon>" "\<epsilon> < k" "\<And>z. z \<in> ball 0 \<epsilon> - {0} \<Longrightarrow> norm(f z) \<le> B"
-proof%unimportant -
+proof -
   obtain \<epsilon> where "0 < \<epsilon>" "\<epsilon> < k/2" and \<epsilon>: "\<And>z. norm z = \<epsilon> \<Longrightarrow> norm(f z) \<le> B"
     using AE [of "k/2"] \<open>0 < k\<close> by auto
   show ?thesis
@@ -1471,13 +1471,13 @@
 qed
   
 
-lemma%important GPicard5:
+lemma GPicard5:
   assumes holf: "f holomorphic_on (ball 0 1 - {0})"
       and f01:  "\<And>z. z \<in> ball 0 1 - {0} \<Longrightarrow> f z \<noteq> 0 \<and> f z \<noteq> 1"
   obtains e B where "0 < e" "e < 1" "0 < B" 
                     "(\<forall>z \<in> ball 0 e - {0}. norm(f z) \<le> B) \<or>
                      (\<forall>z \<in> ball 0 e - {0}. norm(f z) \<ge> B)"
-proof%unimportant -
+proof -
   have [simp]: "1 + of_nat n \<noteq> (0::complex)" for n
     using of_nat_eq_0_iff by fastforce
   have [simp]: "cmod (1 + of_nat n) = 1 + of_nat n" for n
@@ -1614,13 +1614,13 @@
 qed
 
   
-lemma%important GPicard6:
+lemma GPicard6:
   assumes "open M" "z \<in> M" "a \<noteq> 0" and holf: "f holomorphic_on (M - {z})"
       and f0a: "\<And>w. w \<in> M - {z} \<Longrightarrow> f w \<noteq> 0 \<and> f w \<noteq> a"
   obtains r where "0 < r" "ball z r \<subseteq> M" 
                   "bounded(f ` (ball z r - {z})) \<or>
                    bounded((inverse \<circ> f) ` (ball z r - {z}))"
-proof%unimportant -
+proof -
   obtain r where "0 < r" and r: "ball z r \<subseteq> M"
     using assms openE by blast 
   let ?g = "\<lambda>w. f (z + of_real r * w) / a"
@@ -1669,11 +1669,11 @@
 qed
   
 
-theorem%important great_Picard:
+theorem great_Picard:
   assumes "open M" "z \<in> M" "a \<noteq> b" and holf: "f holomorphic_on (M - {z})"
       and fab: "\<And>w. w \<in> M - {z} \<Longrightarrow> f w \<noteq> a \<and> f w \<noteq> b"
   obtains l where "(f \<longlongrightarrow> l) (at z) \<or> ((inverse \<circ> f) \<longlongrightarrow> l) (at z)"
-proof%unimportant -
+proof -
   obtain r where "0 < r" and zrM: "ball z r \<subseteq> M" 
              and r: "bounded((\<lambda>z. f z - a) ` (ball z r - {z})) \<or>
                      bounded((inverse \<circ> (\<lambda>z. f z - a)) ` (ball z r - {z}))"
@@ -1776,7 +1776,7 @@
 qed
 
 
-corollary%important great_Picard_alt:
+corollary great_Picard_alt:
   assumes M: "open M" "z \<in> M" and holf: "f holomorphic_on (M - {z})"
     and non: "\<And>l. \<not> (f \<longlongrightarrow> l) (at z)" "\<And>l. \<not> ((inverse \<circ> f) \<longlongrightarrow> l) (at z)"
   obtains a where "- {a} \<subseteq> f ` (M - {z})"
@@ -1784,11 +1784,11 @@
   by%unimportant (metis great_Picard [OF M _ holf] non)
     
 
-corollary%important great_Picard_infinite:
+corollary great_Picard_infinite:
   assumes M: "open M" "z \<in> M" and holf: "f holomorphic_on (M - {z})"
     and non: "\<And>l. \<not> (f \<longlongrightarrow> l) (at z)" "\<And>l. \<not> ((inverse \<circ> f) \<longlongrightarrow> l) (at z)"
   obtains a where "\<And>w. w \<noteq> a \<Longrightarrow> infinite {x. x \<in> M - {z} \<and> f x = w}"
-proof%unimportant -
+proof -
   have False if "a \<noteq> b" and ab: "finite {x. x \<in> M - {z} \<and> f x = a}" "finite {x. x \<in> M - {z} \<and> f x = b}" for a b
   proof -
     have finab: "finite {x. x \<in> M - {z} \<and> f x \<in> {a,b}}"
@@ -1830,11 +1830,11 @@
     by meson
 qed
 
-corollary%important Casorati_Weierstrass:
+theorem Casorati_Weierstrass:
   assumes "open M" "z \<in> M" "f holomorphic_on (M - {z})"
       and "\<And>l. \<not> (f \<longlongrightarrow> l) (at z)" "\<And>l. \<not> ((inverse \<circ> f) \<longlongrightarrow> l) (at z)"
   shows "closure(f ` (M - {z})) = UNIV"
-proof%unimportant -
+proof -
   obtain a where a: "- {a} \<subseteq> f ` (M - {z})"
     using great_Picard_alt [OF assms] .
   have "UNIV = closure(- {a})"
--- a/src/HOL/Analysis/Homeomorphism.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Homeomorphism.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -8,7 +8,7 @@
 imports Homotopy
 begin
 
-lemma%unimportant homeomorphic_spheres':
+lemma homeomorphic_spheres':
   fixes a ::"'a::euclidean_space" and b ::"'b::euclidean_space"
   assumes "0 < \<delta>" and dimeq: "DIM('a) = DIM('b)"
   shows "(sphere a \<delta>) homeomorphic (sphere b \<delta>)"
@@ -28,7 +28,7 @@
     done
 qed
 
-lemma%unimportant homeomorphic_spheres_gen:
+lemma homeomorphic_spheres_gen:
     fixes a :: "'a::euclidean_space" and b :: "'b::euclidean_space"
   assumes "0 < r" "0 < s" "DIM('a::euclidean_space) = DIM('b::euclidean_space)"
   shows "(sphere a r homeomorphic sphere b s)"
@@ -38,7 +38,7 @@
 
 subsection%important \<open>Homeomorphism of all convex compact sets with nonempty interior\<close>
 
-proposition%important ray_to_rel_frontier:
+proposition ray_to_rel_frontier:
   fixes a :: "'a::real_inner"
   assumes "bounded S"
       and a: "a \<in> rel_interior S"
@@ -46,7 +46,7 @@
       and "l \<noteq> 0"
   obtains d where "0 < d" "(a + d *\<^sub>R l) \<in> rel_frontier S"
            "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S"
-proof%unimportant -
+proof -
   have aaff: "a \<in> affine hull S"
     by (meson a hull_subset rel_interior_subset rev_subsetD)
   let ?D = "{d. 0 < d \<and> a + d *\<^sub>R l \<notin> rel_interior S}"
@@ -139,14 +139,14 @@
     by (rule that [OF \<open>0 < d\<close> infront inint])
 qed
 
-corollary%important ray_to_frontier:
+corollary ray_to_frontier:
   fixes a :: "'a::euclidean_space"
   assumes "bounded S"
       and a: "a \<in> interior S"
       and "l \<noteq> 0"
   obtains d where "0 < d" "(a + d *\<^sub>R l) \<in> frontier S"
            "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (a + e *\<^sub>R l) \<in> interior S"
-proof%unimportant -
+proof -
   have "interior S = rel_interior S"
     using a rel_interior_nonempty_interior by auto
   then have "a \<in> rel_interior S"
@@ -158,7 +158,7 @@
 qed
 
 
-lemma%unimportant segment_to_rel_frontier_aux:
+lemma segment_to_rel_frontier_aux:
   fixes x :: "'a::euclidean_space"
   assumes "convex S" "bounded S" and x: "x \<in> rel_interior S" and y: "y \<in> S" and xy: "x \<noteq> y"
   obtains z where "z \<in> rel_frontier S" "y \<in> closed_segment x z"
@@ -198,7 +198,7 @@
   qed
 qed
 
-lemma%unimportant segment_to_rel_frontier:
+lemma segment_to_rel_frontier:
   fixes x :: "'a::euclidean_space"
   assumes S: "convex S" "bounded S" and x: "x \<in> rel_interior S"
       and y: "y \<in> S" and xy: "\<not>(x = y \<and> S = {x})"
@@ -216,11 +216,11 @@
     using segment_to_rel_frontier_aux [OF S x y] that by blast
 qed
 
-proposition%important rel_frontier_not_sing:
+proposition rel_frontier_not_sing:
   fixes a :: "'a::euclidean_space"
   assumes "bounded S"
     shows "rel_frontier S \<noteq> {a}"
-proof%unimportant (cases "S = {}")
+proof (cases "S = {}")
   case True  then show ?thesis  by simp
 next
   case False
@@ -278,7 +278,7 @@
   qed
 qed
 
-proposition%important
+proposition (*FIXME needs name *)
   fixes S :: "'a::euclidean_space set"
   assumes "compact S" and 0: "0 \<in> rel_interior S"
       and star: "\<And>x. x \<in> S \<Longrightarrow> open_segment 0 x \<subseteq> rel_interior S"
@@ -288,7 +288,7 @@
       and starlike_compact_projective2_0:
             "S homeomorphic cball 0 1 \<inter> affine hull S"
             (is "S homeomorphic ?CBALL")
-proof%unimportant -
+proof -
   have starI: "(u *\<^sub>R x) \<in> rel_interior S" if "x \<in> S" "0 \<le> u" "u < 1" for x u
   proof (cases "x=0 \<or> u=0")
     case True with 0 show ?thesis by force
@@ -540,7 +540,7 @@
     done
 qed
 
-corollary%important
+corollary (* FIXME needs name *)
   fixes S :: "'a::euclidean_space set"
   assumes "compact S" and a: "a \<in> rel_interior S"
       and star: "\<And>x. x \<in> S \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
@@ -548,7 +548,7 @@
             "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S"
       and starlike_compact_projective2:
             "S homeomorphic cball a 1 \<inter> affine hull S"
-proof%unimportant -
+proof -
   have 1: "compact ((+) (-a) ` S)" by (meson assms compact_translation)
   have 2: "0 \<in> rel_interior ((+) (-a) ` S)"
     using a rel_interior_translation [of "- a" S] by (simp cong: image_cong_simp)
@@ -580,12 +580,12 @@
   finally show "S homeomorphic cball a 1 \<inter> affine hull S" .
 qed
 
-corollary%important starlike_compact_projective_special:
+corollary starlike_compact_projective_special:
   assumes "compact S"
     and cb01: "cball (0::'a::euclidean_space) 1 \<subseteq> S"
     and scale: "\<And>x u. \<lbrakk>x \<in> S; 0 \<le> u; u < 1\<rbrakk> \<Longrightarrow> u *\<^sub>R x \<in> S - frontier S"
   shows "S homeomorphic (cball (0::'a::euclidean_space) 1)"
-proof%unimportant -
+proof -
   have "ball 0 1 \<subseteq> interior S"
     using cb01 interior_cball interior_mono by blast
   then have 0: "0 \<in> rel_interior S"
@@ -604,13 +604,13 @@
     using starlike_compact_projective2_0 [OF \<open>compact S\<close> 0 star] by simp
 qed
 
-lemma%important homeomorphic_convex_lemma:
+lemma homeomorphic_convex_lemma:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "convex S" "compact S" "convex T" "compact T"
       and affeq: "aff_dim S = aff_dim T"
     shows "(S - rel_interior S) homeomorphic (T - rel_interior T) \<and>
            S homeomorphic T"
-proof%unimportant (cases "rel_interior S = {} \<or> rel_interior T = {}")
+proof (cases "rel_interior S = {} \<or> rel_interior T = {}")
   case True
     then show ?thesis
       by (metis Diff_empty affeq \<open>convex S\<close> \<open>convex T\<close> aff_dim_empty homeomorphic_empty rel_interior_eq_empty aff_dim_empty)
@@ -708,7 +708,7 @@
     using 1 2 by blast
 qed
 
-lemma%unimportant homeomorphic_convex_compact_sets:
+lemma homeomorphic_convex_compact_sets:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "convex S" "compact S" "convex T" "compact T"
       and affeq: "aff_dim S = aff_dim T"
@@ -716,7 +716,7 @@
 using homeomorphic_convex_lemma [OF assms] assms
 by (auto simp: rel_frontier_def)
 
-lemma%unimportant homeomorphic_rel_frontiers_convex_bounded_sets:
+lemma homeomorphic_rel_frontiers_convex_bounded_sets:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
   assumes "convex S" "bounded S" "convex T" "bounded T"
       and affeq: "aff_dim S = aff_dim T"
@@ -729,7 +729,7 @@
 text\<open>Including the famous stereoscopic projection of the 3-D sphere to the complex plane\<close>
 
 text\<open>The special case with centre 0 and radius 1\<close>
-lemma%unimportant homeomorphic_punctured_affine_sphere_affine_01:
+lemma homeomorphic_punctured_affine_sphere_affine_01:
   assumes "b \<in> sphere 0 1" "affine T" "0 \<in> T" "b \<in> T" "affine p"
       and affT: "aff_dim T = aff_dim p + 1"
     shows "(sphere 0 1 \<inter> T) - {b} homeomorphic p"
@@ -823,12 +823,12 @@
   finally show ?thesis .
 qed
 
-theorem%important homeomorphic_punctured_affine_sphere_affine:
+theorem homeomorphic_punctured_affine_sphere_affine:
   fixes a :: "'a :: euclidean_space"
   assumes "0 < r" "b \<in> sphere a r" "affine T" "a \<in> T" "b \<in> T" "affine p"
       and aff: "aff_dim T = aff_dim p + 1"
     shows "(sphere a r \<inter> T) - {b} homeomorphic p"
-proof%unimportant -
+proof -
   have "a \<noteq> b" using assms by auto
   then have inj: "inj (\<lambda>x::'a. x /\<^sub>R norm (a - b))"
     by (simp add: inj_on_def)
@@ -847,14 +847,14 @@
   finally show ?thesis .
 qed
 
-corollary%important homeomorphic_punctured_sphere_affine:
+corollary homeomorphic_punctured_sphere_affine:
   fixes a :: "'a :: euclidean_space"
   assumes "0 < r" and b: "b \<in> sphere a r"
       and "affine T" and affS: "aff_dim T + 1 = DIM('a)"
     shows "(sphere a r - {b}) homeomorphic T"
   using%unimportant homeomorphic_punctured_affine_sphere_affine [of r b a UNIV T] assms by%unimportant auto
 
-corollary%important homeomorphic_punctured_sphere_hyperplane:
+corollary homeomorphic_punctured_sphere_hyperplane:
   fixes a :: "'a :: euclidean_space"
   assumes "0 < r" and b: "b \<in> sphere a r"
       and "c \<noteq> 0"
@@ -864,12 +864,12 @@
 apply (auto simp: affine_hyperplane of_nat_diff)
 done
 
-proposition%important homeomorphic_punctured_sphere_affine_gen:
+proposition homeomorphic_punctured_sphere_affine_gen:
   fixes a :: "'a :: euclidean_space"
   assumes "convex S" "bounded S" and a: "a \<in> rel_frontier S"
       and "affine T" and affS: "aff_dim S = aff_dim T + 1"
     shows "rel_frontier S - {a} homeomorphic T"
-proof%unimportant -
+proof -
   obtain U :: "'a set" where "affine U" "convex U" and affdS: "aff_dim U = aff_dim S"
     using choose_affine_subset [OF affine_UNIV aff_dim_geq]
     by (meson aff_dim_affine_hull affine_affine_hull affine_imp_convex)
@@ -912,13 +912,13 @@
   is homeomorphic to a closed subset of a convex set, and
   if the set is locally compact we can take the convex set to be the universe.\<close>
 
-proposition%important homeomorphic_closedin_convex:
+proposition homeomorphic_closedin_convex:
   fixes S :: "'m::euclidean_space set"
   assumes "aff_dim S < DIM('n)"
   obtains U and T :: "'n::euclidean_space set"
      where "convex U" "U \<noteq> {}" "closedin (subtopology euclidean U) T"
            "S homeomorphic T"
-proof%unimportant (cases "S = {}")
+proof (cases "S = {}")
   case True then show ?thesis
     by (rule_tac U=UNIV and T="{}" in that) auto
 next
@@ -1009,11 +1009,11 @@
 text\<open> Locally compact sets are closed in an open set and are homeomorphic
   to an absolutely closed set if we have one more dimension to play with.\<close>
 
-lemma%important locally_compact_open_Int_closure:
+lemma locally_compact_open_Int_closure:
   fixes S :: "'a :: metric_space set"
   assumes "locally compact S"
   obtains T where "open T" "S = T \<inter> closure S"
-proof%unimportant -
+proof -
   have "\<forall>x\<in>S. \<exists>T v u. u = S \<inter> T \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<and> open T \<and> compact v"
     by (metis assms locally_compact openin_open)
   then obtain t v where
@@ -1048,14 +1048,14 @@
 qed
 
 
-lemma%unimportant locally_compact_closedin_open:
+lemma locally_compact_closedin_open:
     fixes S :: "'a :: metric_space set"
     assumes "locally compact S"
     obtains T where "open T" "closedin (subtopology euclidean T) S"
   by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int)
 
 
-lemma%unimportant locally_compact_homeomorphism_projection_closed:
+lemma locally_compact_homeomorphism_projection_closed:
   assumes "locally compact S"
   obtains T and f :: "'a \<Rightarrow> 'a :: euclidean_space \<times> 'b :: euclidean_space"
     where "closed T" "homeomorphism S T f fst"
@@ -1108,14 +1108,14 @@
       done
 qed
 
-lemma%unimportant locally_compact_closed_Int_open:
+lemma locally_compact_closed_Int_open:
   fixes S :: "'a :: euclidean_space set"
   shows
     "locally compact S \<longleftrightarrow> (\<exists>U u. closed U \<and> open u \<and> S = U \<inter> u)"
 by (metis closed_closure closed_imp_locally_compact inf_commute locally_compact_Int locally_compact_open_Int_closure open_imp_locally_compact)
 
 
-lemma%unimportant lowerdim_embeddings:
+lemma lowerdim_embeddings:
   assumes  "DIM('a) < DIM('b)"
   obtains f :: "'a::euclidean_space*real \<Rightarrow> 'b::euclidean_space" 
       and g :: "'b \<Rightarrow> 'a*real"
@@ -1161,11 +1161,11 @@
   qed
 qed
 
-proposition%important locally_compact_homeomorphic_closed:
+proposition locally_compact_homeomorphic_closed:
   fixes S :: "'a::euclidean_space set"
   assumes "locally compact S" and dimlt: "DIM('a) < DIM('b)"
   obtains T :: "'b::euclidean_space set" where "closed T" "S homeomorphic T"
-proof%unimportant -
+proof -
   obtain U:: "('a*real)set" and h
     where "closed U" and homU: "homeomorphism S U h fst"
     using locally_compact_homeomorphism_projection_closed assms by metis
@@ -1191,13 +1191,13 @@
 qed
 
 
-lemma%important homeomorphic_convex_compact_lemma:
+lemma homeomorphic_convex_compact_lemma:
   fixes S :: "'a::euclidean_space set"
   assumes "convex S"
     and "compact S"
     and "cball 0 1 \<subseteq> S"
   shows "S homeomorphic (cball (0::'a) 1)"
-proof%unimportant (rule starlike_compact_projective_special[OF assms(2-3)])
+proof (rule starlike_compact_projective_special[OF assms(2-3)])
   fix x u
   assume "x \<in> S" and "0 \<le> u" and "u < (1::real)"
   have "open (ball (u *\<^sub>R x) (1 - u))"
@@ -1223,7 +1223,7 @@
     using frontier_def and interior_subset by auto
 qed
 
-proposition%important homeomorphic_convex_compact_cball:
+proposition homeomorphic_convex_compact_cball:
   fixes e :: real
     and S :: "'a::euclidean_space set"
   assumes "convex S"
@@ -1231,7 +1231,7 @@
     and "interior S \<noteq> {}"
     and "e > 0"
   shows "S homeomorphic (cball (b::'a) e)"
-proof%unimportant -
+proof -
   obtain a where "a \<in> interior S"
     using assms(3) by auto
   then obtain d where "d > 0" and d: "cball a d \<subseteq> S"
@@ -1259,7 +1259,7 @@
     done
 qed
 
-corollary%important homeomorphic_convex_compact:
+corollary homeomorphic_convex_compact:
   fixes S :: "'a::euclidean_space set"
     and T :: "'a set"
   assumes "convex S" "compact S" "interior S \<noteq> {}"
@@ -1268,7 +1268,7 @@
   using assms
   by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym)
 
-subsection%important\<open>Covering spaces and lifting results for them\<close>
+subsection\<open>Covering spaces and lifting results for them\<close>
 
 definition%important covering_space
            :: "'a::topological_space set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
@@ -1281,19 +1281,19 @@
                         pairwise disjnt v \<and>
                         (\<forall>u \<in> v. \<exists>q. homeomorphism u T p q)))"
 
-lemma%unimportant covering_space_imp_continuous: "covering_space c p S \<Longrightarrow> continuous_on c p"
+lemma covering_space_imp_continuous: "covering_space c p S \<Longrightarrow> continuous_on c p"
   by (simp add: covering_space_def)
 
-lemma%unimportant covering_space_imp_surjective: "covering_space c p S \<Longrightarrow> p ` c = S"
+lemma covering_space_imp_surjective: "covering_space c p S \<Longrightarrow> p ` c = S"
   by (simp add: covering_space_def)
 
-lemma%unimportant homeomorphism_imp_covering_space: "homeomorphism S T f g \<Longrightarrow> covering_space S f T"
+lemma homeomorphism_imp_covering_space: "homeomorphism S T f g \<Longrightarrow> covering_space S f T"
   apply (simp add: homeomorphism_def covering_space_def, clarify)
   apply (rule_tac x=T in exI, simp)
   apply (rule_tac x="{S}" in exI, auto)
   done
 
-lemma%unimportant covering_space_local_homeomorphism:
+lemma covering_space_local_homeomorphism:
   assumes "covering_space c p S" "x \<in> c"
   obtains T u q where "x \<in> T" "openin (subtopology euclidean c) T"
                       "p x \<in> u" "openin (subtopology euclidean S) u"
@@ -1304,13 +1304,13 @@
   by (metis IntI UnionE vimage_eq) 
 
 
-lemma%important covering_space_local_homeomorphism_alt:
+lemma covering_space_local_homeomorphism_alt:
   assumes p: "covering_space c p S" and "y \<in> S"
   obtains x T U q where "p x = y"
                         "x \<in> T" "openin (subtopology euclidean c) T"
                         "y \<in> U" "openin (subtopology euclidean S) U"
                           "homeomorphism T U p q"
-proof%unimportant -
+proof -
   obtain x where "p x = y" "x \<in> c"
     using assms covering_space_imp_surjective by blast
   show ?thesis
@@ -1318,11 +1318,11 @@
     using that \<open>p x = y\<close> by blast
 qed
 
-proposition%important covering_space_open_map:
+proposition covering_space_open_map:
   fixes S :: "'a :: metric_space set" and T :: "'b :: metric_space set"
   assumes p: "covering_space c p S" and T: "openin (subtopology euclidean c) T"
     shows "openin (subtopology euclidean S) (p ` T)"
-proof%unimportant -
+proof -
   have pce: "p ` c = S"
    and covs:
        "\<And>x. x \<in> S \<Longrightarrow>
@@ -1368,7 +1368,7 @@
   with openin_subopen show ?thesis by blast
 qed
 
-lemma%important covering_space_lift_unique_gen:
+lemma covering_space_lift_unique_gen:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
   fixes g1 :: "'a \<Rightarrow> 'c::real_normed_vector"
   assumes cov: "covering_space c p S"
@@ -1380,7 +1380,7 @@
       and fg2: "\<And>x. x \<in> T \<Longrightarrow> f x = p(g2 x)"
       and u_compt: "U \<in> components T" and "a \<in> U" "x \<in> U"
     shows "g1 x = g2 x"
-proof%unimportant -
+proof -
   have "U \<subseteq> T" by (rule in_components_subset [OF u_compt])
   define G12 where "G12 \<equiv> {x \<in> U. g1 x - g2 x = 0}"
   have "connected U" by (rule in_components_connected [OF u_compt])
@@ -1427,7 +1427,7 @@
     using \<open>x \<in> U\<close> by force
 qed
 
-proposition%important covering_space_lift_unique:
+proposition covering_space_lift_unique:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
   fixes g1 :: "'a \<Rightarrow> 'c::real_normed_vector"
   assumes "covering_space c p S"
@@ -1437,10 +1437,10 @@
           "continuous_on T g2"  "g2 ` T \<subseteq> c"  "\<And>x. x \<in> T \<Longrightarrow> f x = p(g2 x)"
           "connected T"  "a \<in> T"   "x \<in> T"
    shows "g1 x = g2 x"
-  using%unimportant covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv
-  by%unimportant blast
+  using covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv
+  by blast
 
-lemma%unimportant covering_space_locally:
+lemma covering_space_locally:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes loc: "locally \<phi> C" and cov: "covering_space C p S"
       and pim: "\<And>T. \<lbrakk>T \<subseteq> C; \<phi> T\<rbrakk> \<Longrightarrow> \<psi>(p ` T)"
@@ -1456,14 +1456,14 @@
 qed
 
 
-proposition%important covering_space_locally_eq:
+proposition covering_space_locally_eq:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes cov: "covering_space C p S"
       and pim: "\<And>T. \<lbrakk>T \<subseteq> C; \<phi> T\<rbrakk> \<Longrightarrow> \<psi>(p ` T)"
       and qim: "\<And>q U. \<lbrakk>U \<subseteq> S; continuous_on U q; \<psi> U\<rbrakk> \<Longrightarrow> \<phi>(q ` U)"
     shows "locally \<psi> S \<longleftrightarrow> locally \<phi> C"
          (is "?lhs = ?rhs")
-proof%unimportant
+proof
   assume L: ?lhs
   show ?rhs
   proof (rule locallyI)
@@ -1518,7 +1518,7 @@
     using cov covering_space_locally pim by blast
 qed
 
-lemma%unimportant covering_space_locally_compact_eq:
+lemma covering_space_locally_compact_eq:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "covering_space C p S"
   shows "locally compact S \<longleftrightarrow> locally compact C"
@@ -1526,7 +1526,7 @@
    apply (meson assms compact_continuous_image continuous_on_subset covering_space_imp_continuous)
   using compact_continuous_image by blast
 
-lemma%unimportant covering_space_locally_connected_eq:
+lemma covering_space_locally_connected_eq:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "covering_space C p S"
     shows "locally connected S \<longleftrightarrow> locally connected C"
@@ -1534,7 +1534,7 @@
    apply (meson connected_continuous_image assms continuous_on_subset covering_space_imp_continuous)
   using connected_continuous_image by blast
 
-lemma%unimportant covering_space_locally_path_connected_eq:
+lemma covering_space_locally_path_connected_eq:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "covering_space C p S"
     shows "locally path_connected S \<longleftrightarrow> locally path_connected C"
@@ -1543,26 +1543,26 @@
   using path_connected_continuous_image by blast
 
 
-lemma%unimportant covering_space_locally_compact:
+lemma covering_space_locally_compact:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "locally compact C" "covering_space C p S"
   shows "locally compact S"
   using assms covering_space_locally_compact_eq by blast
 
 
-lemma%unimportant covering_space_locally_connected:
+lemma covering_space_locally_connected:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "locally connected C" "covering_space C p S"
   shows "locally connected S"
   using assms covering_space_locally_connected_eq by blast
 
-lemma%unimportant covering_space_locally_path_connected:
+lemma covering_space_locally_path_connected:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "locally path_connected C" "covering_space C p S"
   shows "locally path_connected S"
   using assms covering_space_locally_path_connected_eq by blast
 
-proposition%important covering_space_lift_homotopy:
+proposition covering_space_lift_homotopy:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
     and h :: "real \<times> 'c::real_normed_vector \<Rightarrow> 'b"
   assumes cov: "covering_space C p S"
@@ -1574,7 +1574,7 @@
                     "k ` ({0..1} \<times> U) \<subseteq> C"
                     "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = f y"
                     "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p(k z)"
-proof%unimportant -
+proof -
   have "\<exists>V k. openin (subtopology euclidean U) V \<and> y \<in> V \<and>
               continuous_on ({0..1} \<times> V) k \<and> k ` ({0..1} \<times> V) \<subseteq> C \<and>
               (\<forall>z \<in> V. k(0, z) = f z) \<and> (\<forall>z \<in> {0..1} \<times> V. h z = p(k z))"
@@ -1912,7 +1912,7 @@
   qed (auto simp: contk)
 qed
 
-corollary%important covering_space_lift_homotopy_alt:
+corollary covering_space_lift_homotopy_alt:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
     and h :: "'c::real_normed_vector \<times> real \<Rightarrow> 'b"
   assumes cov: "covering_space C p S"
@@ -1924,7 +1924,7 @@
                   "k ` (U \<times> {0..1}) \<subseteq> C"
                   "\<And>y. y \<in> U \<Longrightarrow> k(y, 0) = f y"
                   "\<And>z. z \<in> U \<times> {0..1} \<Longrightarrow> h z = p(k z)"
-proof%unimportant -
+proof -
   have "continuous_on ({0..1} \<times> U) (h \<circ> (\<lambda>z. (snd z, fst z)))"
     by (intro continuous_intros continuous_on_subset [OF conth]) auto
   then obtain k where contk: "continuous_on ({0..1} \<times> U) k"
@@ -1940,7 +1940,7 @@
     done
 qed
 
-corollary%important covering_space_lift_homotopic_function:
+corollary covering_space_lift_homotopic_function:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" and g:: "'c::real_normed_vector \<Rightarrow> 'a"
   assumes cov: "covering_space C p S"
       and contg: "continuous_on U g"
@@ -1948,7 +1948,7 @@
       and pgeq: "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
       and hom: "homotopic_with (\<lambda>x. True) U S f f'"
     obtains g' where "continuous_on U g'" "image g' U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g' y) = f' y"
-proof%unimportant -
+proof -
   obtain h where conth: "continuous_on ({0..1::real} \<times> U) h"
              and him: "h ` ({0..1} \<times> U) \<subseteq> S"
              and h0:  "\<And>x. h(0, x) = f x"
@@ -1972,12 +1972,12 @@
   qed
 qed
 
-corollary%important covering_space_lift_inessential_function:
+corollary covering_space_lift_inessential_function:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" and U :: "'c::real_normed_vector set"
   assumes cov: "covering_space C p S"
       and hom: "homotopic_with (\<lambda>x. True) U S f (\<lambda>x. a)"
   obtains g where "continuous_on U g" "g ` U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
-proof%unimportant (cases "U = {}")
+proof (cases "U = {}")
   case True
   then show ?thesis
     using that continuous_on_empty by blast
@@ -1997,14 +1997,14 @@
 
 subsection%important\<open> Lifting of general functions to covering space\<close>
 
-proposition%important covering_space_lift_path_strong:
+proposition covering_space_lift_path_strong:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
     and f :: "'c::real_normed_vector \<Rightarrow> 'b"
   assumes cov: "covering_space C p S" and "a \<in> C"
       and "path g" and pag: "path_image g \<subseteq> S" and pas: "pathstart g = p a"
     obtains h where "path h" "path_image h \<subseteq> C" "pathstart h = a"
                 and "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = g t"
-proof%unimportant -
+proof -
   obtain k:: "real \<times> 'c \<Rightarrow> 'a"
     where contk: "continuous_on ({0..1} \<times> {undefined}) k"
       and kim: "k ` ({0..1} \<times> {undefined}) \<subseteq> C"
@@ -2033,11 +2033,11 @@
   qed
 qed
 
-corollary%important covering_space_lift_path:
+corollary covering_space_lift_path:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes cov: "covering_space C p S" and "path g" and pig: "path_image g \<subseteq> S"
   obtains h where "path h" "path_image h \<subseteq> C" "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = g t"
-proof%unimportant -
+proof -
   obtain a where "a \<in> C" "pathstart g = p a"
     by (metis pig cov covering_space_imp_surjective imageE pathstart_in_path_image subsetCE)
   show ?thesis
@@ -2046,7 +2046,7 @@
 qed
 
   
-proposition%important covering_space_lift_homotopic_paths:
+proposition covering_space_lift_homotopic_paths:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes cov: "covering_space C p S"
       and "path g1" and pig1: "path_image g1 \<subseteq> S"
@@ -2056,7 +2056,7 @@
       and "path h2" and pih2: "path_image h2 \<subseteq> C" and ph2: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h2 t) = g2 t"
       and h1h2: "pathstart h1 = pathstart h2"
     shows "homotopic_paths C h1 h2"
-proof%unimportant -
+proof -
   obtain h :: "real \<times> real \<Rightarrow> 'b"
      where conth: "continuous_on ({0..1} \<times> {0..1}) h"
        and him: "h ` ({0..1} \<times> {0..1}) \<subseteq> S"
@@ -2117,7 +2117,7 @@
 qed
 
 
-corollary%important covering_space_monodromy:
+corollary covering_space_monodromy:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes cov: "covering_space C p S"
       and "path g1" and pig1: "path_image g1 \<subseteq> S"
@@ -2131,7 +2131,7 @@
   by%unimportant blast
 
 
-corollary%important covering_space_lift_homotopic_path:
+corollary covering_space_lift_homotopic_path:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes cov: "covering_space C p S"
       and hom: "homotopic_paths S f f'"
@@ -2140,7 +2140,7 @@
       and pgeq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(g t) = f t"
   obtains g' where "path g'" "path_image g' \<subseteq> C"
                    "pathstart g' = a" "pathfinish g' = b" "\<And>t. t \<in> {0..1} \<Longrightarrow> p(g' t) = f' t"
-proof%unimportant (rule covering_space_lift_path_strong [OF cov, of a f'])
+proof (rule covering_space_lift_path_strong [OF cov, of a f'])
   show "a \<in> C"
     using a pig by auto
   show "path f'" "path_image f' \<subseteq> S"
@@ -2150,7 +2150,7 @@
 qed (metis (mono_tags, lifting) assms cov covering_space_monodromy hom homotopic_paths_imp_path homotopic_paths_imp_subset pgeq pig)
 
 
-proposition%important covering_space_lift_general:
+proposition covering_space_lift_general:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
     and f :: "'c::real_normed_vector \<Rightarrow> 'b"
   assumes cov: "covering_space C p S" and "a \<in> C" "z \<in> U"
@@ -2162,7 +2162,7 @@
                              pathstart q = a \<and> pathfinish q = a \<and>
                              homotopic_paths S (f \<circ> r) (p \<circ> q)"
   obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
-proof%unimportant -
+proof -
   have *: "\<exists>g h. path g \<and> path_image g \<subseteq> U \<and>
                  pathstart g = z \<and> pathfinish g = y \<and>
                  path h \<and> path_image h \<subseteq> C \<and> pathstart h = a \<and>
@@ -2405,7 +2405,7 @@
 qed
 
 
-corollary%important covering_space_lift_stronger:
+corollary covering_space_lift_stronger:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
     and f :: "'c::real_normed_vector \<Rightarrow> 'b"
   assumes cov: "covering_space C p S" "a \<in> C" "z \<in> U"
@@ -2415,7 +2415,7 @@
       and hom: "\<And>r. \<lbrakk>path r; path_image r \<subseteq> U; pathstart r = z; pathfinish r = z\<rbrakk>
                      \<Longrightarrow> \<exists>b. homotopic_paths S (f \<circ> r) (linepath b b)"
   obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
-proof%unimportant (rule covering_space_lift_general [OF cov U contf fim feq])
+proof (rule covering_space_lift_general [OF cov U contf fim feq])
   fix r
   assume "path r" "path_image r \<subseteq> U" "pathstart r = z" "pathfinish r = z"
   then obtain b where b: "homotopic_paths S (f \<circ> r) (linepath b b)"
@@ -2432,7 +2432,7 @@
     by (force simp: \<open>a \<in> C\<close>)
 qed auto
 
-corollary%important covering_space_lift_strong:
+corollary covering_space_lift_strong:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
     and f :: "'c::real_normed_vector \<Rightarrow> 'b"
   assumes cov: "covering_space C p S" "a \<in> C" "z \<in> U"
@@ -2440,7 +2440,7 @@
       and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
       and feq: "f z = p a"
   obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
-proof%unimportant (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq])
+proof (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq])
   show "path_connected U"
     using scU simply_connected_eq_contractible_loop_some by blast
   fix r
@@ -2453,14 +2453,14 @@
     by blast
 qed blast
 
-corollary%important covering_space_lift:
+corollary covering_space_lift:
   fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
     and f :: "'c::real_normed_vector \<Rightarrow> 'b"
   assumes cov: "covering_space C p S"
       and U: "simply_connected U"  "locally path_connected U"
       and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
     obtains g where "continuous_on U g" "g ` U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
-proof%unimportant (cases "U = {}")
+proof (cases "U = {}")
   case True
   with that show ?thesis by auto
 next
--- a/src/HOL/Analysis/Improper_Integral.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Improper_Integral.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -28,11 +28,11 @@
   unfolding equiintegrable_on_def Ball_def
   by (erule conj_forward imp_forward all_forward ex_forward | blast)+
 
-lemma%important equiintegrable_on_Un:
+lemma equiintegrable_on_Un:
   assumes "F equiintegrable_on I" "G equiintegrable_on I"
   shows "(F \<union> G) equiintegrable_on I"
   unfolding equiintegrable_on_def
-proof%unimportant (intro conjI impI allI)
+proof (intro conjI impI allI)
   show "\<forall>f\<in>F \<union> G. f integrable_on I"
     using assms unfolding equiintegrable_on_def by blast
   show "\<exists>\<gamma>. gauge \<gamma> \<and>
@@ -76,12 +76,12 @@
 
 text\<open> Main limit theorem for an equiintegrable sequence.\<close>
 
-theorem%important equiintegrable_limit:
+theorem equiintegrable_limit:
   fixes g :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
   assumes feq: "range f equiintegrable_on cbox a b"
       and to_g: "\<And>x. x \<in> cbox a b \<Longrightarrow> (\<lambda>n. f n x) \<longlonglongrightarrow> g x"
     shows "g integrable_on cbox a b \<and> (\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> integral (cbox a b) g"
-proof%unimportant -
+proof -
   have "Cauchy (\<lambda>n. integral(cbox a b) (f n))"
   proof (clarsimp simp add: Cauchy_def)
     fix e::real
@@ -151,10 +151,10 @@
 qed
 
 
-lemma%important equiintegrable_reflect:
+lemma equiintegrable_reflect:
   assumes "F equiintegrable_on cbox a b"
   shows "(\<lambda>f. f \<circ> uminus) ` F equiintegrable_on cbox (-b) (-a)"
-proof%unimportant -
+proof -
   have "\<exists>\<gamma>. gauge \<gamma> \<and>
             (\<forall>f \<D>. f \<in> (\<lambda>f. f \<circ> uminus) ` F \<and> \<D> tagged_division_of cbox (- b) (- a) \<and> \<gamma> fine \<D> \<longrightarrow>
                    norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox (- b) (- a)) f) < e)"
@@ -246,13 +246,13 @@
 qed
 
 
-lemma%important content_division_lemma1:
+lemma content_division_lemma1:
   assumes div: "\<D> division_of S" and S: "S \<subseteq> cbox a b" and i: "i \<in> Basis"
       and mt: "\<And>K. K \<in> \<D> \<Longrightarrow> content K \<noteq> 0"
       and disj: "(\<forall>K \<in> \<D>. K \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}) \<or> (\<forall>K \<in> \<D>. K \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {})"
    shows "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
           \<le> content(cbox a b)"   (is "?lhs \<le> ?rhs")
-proof%unimportant -
+proof -
   have "finite \<D>"
     using div by blast
   define extend where
@@ -409,13 +409,13 @@
 qed
 
 
-proposition%important sum_content_area_over_thin_division:
+proposition sum_content_area_over_thin_division:
   assumes div: "\<D> division_of S" and S: "S \<subseteq> cbox a b" and i: "i \<in> Basis"
     and "a \<bullet> i \<le> c" "c \<le> b \<bullet> i"
     and nonmt: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<inter> {x. x \<bullet> i = c} \<noteq> {}"
   shows "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
           \<le> 2 * content(cbox a b)"
-proof%unimportant (cases "content(cbox a b) = 0")
+proof (cases "content(cbox a b) = 0")
   case True
   have "(\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) = 0"
     using S div by (force intro!: sum.neutral content_0_subset [OF True])
@@ -609,7 +609,7 @@
 
 
 
-proposition%important bounded_equiintegral_over_thin_tagged_partial_division:
+proposition bounded_equiintegral_over_thin_tagged_partial_division:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F" and "0 < \<epsilon>"
       and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
@@ -617,7 +617,7 @@
              "\<And>c i S h. \<lbrakk>c \<in> cbox a b; i \<in> Basis; S tagged_partial_division_of cbox a b;
                          \<gamma> fine S; h \<in> F; \<And>x K. (x,K) \<in> S \<Longrightarrow> (K \<inter> {x. x \<bullet> i = c \<bullet> i} \<noteq> {})\<rbrakk>
                         \<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
-proof%unimportant (cases "content(cbox a b) = 0")
+proof (cases "content(cbox a b) = 0")
   case True
   show ?thesis
   proof
@@ -813,13 +813,13 @@
 
 
 
-proposition%important equiintegrable_halfspace_restrictions_le:
+proposition equiintegrable_halfspace_restrictions_le:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
     and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
   shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i \<le> c then h x else 0)})
          equiintegrable_on cbox a b"
-proof%unimportant (cases "content(cbox a b) = 0")
+proof (cases "content(cbox a b) = 0")
   case True
   then show ?thesis by simp
 next
@@ -1172,13 +1172,13 @@
 
 
 
-corollary%important equiintegrable_halfspace_restrictions_ge:
+corollary equiintegrable_halfspace_restrictions_ge:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
     and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
   shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i \<ge> c then h x else 0)})
          equiintegrable_on cbox a b"
-proof%unimportant -
+proof -
   have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>(\<lambda>f. f \<circ> uminus) ` F. {\<lambda>x. if x \<bullet> i \<le> c then h x else 0})
            equiintegrable_on  cbox (- b) (- a)"
   proof (rule equiintegrable_halfspace_restrictions_le)
@@ -1203,11 +1203,11 @@
 qed
 
 
-proposition%important equiintegrable_closed_interval_restrictions:
+proposition equiintegrable_closed_interval_restrictions:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes f: "f integrable_on cbox a b"
   shows "(\<Union>c d. {(\<lambda>x. if x \<in> cbox c d then f x else 0)}) equiintegrable_on cbox a b"
-proof%unimportant -
+proof -
   let ?g = "\<lambda>B c d x. if \<forall>i\<in>B. c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i then f x else 0"
   have *: "insert f (\<Union>c d. {?g B c d}) equiintegrable_on cbox a b" if "B \<subseteq> Basis" for B
   proof -
@@ -1266,14 +1266,14 @@
 
 subsection%important\<open>Continuity of the indefinite integral\<close>
 
-proposition%important indefinite_integral_continuous:
+proposition indefinite_integral_continuous:
   fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
   assumes int_f: "f integrable_on cbox a b"
       and c: "c \<in> cbox a b" and d: "d \<in> cbox a b" "0 < \<epsilon>"
   obtains \<delta> where "0 < \<delta>"
               "\<And>c' d'. \<lbrakk>c' \<in> cbox a b; d' \<in> cbox a b; norm(c' - c) \<le> \<delta>; norm(d' - d) \<le> \<delta>\<rbrakk>
                         \<Longrightarrow> norm(integral(cbox c' d') f - integral(cbox c d) f) < \<epsilon>"
-proof%unimportant -
+proof -
   { assume "\<exists>c' d'. c' \<in> cbox a b \<and> d' \<in> cbox a b \<and> norm(c' - c) \<le> \<delta> \<and> norm(d' - d) \<le> \<delta> \<and>
                     norm(integral(cbox c' d') f - integral(cbox c d) f) \<ge> \<epsilon>"
                     (is "\<exists>c' d'. ?\<Phi> c' d' \<delta>") if "0 < \<delta>" for \<delta>
@@ -1358,11 +1358,11 @@
     by (meson not_le that)
 qed
 
-corollary%important indefinite_integral_uniformly_continuous:
+corollary indefinite_integral_uniformly_continuous:
   fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
   assumes "f integrable_on cbox a b"
   shows "uniformly_continuous_on (cbox (Pair a a) (Pair b b)) (\<lambda>y. integral (cbox (fst y) (snd y)) f)"
-proof%unimportant -
+proof -
   show ?thesis
   proof (rule compact_uniformly_continuous, clarsimp simp add: continuous_on_iff)
     fix c d and \<epsilon>::real
@@ -1383,11 +1383,11 @@
 qed
 
 
-corollary%important bounded_integrals_over_subintervals:
+corollary bounded_integrals_over_subintervals:
   fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
   assumes "f integrable_on cbox a b"
   shows "bounded {integral (cbox c d) f |c d. cbox c d \<subseteq> cbox a b}"
-proof%unimportant -
+proof -
   have "bounded ((\<lambda>y. integral (cbox (fst y) (snd y)) f) ` cbox (a, a) (b, b))"
        (is "bounded ?I")
     by (blast intro: bounded_cbox bounded_uniformly_continuous_image indefinite_integral_uniformly_continuous [OF assms])
@@ -1414,7 +1414,7 @@
 We only need to assume that the integrals are bounded, and we get absolute integrability,
 but we also need a (rather weak) bound assumption on the function.\<close>
 
-theorem%important absolutely_integrable_improper:
+theorem absolutely_integrable_improper:
   fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
   assumes int_f: "\<And>c d. cbox c d \<subseteq> box a b \<Longrightarrow> f integrable_on cbox c d"
       and bo: "bounded {integral (cbox c d) f |c d. cbox c d \<subseteq> box a b}"
@@ -1422,7 +1422,7 @@
           \<Longrightarrow> \<exists>g. g absolutely_integrable_on cbox a b \<and>
                   ((\<forall>x \<in> cbox a b. f x \<bullet> i \<le> g x) \<or> (\<forall>x \<in> cbox a b. f x \<bullet> i \<ge> g x))"
       shows "f absolutely_integrable_on cbox a b"
-proof%unimportant (cases "content(cbox a b) = 0")
+proof (cases "content(cbox a b) = 0")
   case True
   then show ?thesis
     by auto
--- a/src/HOL/Analysis/Interval_Integral.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Interval_Integral.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -4,15 +4,15 @@
 Lebesgue integral over an interval (with endpoints possibly +-\<infinity>)
 *)
 
-theory Interval_Integral
+theory Interval_Integral (*FIX ME rename? Lebesgue  *)
   imports Equivalence_Lebesgue_Henstock_Integration
 begin
 
-lemma%important continuous_on_vector_derivative:
+lemma continuous_on_vector_derivative:
   "(\<And>x. x \<in> S \<Longrightarrow> (f has_vector_derivative f' x) (at x within S)) \<Longrightarrow> continuous_on S f"
-  by%unimportant (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous)
+  by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous)
 
-definition%important "einterval a b = {x. a < ereal x \<and> ereal x < b}"
+definition "einterval a b = {x. a < ereal x \<and> ereal x < b}"
 
 lemma einterval_eq[simp]:
   shows einterval_eq_Icc: "einterval (ereal a) (ereal b) = {a <..< b}"
@@ -37,16 +37,16 @@
 lemma borel_einterval[measurable]: "einterval a b \<in> sets borel"
   unfolding einterval_def by measurable
 
-subsection\<open>Approximating a (possibly infinite) interval\<close>
+subsection%important \<open>Approximating a (possibly infinite) interval\<close>
 
 lemma filterlim_sup1: "(LIM x F. f x :> G1) \<Longrightarrow> (LIM x F. f x :> (sup G1 G2))"
  unfolding filterlim_def by (auto intro: le_supI1)
 
-lemma%important ereal_incseq_approx:
+lemma ereal_incseq_approx:
   fixes a b :: ereal
   assumes "a < b"
   obtains X :: "nat \<Rightarrow> real" where "incseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X \<longlonglongrightarrow> b"
-proof%unimportant (cases b)
+proof (cases b)
   case PInf
   with \<open>a < b\<close> have "a = -\<infinity> \<or> (\<exists>r. a = ereal r)"
     by (cases a) auto
@@ -78,12 +78,12 @@
        (auto simp: real incseq_def intro!: divide_left_mono)
 qed (insert \<open>a < b\<close>, auto)
 
-lemma%important ereal_decseq_approx:
+lemma ereal_decseq_approx:
   fixes a b :: ereal
   assumes "a < b"
   obtains X :: "nat \<Rightarrow> real" where
     "decseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X \<longlonglongrightarrow> a"
-proof%unimportant -
+proof -
   have "-b < -a" using \<open>a < b\<close> by simp
   from ereal_incseq_approx[OF this] guess X .
   then show thesis
@@ -93,14 +93,14 @@
     done
 qed
 
-lemma%important einterval_Icc_approximation:
+proposition einterval_Icc_approximation:
   fixes a b :: ereal
   assumes "a < b"
   obtains u l :: "nat \<Rightarrow> real" where
     "einterval a b = (\<Union>i. {l i .. u i})"
     "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
     "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
-proof%unimportant -
+proof -
   from dense[OF \<open>a < b\<close>] obtain c where "a < c" "c < b" by safe
   from ereal_incseq_approx[OF \<open>c < b\<close>] guess u . note u = this
   from ereal_decseq_approx[OF \<open>a < c\<close>] guess l . note l = this
@@ -202,17 +202,17 @@
     interval_lebesgue_integrable M a b (\<lambda>x. c * f x)"
   by (simp add: interval_lebesgue_integrable_def)
 
-lemma%important interval_lebesgue_integrable_mult_left [intro, simp]:
+lemma interval_lebesgue_integrable_mult_left [intro, simp]:
   fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
   shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
     interval_lebesgue_integrable M a b (\<lambda>x. f x * c)"
-  by%unimportant (simp add: interval_lebesgue_integrable_def)
+  by (simp add: interval_lebesgue_integrable_def)
 
-lemma%important interval_lebesgue_integrable_divide [intro, simp]:
+lemma interval_lebesgue_integrable_divide [intro, simp]:
   fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, field, second_countable_topology}"
   shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
     interval_lebesgue_integrable M a b (\<lambda>x. f x / c)"
-  by%unimportant (simp add: interval_lebesgue_integrable_def)
+  by (simp add: interval_lebesgue_integrable_def)
 
 lemma interval_lebesgue_integral_mult_right [simp]:
   fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
@@ -220,11 +220,11 @@
     c * interval_lebesgue_integral M a b f"
   by (simp add: interval_lebesgue_integral_def)
 
-lemma%important interval_lebesgue_integral_mult_left [simp]:
+lemma interval_lebesgue_integral_mult_left [simp]:
   fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
   shows "interval_lebesgue_integral M a b (\<lambda>x. f x * c) =
     interval_lebesgue_integral M a b f * c"
-  by%unimportant (simp add: interval_lebesgue_integral_def)
+  by (simp add: interval_lebesgue_integral_def)
 
 lemma interval_lebesgue_integral_divide [simp]:
   fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, field, second_countable_topology}"
@@ -242,11 +242,11 @@
   unfolding interval_lebesgue_integral_def
   by (auto simp: interval_lebesgue_integral_def set_integral_complex_of_real)
 
-lemma%important interval_lebesgue_integral_le_eq:
+lemma interval_lebesgue_integral_le_eq:
   fixes a b f
   assumes "a \<le> b"
   shows "interval_lebesgue_integral M a b f = (LINT x : einterval a b | M. f x)"
-using%unimportant assms by (auto simp: interval_lebesgue_integral_def)
+  using assms by (auto simp: interval_lebesgue_integral_def)
 
 lemma interval_lebesgue_integral_gt_eq:
   fixes a b f
@@ -271,9 +271,9 @@
     interval_lebesgue_integrable lborel b a f"
   by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integrable_def einterval_same)
 
-lemma%important interval_integral_reflect:
+lemma interval_integral_reflect:
   "(LBINT x=a..b. f x) = (LBINT x=-b..-a. f (-x))"
-proof%unimportant (induct a b rule: linorder_wlog)
+proof (induct a b rule: linorder_wlog)
   case (sym a b) then show ?case
     by (auto simp: interval_lebesgue_integral_def interval_integrable_endpoints_reverse
              split: if_split_asm)
@@ -299,7 +299,7 @@
 lemma interval_integral_to_infinity_eq: "(LINT x=ereal a..\<infinity> | M. f x) = (LINT x : {a<..} | M. f x)"
   unfolding interval_lebesgue_integral_def by auto
 
-lemma%important interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \<infinity> f) =
+proposition interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \<infinity> f) =
   (set_integrable M {a<..} f)"
   unfolding%unimportant interval_lebesgue_integrable_def by auto
 
@@ -317,12 +317,12 @@
   unfolding interval_lebesgue_integral_def interval_lebesgue_integrable_def einterval_eq
   by (auto simp: less_imp_le field_simps measure_def set_integrable_def set_lebesgue_integral_def)
 
-lemma%important interval_integral_cong_AE:
+lemma interval_integral_cong_AE:
   assumes [measurable]: "f \<in> borel_measurable borel" "g \<in> borel_measurable borel"
   assumes "AE x \<in> einterval (min a b) (max a b) in lborel. f x = g x"
   shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g"
   using assms
-proof%unimportant (induct a b rule: linorder_wlog)
+proof (induct a b rule: linorder_wlog)
   case (sym a b) then show ?case
     by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b])
 next
@@ -331,11 +331,11 @@
              intro!: set_lebesgue_integral_cong_AE)
 qed
 
-lemma%important interval_integral_cong:
+lemma interval_integral_cong:
   assumes "\<And>x. x \<in> einterval (min a b) (max a b) \<Longrightarrow> f x = g x"
   shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g"
   using assms
-proof%unimportant (induct a b rule: linorder_wlog)
+proof (induct a b rule: linorder_wlog)
   case (sym a b) then show ?case
     by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b])
 next
@@ -407,11 +407,11 @@
   apply (auto simp: eq anti_eq einterval_iff split: split_indicator)
   done
 
-lemma%important interval_integral_sum:
+lemma interval_integral_sum:
   fixes a b c :: ereal
   assumes integrable: "interval_lebesgue_integrable lborel (min a (min b c)) (max a (max b c)) f"
   shows "(LBINT x=a..b. f x) + (LBINT x=b..c. f x) = (LBINT x=a..c. f x)"
-proof%unimportant -
+proof -
   let ?I = "\<lambda>a b. LBINT x=a..b. f x"
   { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \<le> b" "b \<le> c"
     then have ord: "a \<le> b" "b \<le> c" "a \<le> c" and f': "set_integrable lborel (einterval a c) f"
@@ -446,11 +446,11 @@
        (simp_all add: min_absorb1 min_absorb2 max_absorb1 max_absorb2 field_simps 1 2 3)
 qed
 
-lemma%important interval_integrable_isCont:
+lemma interval_integrable_isCont:
   fixes a b and f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
   shows "(\<And>x. min a b \<le> x \<Longrightarrow> x \<le> max a b \<Longrightarrow> isCont f x) \<Longrightarrow>
     interval_lebesgue_integrable lborel a b f"
-proof%unimportant (induct a b rule: linorder_wlog)
+proof (induct a b rule: linorder_wlog)
   case (le a b) then show ?case
     unfolding interval_lebesgue_integrable_def set_integrable_def
     by (auto simp: interval_lebesgue_integrable_def
@@ -478,7 +478,7 @@
 
 subsection%important\<open>General limit approximation arguments\<close>
 
-lemma%important interval_integral_Icc_approx_nonneg:
+proposition interval_integral_Icc_approx_nonneg:
   fixes a b :: ereal
   assumes "a < b"
   fixes u l :: "nat \<Rightarrow> real"
@@ -493,7 +493,7 @@
   shows
     "set_integrable lborel (einterval a b) f"
     "(LBINT x=a..b. f x) = C"
-proof%unimportant -
+proof -
   have 1 [unfolded set_integrable_def]: "\<And>i. set_integrable lborel {l i..u i} f" by (rule f_integrable)
   have 2: "AE x in lborel. mono (\<lambda>n. indicator {l n..u n} x *\<^sub>R f x)"
   proof -
@@ -534,7 +534,7 @@
     by (rule integrable_monotone_convergence[OF 1 2 3 4 5])
 qed
 
-lemma%important interval_integral_Icc_approx_integrable:
+proposition interval_integral_Icc_approx_integrable:
   fixes u l :: "nat \<Rightarrow> real" and a b :: ereal
   fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
   assumes "a < b"
@@ -543,7 +543,7 @@
     "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
   assumes f_integrable: "set_integrable lborel (einterval a b) f"
   shows "(\<lambda>i. LBINT x=l i.. u i. f x) \<longlonglongrightarrow> (LBINT x=a..b. f x)"
-proof%unimportant -
+proof -
   have "(\<lambda>i. LBINT x:{l i.. u i}. f x) \<longlonglongrightarrow> (LBINT x:einterval a b. f x)"
     unfolding set_lebesgue_integral_def
   proof (rule integral_dominated_convergence)
@@ -580,13 +580,13 @@
   TODO: make the older versions corollaries of these (using continuous_at_imp_continuous_on, etc.)
 *)
 
-lemma%important interval_integral_FTC_finite:
+lemma interval_integral_FTC_finite:
   fixes f F :: "real \<Rightarrow> 'a::euclidean_space" and a b :: real
   assumes f: "continuous_on {min a b..max a b} f"
   assumes F: "\<And>x. min a b \<le> x \<Longrightarrow> x \<le> max a b \<Longrightarrow> (F has_vector_derivative (f x)) (at x within
     {min a b..max a b})"
   shows "(LBINT x=a..b. f x) = F b - F a"
-proof%unimportant (cases "a \<le> b")
+proof (cases "a \<le> b")
   case True
   have "(LBINT x=a..b. f x) = (LBINT x. indicat_real {a..b} x *\<^sub>R f x)"
     by (simp add: True interval_integral_Icc set_lebesgue_integral_def)
@@ -617,7 +617,7 @@
 qed
 
 
-lemma%important interval_integral_FTC_nonneg:
+lemma interval_integral_FTC_nonneg:
   fixes f F :: "real \<Rightarrow> real" and a b :: ereal
   assumes "a < b"
   assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV F x :> f x"
@@ -628,7 +628,7 @@
   shows
     "set_integrable lborel (einterval a b) f"
     "(LBINT x=a..b. f x) = B - A"
-proof%unimportant -
+proof -
   obtain u l where approx:
     "einterval a b = (\<Union>i. {l i .. u i})"
     "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
@@ -669,7 +669,7 @@
     by (rule interval_integral_Icc_approx_nonneg [OF \<open>a < b\<close> approx 1 f_nonneg 2 3])
 qed
 
-lemma%important interval_integral_FTC_integrable:
+theorem interval_integral_FTC_integrable:
   fixes f F :: "real \<Rightarrow> 'a::euclidean_space" and a b :: ereal
   assumes "a < b"
   assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> (F has_vector_derivative f x) (at x)"
@@ -678,7 +678,7 @@
   assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
   assumes B: "((F \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
   shows "(LBINT x=a..b. f x) = B - A"
-proof%unimportant -
+proof -
   obtain u l where approx:
     "einterval a b = (\<Union>i. {l i .. u i})"
     "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
@@ -714,14 +714,14 @@
   einterval.
 *)
 
-lemma%important interval_integral_FTC2:
+theorem interval_integral_FTC2:
   fixes a b c :: real and f :: "real \<Rightarrow> 'a::euclidean_space"
   assumes "a \<le> c" "c \<le> b"
   and contf: "continuous_on {a..b} f"
   fixes x :: real
   assumes "a \<le> x" and "x \<le> b"
   shows "((\<lambda>u. LBINT y=c..u. f y) has_vector_derivative (f x)) (at x within {a..b})"
-proof%unimportant -
+proof -
   let ?F = "(\<lambda>u. LBINT y=a..u. f y)"
   have intf: "set_integrable lborel {a..b} f"
     by (rule borel_integrable_atLeastAtMost', rule contf)
@@ -746,11 +746,11 @@
   qed (insert assms, auto)
 qed
 
-lemma%important einterval_antiderivative:
+proposition einterval_antiderivative:
   fixes a b :: ereal and f :: "real \<Rightarrow> 'a::euclidean_space"
   assumes "a < b" and contf: "\<And>x :: real. a < x \<Longrightarrow> x < b \<Longrightarrow> isCont f x"
   shows "\<exists>F. \<forall>x :: real. a < x \<longrightarrow> x < b \<longrightarrow> (F has_vector_derivative f x) (at x)"
-proof%unimportant -
+proof -
   from einterval_nonempty [OF \<open>a < b\<close>] obtain c :: real where [simp]: "a < c" "c < b"
     by (auto simp: einterval_def)
   let ?F = "(\<lambda>u. LBINT y=c..u. f y)"
@@ -786,14 +786,14 @@
 text\<open>Once again, three versions: first, for finite intervals, and then two versions for
     arbitrary intervals.\<close>
 
-lemma%important interval_integral_substitution_finite:
+theorem interval_integral_substitution_finite:
   fixes a b :: real and f :: "real \<Rightarrow> 'a::euclidean_space"
   assumes "a \<le> b"
   and derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_real_derivative (g' x)) (at x within {a..b})"
   and contf : "continuous_on (g ` {a..b}) f"
   and contg': "continuous_on {a..b} g'"
   shows "LBINT x=a..b. g' x *\<^sub>R f (g x) = LBINT y=g a..g b. f y"
-proof%unimportant-
+proof-
   have v_derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_vector_derivative (g' x)) (at x within {a..b})"
     using derivg unfolding has_field_derivative_iff_has_vector_derivative .
   then have contg [simp]: "continuous_on {a..b} g"
@@ -833,7 +833,7 @@
 
 (* TODO: is it possible to lift the assumption here that g' is nonnegative? *)
 
-lemma%important interval_integral_substitution_integrable:
+theorem interval_integral_substitution_integrable:
   fixes f :: "real \<Rightarrow> 'a::euclidean_space" and a b u v :: ereal
   assumes "a < b"
   and deriv_g: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV g x :> g' x"
@@ -845,7 +845,7 @@
   and integrable: "set_integrable lborel (einterval a b) (\<lambda>x. g' x *\<^sub>R f (g x))"
   and integrable2: "set_integrable lborel (einterval A B) (\<lambda>x. f x)"
   shows "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))"
-proof%unimportant -
+proof -
   obtain u l where approx [simp]:
     "einterval a b = (\<Union>i. {l i .. u i})"
     "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
@@ -932,7 +932,7 @@
    An alternative: make the second one the main one, and then have another lemma
    that says that if f is nonnegative and all the other hypotheses hold, then it is integrable. *)
 
-lemma%important interval_integral_substitution_nonneg:
+theorem interval_integral_substitution_nonneg:
   fixes f g g':: "real \<Rightarrow> real" and a b u v :: ereal
   assumes "a < b"
   and deriv_g: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV g x :> g' x"
@@ -946,7 +946,7 @@
   shows
     "set_integrable lborel (einterval A B) f"
     "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))"
-proof%unimportant -
+proof -
   from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx [simp] = this
   note less_imp_le [simp]
   have aless[simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
@@ -1079,17 +1079,17 @@
 translations
   "CLBINT x=a..b. f" == "CONST complex_interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
 
-lemma%important interval_integral_norm:
+proposition interval_integral_norm:
   fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
   shows "interval_lebesgue_integrable lborel a b f \<Longrightarrow> a \<le> b \<Longrightarrow>
     norm (LBINT t=a..b. f t) \<le> LBINT t=a..b. norm (f t)"
   using%unimportant integral_norm_bound[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
   by%unimportant (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def)
 
-lemma%important interval_integral_norm2:
+proposition interval_integral_norm2:
   "interval_lebesgue_integrable lborel a b f \<Longrightarrow>
     norm (LBINT t=a..b. f t) \<le> \<bar>LBINT t=a..b. norm (f t)\<bar>"
-proof%unimportant (induct a b rule: linorder_wlog)
+proof (induct a b rule: linorder_wlog)
   case (sym a b) then show ?case
     by (simp add: interval_integral_endpoints_reverse[of a b] interval_integrable_endpoints_reverse[of a b])
 next
--- a/src/HOL/Analysis/Tagged_Division.thy	Thu Jan 17 16:22:21 2019 -0500
+++ b/src/HOL/Analysis/Tagged_Division.thy	Thu Jan 17 16:28:07 2019 -0500
@@ -32,7 +32,7 @@
 
 
 subsection%unimportant \<open>Sundries\<close>
-(*FIXME restructure this entire section consists of single lemma *)
+
 
 text\<open>A transitive relation is well-founded if all initial segments are finite.\<close>
 lemma wf_finite_segments:
@@ -55,7 +55,7 @@
     by (simp add: field_simps)
 qed
 
-subsection%unimportant \<open>Some useful lemmas about intervals\<close>
+subsection%important \<open>Some useful lemmas about intervals\<close>
 
 lemma interior_subset_union_intervals:
   assumes "i = cbox a b"
@@ -152,22 +152,22 @@
 
 lemmas interval_bounds = interval_upperbound interval_lowerbound
 
-lemma%important
+lemma
   fixes X::"real set"
   shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X"
     and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X"
-  by%unimportant (auto simp: interval_upperbound_def interval_lowerbound_def)
+  by (auto simp: interval_upperbound_def interval_lowerbound_def)
 
-lemma%important interval_bounds'[simp]:
+lemma interval_bounds'[simp]:
   assumes "cbox a b \<noteq> {}"
   shows "interval_upperbound (cbox a b) = b"
     and "interval_lowerbound (cbox a b) = a"
-  using%unimportant assms unfolding box_ne_empty by auto
+  using assms unfolding box_ne_empty by auto
 
-lemma%important interval_upperbound_Times:
+lemma interval_upperbound_Times:
   assumes "A \<noteq> {}" and "B \<noteq> {}"
   shows "interval_upperbound (A \<times> B) = (interval_upperbound A, interval_upperbound B)"
-proof%unimportant-
+proof-
   from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
   have "(\<Sum>i\<in>Basis. (SUP x\<in>A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x\<in>A. x \<bullet> i) *\<^sub>R i)"
       by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
@@ -178,10 +178,10 @@
       by (subst sum_Basis_prod_eq) (auto simp add: sum_prod)
 qed
 
-lemma%important interval_lowerbound_Times:
+lemma interval_lowerbound_Times:
   assumes "A \<noteq> {}" and "B \<noteq> {}"
   shows "interval_lowerbound (A \<times> B) = (interval_lowerbound A, interval_lowerbound B)"
-proof%unimportant-
+proof-
   from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
   have "(\<Sum>i\<in>Basis. (INF x\<in>A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x\<in>A. x \<bullet> i) *\<^sub>R i)"
       by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
@@ -226,11 +226,11 @@
   using equation_minus_iff
   by (auto simp: gauge_def surj_def intro!: open_surjective_linear_image linear_uminus)
 
-lemma%important gauge_Inter:
+lemma gauge_Inter:
   assumes "finite S"
     and "\<And>u. u\<in>S \<Longrightarrow> gauge (f u)"
   shows "gauge (\<lambda>x. \<Inter>{f u x | u. u \<in> S})"
-proof%unimportant -
+proof -
   have *: "\<And>x. {f u x |u. u \<in> S} = (\<lambda>u. f u x) ` S"
     by auto
   show ?thesis
@@ -238,12 +238,12 @@
     using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto
 qed
 
-lemma%important gauge_existence_lemma:
+lemma gauge_existence_lemma:
   "(\<forall>x. \<exists>d :: real. p x \<longrightarrow> 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. p x \<longrightarrow> q d x)"
-  by%unimportant (metis zero_less_one)
+  by (metis zero_less_one)
 
-subsection%unimportant \<open>Attempt a systematic general set of "offset" results for components\<close>
-(*FIXME Restructure *)
+subsection%important \<open>Attempt a systematic general set of "offset" results for components\<close>
+(*FIXME Restructure, the subsection consists only of 1 lemma *)
 lemma gauge_modify:
   assumes "(\<forall>S. open S \<longrightarrow> open {x. f(x) \<in> S})" "gauge d"
   shows "gauge (\<lambda>x. {y. f y \<in> d (f x)})"
@@ -605,10 +605,10 @@
   qed
 qed
 
-proposition%important partial_division_extend_interval:
+proposition partial_division_extend_interval:
   assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> cbox a b"
   obtains q where "p \<subseteq> q" "q division_of cbox a (b::'a::euclidean_space)"
-proof%unimportant (cases "p = {}")
+proof (cases "p = {}")
   case True
   obtain q where "q division_of (cbox a b)"
     by (rule elementary_interval)
@@ -687,11 +687,11 @@
   "p division_of S \<Longrightarrow> \<exists>a b. S \<subseteq> cbox a (b::'a::euclidean_space)"
   by (meson bounded_subset_cbox_symmetric elementary_bounded)
 
-lemma%important division_union_intervals_exists:
+proposition division_union_intervals_exists:
   fixes a b :: "'a::euclidean_space"
   assumes "cbox a b \<noteq> {}"
   obtains p where "(insert (cbox a b) p) division_of (cbox a b \<union> cbox c d)"
-proof%unimportant (cases "cbox c d = {}")
+proof (cases "cbox c d = {}")
   case True
   with assms that show ?thesis by force
 next
@@ -739,11 +739,11 @@
   shows "\<Union>f division_of \<Union>\<Union>f"
   using assms  by (auto intro!: division_ofI)
 
-lemma%important elementary_union_interval:
+lemma elementary_union_interval:
   fixes a b :: "'a::euclidean_space"
   assumes "p division_of \<Union>p"
   obtains q where "q division_of (cbox a b \<union> \<Union>p)"
-proof%unimportant (cases "p = {} \<or> cbox a b = {}")
+proof (cases "p = {} \<or> cbox a b = {}")
   case True
   obtain p where "p division_of (cbox a b)"
     by (rule elementary_interval)
@@ -855,11 +855,11 @@
     using that by auto
 qed
 
-lemma%important elementary_union:
+lemma elementary_union:
   fixes S T :: "'a::euclidean_space set"
   assumes "ps division_of S" "pt division_of T"
   obtains p where "p division_of (S \<union> T)"
-proof%unimportant -
+proof -
   have *: "S \<union> T = \<Union>ps \<union> \<Union>pt"
     using assms unfolding division_of_def by auto
   show ?thesis
@@ -868,13 +868,13 @@
     by (simp add: * that)
 qed
 
-lemma%important partial_division_extend:
+lemma partial_division_extend:
   fixes T :: "'a::euclidean_space set"
   assumes "p division_of S"
     and "q division_of T"
     and "S \<subseteq> T"
   obtains r where "p \<subseteq> r" and "r division_of T"
-proof%unimportant -
+proof -
   note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)]
   obtain a b where ab: "T \<subseteq> cbox a b"
     using elementary_subset_cbox[OF assms(2)] by auto
@@ -931,7 +931,7 @@
 qed
 
 
-lemma%important division_split:
+lemma division_split:
   fixes a :: "'a::euclidean_space"
   assumes "p division_of (cbox a b)"
     and k: "k\<in>Basis"
@@ -939,7 +939,7 @@
       (is "?p1 division_of ?I1")
     and "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
       (is "?p2 division_of ?I2")
-proof%unimportant (rule_tac[!] division_ofI)
+proof (rule_tac[!] division_ofI)
   note p = division_ofD[OF assms(1)]
   show "finite ?p1" "finite ?p2"
     using p(1) by auto
@@ -1020,14 +1020,14 @@
 lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s"
   unfolding tagged_division_of_def tagged_partial_division_of_def by auto
 
-lemma%important tagged_division_of:
+lemma tagged_division_of:
   "s tagged_division_of i \<longleftrightarrow>
     finite s \<and>
     (\<forall>x K. (x, K) \<in> s \<longrightarrow> x \<in> K \<and> K \<subseteq> i \<and> (\<exists>a b. K = cbox a b)) \<and>
     (\<forall>x1 K1 x2 K2. (x1, K1) \<in> s \<and> (x2, K2) \<in> s \<and> (x1, K1) \<noteq> (x2, K2) \<longrightarrow>
       interior K1 \<inter> interior K2 = {}) \<and>
     (\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
-  unfolding%unimportant tagged_division_of_def tagged_partial_division_of_def by auto
+  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
 
 lemma tagged_division_ofI:
   assumes "finite s"
@@ -1052,10 +1052,10 @@
     and "(\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
   using assms unfolding tagged_division_of by blast+
 
-lemma%important division_of_tagged_division:
+lemma division_of_tagged_division:
   assumes "s tagged_division_of i"
   shows "(snd ` s) division_of i"
-proof%unimportant (rule division_ofI)
+proof (rule division_ofI)
   note assm = tagged_division_ofD[OF assms]
   show "\<Union>(snd ` s) = i" "finite (snd ` s)"
     using assm by auto
@@ -1073,10 +1073,10 @@
     using assm(5) k'(2) xk by blast
 qed
 
-lemma%important partial_division_of_tagged_division:
+lemma partial_division_of_tagged_division:
   assumes "s tagged_partial_division_of i"
   shows "(snd ` s) division_of \<Union>(snd ` s)"
-proof%unimportant (rule division_ofI)
+proof (rule division_ofI)
   note assm = tagged_partial_division_ofD[OF assms]
   show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)"
     using assm by auto
@@ -1121,12 +1121,12 @@
   unfolding box_real[symmetric]
   by (rule tagged_division_of_self)
 
-lemma%important tagged_division_Un:
+lemma tagged_division_Un:
   assumes "p1 tagged_division_of s1"
     and "p2 tagged_division_of s2"
     and "interior s1 \<inter> interior s2 = {}"
   shows "(p1 \<union> p2) tagged_division_of (s1 \<union> s2)"
-proof%unimportant (rule tagged_division_ofI)
+proof (rule tagged_division_ofI)
   note p1 = tagged_division_ofD[OF assms(1)]
   note p2 = tagged_division_ofD[OF assms(2)]
   show "finite (p1 \<union> p2)"
@@ -1149,12 +1149,12 @@
     by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2))
 qed
 
-lemma%important tagged_division_Union:
+lemma tagged_division_Union:
   assumes "finite I"
     and tag: "\<And>i. i\<in>I \<Longrightarrow> pfn i tagged_division_of i"
     and disj: "\<And>i1 i2. \<lbrakk>i1 \<in> I; i2 \<in> I; i1 \<noteq> i2\<rbrakk> \<Longrightarrow> interior(i1) \<inter> interior(i2) = {}"
   shows "\<Union>(pfn ` I) tagged_division_of (\<Union>I)"
-proof%unimportant (rule tagged_division_ofI)
+proof (rule tagged_division_ofI)
   note assm = tagged_division_ofD[OF tag]
   show "finite (\<Union>(pfn ` I))"
     using assms by auto
@@ -1229,13 +1229,13 @@
   unfolding box_real[symmetric]
   by (rule tagged_division_Un_interval)
 
-lemma%important tagged_division_split_left_inj:
+lemma tagged_division_split_left_inj:
   assumes d: "d tagged_division_of i"
   and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
   and "K1 \<noteq> K2"
   and eq: "K1 \<inter> {x. x \<bullet> k \<le> c} = K2 \<inter> {x. x \<bullet> k \<le> c}"
     shows "interior (K1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
-proof%unimportant -
+proof -
   have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
     using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
   then show ?thesis
@@ -1255,11 +1255,11 @@
     using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
 qed
 
-lemma%important (in comm_monoid_set) over_tagged_division_lemma:
+lemma (in comm_monoid_set) over_tagged_division_lemma:
   assumes "p tagged_division_of i"
     and "\<And>u v. box u v = {} \<Longrightarrow> d (cbox u v) = \<^bold>1"
   shows "F (\<lambda>(_, k). d k) p = F d (snd ` p)"
-proof%unimportant -
+proof -
   have *: "(\<lambda>(_ ,k). d k) = d \<circ> snd"
     by (simp add: fun_eq_iff)
   note assm = tagged_division_ofD[OF assms(1)]
@@ -1293,7 +1293,7 @@
   \<^typ>\<open>bool\<close>.\<close>
 
 paragraph%important \<open>Using additivity of lifted function to encode definedness.\<close>
-text%important \<open>%whitespace\<close>
+text \<open>%whitespace\<close>
 definition%important lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option"
 where
   "lift_option f a' b' = Option.bind a' (\<lambda>a. Option.bind b' (\<lambda>b. Some (f a b)))"
@@ -1307,7 +1307,7 @@
 lemma%important comm_monoid_lift_option:
   assumes "comm_monoid f z"
   shows "comm_monoid (lift_option f) (Some z)"
-proof%unimportant -
+proof -
   from assms interpret comm_monoid f z .
   show ?thesis
     by standard (auto simp: lift_option_def ac_simps split: bind_split)
@@ -1334,16 +1334,16 @@
 
 
 paragraph \<open>Division points\<close>
-text%important \<open>%whitespace\<close>
+text \<open>%whitespace\<close>
 definition%important "division_points (k::('a::euclidean_space) set) d =
    {(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
      (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
 
-lemma%important division_points_finite:
+lemma division_points_finite:
   fixes i :: "'a::euclidean_space set"
   assumes "d division_of i"
   shows "finite (division_points i d)"
-proof%unimportant -
+proof -
   note assm = division_ofD[OF assms]
   let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)\<bullet>j < x \<and> x < (interval_upperbound i)\<bullet>j \<and>
     (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
@@ -1353,7 +1353,7 @@
     unfolding * using assm by auto
 qed
 
-lemma%important division_points_subset:
+lemma division_points_subset:
   fixes a :: "'a::euclidean_space"
   assumes "d division_of (cbox a b)"
     and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
@@ -1362,7 +1362,7 @@
       division_points (cbox a b) d" (is ?t1)
     and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> \<not>(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq>
       division_points (cbox a b) d" (is ?t2)
-proof%unimportant -
+proof -
   note assm = division_ofD[OF assms(1)]
   have *: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
     "\<forall>i\<in>Basis. a\<bullet>i \<le> (\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else  b \<bullet> i) *\<^sub>R i) \<bullet> i"
@@ -1423,7 +1423,7 @@
     by force
 qed
 
-lemma%important division_points_psubset:
+lemma division_points_psubset:
   fixes a :: "'a::euclidean_space"
   assumes d: "d division_of (cbox a b)"
       and altb: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
@@ -1434,7 +1434,7 @@
          division_points (cbox a b) d" (is "?D1 \<subset> ?D")
     and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} \<subset>
          division_points (cbox a b) d" (is "?D2 \<subset> ?D")
-proof%unimportant -
+proof -
   have ab: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
     using altb by (auto intro!:less_imp_le)
   obtain u v where l: "l = cbox u v"
@@ -1485,13 +1485,13 @@
     using \<open>K1 \<in> \<D>\<close> \<open>K1 \<noteq> K2\<close> by auto
 qed
 
-lemma%important division_split_right_inj:
+lemma division_split_right_inj:
   fixes S :: "'a::euclidean_space set"
   assumes div: "\<D> division_of S"
     and eq: "K1 \<inter> {x::'a. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}"
     and "K1 \<in> \<D>" "K2 \<in> \<D>" "K1 \<noteq> K2"
   shows "interior (K1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
-proof%unimportant -
+proof -
   have "interior K2 \<inter> interior {a. a \<bullet> k \<ge> c} = interior K1 \<inter> interior {a. a \<bullet> k \<ge> c}"
     by (metis (no_types) eq interior_Int)
   moreover have "\<And>A. interior A \<inter> interior K2 = {} \<or> A = K2 \<or> A \<notin> \<D>"
@@ -1515,13 +1515,13 @@
     unfolding * ** interval_split[OF assms] by (rule refl)
 qed
 
-lemma%important division_doublesplit:
+lemma division_doublesplit:
   fixes a :: "'a::euclidean_space"
   assumes "p division_of (cbox a b)"
     and k: "k \<in> Basis"
   shows "(\<lambda>l. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e}) ` {l\<in>p. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e} \<noteq> {}}
          division_of  (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e})"
-proof%unimportant -
+proof -
   have *: "\<And>x c. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
     by auto
   have **: "\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'"
@@ -1560,9 +1560,9 @@
     using box_empty_imp [of One "-One"] by simp
 qed
        
-lemma%important division:
+lemma division:
   "F g d = g (cbox a b)" if "d division_of (cbox a b)"
-proof%unimportant -
+proof -
   define C where [abs_def]: "C = card (division_points (cbox a b) d)"
   then show ?thesis
   using that proof (induction C arbitrary: a b d rule: less_induct)
@@ -1747,10 +1747,10 @@
   qed
 qed
 
-lemma%important tagged_division:
+proposition tagged_division:
   assumes "d tagged_division_of (cbox a b)"
   shows "F (\<lambda>(_, l). g l) d = g (cbox a b)"
-proof%unimportant -
+proof -
   have "F (\<lambda>(_, k). g k) d = F g (snd ` d)"
     using assms box_empty_imp by (rule over_tagged_division_lemma)
   then show ?thesis
@@ -1830,12 +1830,12 @@
 
 subsection%important \<open>Special case of additivity we need for the FTC\<close>
 (*fix add explanation here *)
-lemma%important additive_tagged_division_1:
+lemma additive_tagged_division_1:
   fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
   assumes "a \<le> b"
     and "p tagged_division_of {a..b}"
   shows "sum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
-proof%unimportant -
+proof -
   let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
   interpret operative_real plus 0 ?f
     rewrites "comm_monoid_set.F (+) 0 = sum"
@@ -1881,21 +1881,21 @@
 lemma fine_Un: "d fine p1 \<Longrightarrow> d fine p2 \<Longrightarrow> d fine (p1 \<union> p2)"
   unfolding fine_def by blast
 
-lemma%important fine_Union: "(\<And>p. p \<in> ps \<Longrightarrow> d fine p) \<Longrightarrow> d fine (\<Union>ps)"
+lemma fine_Union: "(\<And>p. p \<in> ps \<Longrightarrow> d fine p) \<Longrightarrow> d fine (\<Union>ps)"
   unfolding fine_def by auto
 
-lemma%unimportant fine_subset: "p \<subseteq> q \<Longrightarrow> d fine q \<Longrightarrow> d fine p"
+lemma fine_subset: "p \<subseteq> q \<Longrightarrow> d fine q \<Longrightarrow> d fine p"
   unfolding fine_def by blast
 
 subsection%important \<open>Some basic combining lemmas\<close>
 
-lemma%important tagged_division_Union_exists:
+lemma tagged_division_Union_exists:
   assumes "finite I"
     and "\<forall>i\<in>I. \<exists>p. p tagged_division_of i \<and> d fine p"
     and "\<forall>i1\<in>I. \<forall>i2\<in>I. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}"
     and "\<Union>I = i"
    obtains p where "p tagged_division_of i" and "d fine p"
-proof%unimportant -
+proof -
   obtain pfn where pfn:
     "\<And>x. x \<in> I \<Longrightarrow> pfn x tagged_division_of x"
     "\<And>x. x \<in> I \<Longrightarrow> d fine pfn x"
@@ -1913,17 +1913,17 @@
   fixes i :: "'n::euclidean_space set"
   shows "s division_of i \<Longrightarrow> closed i"
   unfolding division_of_def by fastforce
-
+(* FIXME structure here, do not have one lemma in each subsection *)
 subsection%important \<open>General bisection principle for intervals; might be useful elsewhere\<close>
-
-lemma%important interval_bisection_step:
+(* FIXME  move? *)
+lemma interval_bisection_step:
   fixes type :: "'a::euclidean_space"
   assumes emp: "P {}"
     and Un: "\<And>S T. \<lbrakk>P S; P T; interior(S) \<inter> interior(T) = {}\<rbrakk> \<Longrightarrow> P (S \<union> T)"
     and non: "\<not> P (cbox a (b::'a))"
   obtains c d where "\<not> P (cbox c d)"
     and "\<And>i. i \<in> Basis \<Longrightarrow> a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
-proof%unimportant -
+proof -
   have "cbox a b \<noteq> {}"
     using emp non by metis
   then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
@@ -2029,14 +2029,14 @@
       by (metis (no_types, lifting) assms(3) that)
 qed
 
-lemma%important interval_bisection:
+lemma interval_bisection:
   fixes type :: "'a::euclidean_space"
   assumes "P {}"
     and Un: "\<And>S T. \<lbrakk>P S; P T; interior(S) \<inter> interior(T) = {}\<rbrakk> \<Longrightarrow> P (S \<union> T)"
     and "\<not> P (cbox a (b::'a))"
   obtains x where "x \<in> cbox a b"
     and "\<forall>e>0. \<exists>c d. x \<in> cbox c d \<and> cbox c d \<subseteq> ball x e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
-proof%unimportant -
+proof -
   have "\<forall>x. \<exists>y. \<not> P (cbox (fst x) (snd x)) \<longrightarrow> (\<not> P (cbox (fst y) (snd y)) \<and>
     (\<forall>i\<in>Basis. fst x\<bullet>i \<le> fst y\<bullet>i \<and> fst y\<bullet>i \<le> snd y\<bullet>i \<and> snd y\<bullet>i \<le> snd x\<bullet>i \<and>
        2 * (snd y\<bullet>i - fst y\<bullet>i) \<le> snd x\<bullet>i - fst x\<bullet>i))" (is "\<forall>x. ?P x")
@@ -2182,11 +2182,11 @@
 
 subsection%important \<open>Cousin's lemma\<close>
 
-lemma%important fine_division_exists:
+lemma fine_division_exists: (*FIXME rename(?) *)
   fixes a b :: "'a::euclidean_space"
   assumes "gauge g"
   obtains p where "p tagged_division_of (cbox a b)" "g fine p"
-proof%unimportant (cases "\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p")
+proof (cases "\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p")
   case True
   then show ?thesis
     using that by auto
@@ -2228,14 +2228,14 @@
 
 subsection%important \<open>A technical lemma about "refinement" of division\<close>
 
-lemma%important tagged_division_finer:
+lemma tagged_division_finer:
   fixes p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
   assumes ptag: "p tagged_division_of (cbox a b)"
     and "gauge d"
   obtains q where "q tagged_division_of (cbox a b)"
     and "d fine q"
     and "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
-proof%unimportant -
+proof -
   have p: "finite p" "p tagged_partial_division_of (cbox a b)"
     using ptag unfolding tagged_division_of_def by auto
   have "(\<exists>q. q tagged_division_of (\<Union>{k. \<exists>x. (x,k) \<in> p}) \<and> d fine q \<and> (\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q))" 
@@ -2307,7 +2307,7 @@
   lemma is an obvious multidimensional generalization of Lemma 3, p65 of Swartz's
   "Introduction to Gauge Integrals". \<close>
 
-proposition%important covering_lemma:
+proposition covering_lemma:
   assumes "S \<subseteq> cbox a b" "box a b \<noteq> {}" "gauge g"
   obtains \<D> where
     "countable \<D>"  "\<Union>\<D> \<subseteq> cbox a b"
@@ -2316,7 +2316,7 @@
     "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> g x"
     "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
     "S \<subseteq> \<Union>\<D>"
-proof%unimportant -
+proof -
   have aibi: "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" and normab: "0 < norm(b - a)"
     using \<open>box a b \<noteq> {}\<close> box_eq_empty box_sing by fastforce+
   let ?K0 = "\<lambda>(n, f::'a\<Rightarrow>nat).
@@ -2578,11 +2578,11 @@
 definition%important division_filter :: "'a::euclidean_space set \<Rightarrow> ('a \<times> 'a set) set filter"
   where "division_filter s = (INF g\<in>{g. gauge g}. principal {p. p tagged_division_of s \<and> g fine p})"
 
-lemma%important eventually_division_filter:
+proposition eventually_division_filter:
   "(\<forall>\<^sub>F p in division_filter s. P p) \<longleftrightarrow>
     (\<exists>g. gauge g \<and> (\<forall>p. p tagged_division_of s \<and> g fine p \<longrightarrow> P p))"
-  unfolding%unimportant division_filter_def
-proof%unimportant (subst eventually_INF_base; clarsimp)
+  unfolding division_filter_def
+proof (subst eventually_INF_base; clarsimp)
   fix g1 g2 :: "'a \<Rightarrow> 'a set" show "gauge g1 \<Longrightarrow> gauge g2 \<Longrightarrow> \<exists>x. gauge x \<and>
     {p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g1 fine p} \<and>
     {p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g2 fine p}"
@@ -2593,7 +2593,7 @@
   unfolding trivial_limit_def eventually_division_filter
   by (auto elim: fine_division_exists)
 
-lemma%important eventually_division_filter_tagged_division:
+lemma eventually_division_filter_tagged_division:
   "eventually (\<lambda>p. p tagged_division_of s) (division_filter s)"
   unfolding eventually_division_filter by (intro exI[of _ "\<lambda>x. ball x 1"]) auto