src/HOL/Analysis/Borel_Space.thy
changeset 68833 fde093888c16
parent 68635 8094b853a92f
child 69022 e2858770997a
--- a/src/HOL/Analysis/Borel_Space.thy	Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Borel_Space.thy	Tue Aug 28 13:28:39 2018 +0100
@@ -3,55 +3,55 @@
     Author:     Armin Heller, TU München
 *)
 
-section \<open>Borel spaces\<close>
+section%important \<open>Borel spaces\<close>
 
 theory Borel_Space
 imports
   Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits
 begin
 
-lemma sets_Collect_eventually_sequentially[measurable]:
+lemma%unimportant sets_Collect_eventually_sequentially[measurable]:
   "(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M"
   unfolding eventually_sequentially by simp
 
-lemma topological_basis_trivial: "topological_basis {A. open A}"
+lemma%unimportant topological_basis_trivial: "topological_basis {A. open A}"
   by (auto simp: topological_basis_def)
 
-lemma open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}"
-proof -
+lemma%important open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}"
+proof%unimportant -
   have "{A \<times> B :: ('a \<times> 'b) set | A B. open A \<and> open B} = ((\<lambda>(a, b). a \<times> b) ` ({A. open A} \<times> {A. open A}))"
     by auto
   then show ?thesis
     by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis)
 qed
 
-definition "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
+definition%important "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
 
-lemma mono_onI:
+lemma%unimportant mono_onI:
   "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A"
   unfolding mono_on_def by simp
 
-lemma mono_onD:
+lemma%unimportant mono_onD:
   "\<lbrakk>mono_on f A; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s"
   unfolding mono_on_def by simp
 
-lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
+lemma%unimportant mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
   unfolding mono_def mono_on_def by auto
 
-lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
+lemma%unimportant mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
   unfolding mono_on_def by auto
 
-definition "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
+definition%important "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
 
-lemma strict_mono_onI:
+lemma%unimportant strict_mono_onI:
   "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A"
   unfolding strict_mono_on_def by simp
 
-lemma strict_mono_onD:
+lemma%unimportant strict_mono_onD:
   "\<lbrakk>strict_mono_on f A; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s"
   unfolding strict_mono_on_def by simp
 
-lemma mono_on_greaterD:
+lemma%unimportant mono_on_greaterD:
   assumes "mono_on g A" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)"
   shows "x > y"
 proof (rule ccontr)
@@ -61,7 +61,7 @@
   with assms(4) show False by simp
 qed
 
-lemma strict_mono_inv:
+lemma%unimportant strict_mono_inv:
   fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)"
   assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x"
   shows "strict_mono g"
@@ -72,7 +72,7 @@
   with inv show "g x < g y" by simp
 qed
 
-lemma strict_mono_on_imp_inj_on:
+lemma%unimportant strict_mono_on_imp_inj_on:
   assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder)) A"
   shows "inj_on f A"
 proof (rule inj_onI)
@@ -82,7 +82,7 @@
        (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x])
 qed
 
-lemma strict_mono_on_leD:
+lemma%unimportant strict_mono_on_leD:
   assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A" "x \<in> A" "y \<in> A" "x \<le> y"
   shows "f x \<le> f y"
 proof (insert le_less_linear[of y x], elim disjE)
@@ -91,17 +91,17 @@
   thus ?thesis by (rule less_imp_le)
 qed (insert assms, simp)
 
-lemma strict_mono_on_eqD:
+lemma%unimportant strict_mono_on_eqD:
   fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)"
   assumes "strict_mono_on f A" "f x = f y" "x \<in> A" "y \<in> A"
   shows "y = x"
   using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD)
 
-lemma mono_on_imp_deriv_nonneg:
+lemma%important mono_on_imp_deriv_nonneg:
   assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)"
   assumes "x \<in> interior A"
   shows "D \<ge> 0"
-proof (rule tendsto_lowerbound)
+proof%unimportant (rule tendsto_lowerbound)
   let ?A' = "(\<lambda>y. y - x) ` interior A"
   from deriv show "((\<lambda>h. (f (x + h) - f x) / h) \<longlongrightarrow> D) (at 0)"
       by (simp add: field_has_derivative_at has_field_derivative_def)
@@ -124,16 +124,16 @@
   qed
 qed simp
 
-lemma strict_mono_on_imp_mono_on:
+lemma%unimportant strict_mono_on_imp_mono_on:
   "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A \<Longrightarrow> mono_on f A"
   by (rule mono_onI, rule strict_mono_on_leD)
 
-lemma mono_on_ctble_discont:
+lemma%important mono_on_ctble_discont:
   fixes f :: "real \<Rightarrow> real"
   fixes A :: "real set"
   assumes "mono_on f A"
   shows "countable {a\<in>A. \<not> continuous (at a within A) f}"
-proof -
+proof%unimportant -
   have mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
     using \<open>mono_on f A\<close> by (simp add: mono_on_def)
   have "\<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}. \<exists>q :: nat \<times> rat.
@@ -196,12 +196,12 @@
     by (rule countableI')
 qed
 
-lemma mono_on_ctble_discont_open:
+lemma%important mono_on_ctble_discont_open:
   fixes f :: "real \<Rightarrow> real"
   fixes A :: "real set"
   assumes "open A" "mono_on f A"
   shows "countable {a\<in>A. \<not>isCont f a}"
-proof -
+proof%unimportant -
   have "{a\<in>A. \<not>isCont f a} = {a\<in>A. \<not>(continuous (at a within A) f)}"
     by (auto simp add: continuous_within_open [OF _ \<open>open A\<close>])
   thus ?thesis
@@ -209,13 +209,13 @@
     by (rule mono_on_ctble_discont, rule assms)
 qed
 
-lemma mono_ctble_discont:
+lemma%important mono_ctble_discont:
   fixes f :: "real \<Rightarrow> real"
   assumes "mono f"
   shows "countable {a. \<not> isCont f a}"
-using assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto
+using%unimportant assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto
 
-lemma has_real_derivative_imp_continuous_on:
+lemma%important has_real_derivative_imp_continuous_on:
   assumes "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)"
   shows "continuous_on A f"
   apply (intro differentiable_imp_continuous_on, unfold differentiable_on_def)
@@ -224,11 +224,11 @@
   apply simp_all
   done
 
-lemma closure_contains_Sup:
+lemma%important closure_contains_Sup:
   fixes S :: "real set"
   assumes "S \<noteq> {}" "bdd_above S"
   shows "Sup S \<in> closure S"
-proof-
+proof%unimportant -
   have "Inf (uminus ` S) \<in> closure (uminus ` S)"
       using assms by (intro closure_contains_Inf) auto
   also have "Inf (uminus ` S) = -Sup S" by (simp add: Inf_real_def)
@@ -237,22 +237,22 @@
   finally show ?thesis by auto
 qed
 
-lemma closed_contains_Sup:
+lemma%unimportant closed_contains_Sup:
   fixes S :: "real set"
   shows "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> closed S \<Longrightarrow> Sup S \<in> S"
   by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup)
 
-lemma closed_subset_contains_Sup:
+lemma%unimportant closed_subset_contains_Sup:
   fixes A C :: "real set"
   shows "closed C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> Sup A \<in> C"
   by (metis closure_contains_Sup closure_minimal subset_eq)
 
-lemma deriv_nonneg_imp_mono:
+lemma%important deriv_nonneg_imp_mono:
   assumes deriv: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
   assumes nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
   assumes ab: "a \<le> b"
   shows "g a \<le> g b"
-proof (cases "a < b")
+proof%unimportant (cases "a < b")
   assume "a < b"
   from deriv have "\<And>x. \<lbrakk>x \<ge> a; x \<le> b\<rbrakk> \<Longrightarrow> (g has_real_derivative g' x) (at x)" by simp
   with MVT2[OF \<open>a < b\<close>] and deriv
@@ -261,11 +261,11 @@
   with g_ab show ?thesis by simp
 qed (insert ab, simp)
 
-lemma continuous_interval_vimage_Int:
+lemma%important continuous_interval_vimage_Int:
   assumes "continuous_on {a::real..b} g" and mono: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
   assumes "a \<le> b" "(c::real) \<le> d" "{c..d} \<subseteq> {g a..g b}"
   obtains c' d' where "{a..b} \<inter> g -` {c..d} = {c'..d'}" "c' \<le> d'" "g c' = c" "g d' = d"
-proof-
+proof%unimportant-
   let ?A = "{a..b} \<inter> g -` {c..d}"
   from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
   obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto
@@ -297,48 +297,48 @@
   ultimately show ?thesis using that by blast
 qed
 
-subsection \<open>Generic Borel spaces\<close>
+subsection%important \<open>Generic Borel spaces\<close>
 
-definition (in topological_space) borel :: "'a measure" where
+definition%important (in topological_space) borel :: "'a measure" where
   "borel = sigma UNIV {S. open S}"
 
 abbreviation "borel_measurable M \<equiv> measurable M borel"
 
-lemma in_borel_measurable:
+lemma%important in_borel_measurable:
    "f \<in> borel_measurable M \<longleftrightarrow>
     (\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)"
-  by (auto simp add: measurable_def borel_def)
+  by%unimportant (auto simp add: measurable_def borel_def)
 
-lemma in_borel_measurable_borel:
+lemma%important in_borel_measurable_borel:
    "f \<in> borel_measurable M \<longleftrightarrow>
     (\<forall>S \<in> sets borel.
       f -` S \<inter> space M \<in> sets M)"
-  by (auto simp add: measurable_def borel_def)
+  by%unimportant (auto simp add: measurable_def borel_def)
 
-lemma space_borel[simp]: "space borel = UNIV"
+lemma%unimportant space_borel[simp]: "space borel = UNIV"
   unfolding borel_def by auto
 
-lemma space_in_borel[measurable]: "UNIV \<in> sets borel"
+lemma%unimportant space_in_borel[measurable]: "UNIV \<in> sets borel"
   unfolding borel_def by auto
 
-lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
+lemma%unimportant sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
   unfolding borel_def by (rule sets_measure_of) simp
 
-lemma measurable_sets_borel:
+lemma%unimportant measurable_sets_borel:
     "\<lbrakk>f \<in> measurable borel M; A \<in> sets M\<rbrakk> \<Longrightarrow> f -` A \<in> sets borel"
   by (drule (1) measurable_sets) simp
 
-lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
+lemma%unimportant pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
   unfolding borel_def pred_def by auto
 
-lemma borel_open[measurable (raw generic)]:
+lemma%unimportant borel_open[measurable (raw generic)]:
   assumes "open A" shows "A \<in> sets borel"
 proof -
   have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
   thus ?thesis unfolding borel_def by auto
 qed
 
-lemma borel_closed[measurable (raw generic)]:
+lemma%unimportant borel_closed[measurable (raw generic)]:
   assumes "closed A" shows "A \<in> sets borel"
 proof -
   have "space borel - (- A) \<in> sets borel"
@@ -346,11 +346,11 @@
   thus ?thesis by simp
 qed
 
-lemma borel_singleton[measurable]:
+lemma%unimportant borel_singleton[measurable]:
   "A \<in> sets borel \<Longrightarrow> insert x A \<in> sets (borel :: 'a::t1_space measure)"
   unfolding insert_def by (rule sets.Un) auto
 
-lemma sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
+lemma%unimportant sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
 proof -
   have "(\<Union>a\<in>A. {a}) \<in> sets borel" for A :: "'a set"
     by (intro sets.countable_UN') auto
@@ -358,16 +358,16 @@
     by auto
 qed
 
-lemma borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
+lemma%unimportant borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
   unfolding Compl_eq_Diff_UNIV by simp
 
-lemma borel_measurable_vimage:
+lemma%unimportant borel_measurable_vimage:
   fixes f :: "'a \<Rightarrow> 'x::t2_space"
   assumes borel[measurable]: "f \<in> borel_measurable M"
   shows "f -` {x} \<inter> space M \<in> sets M"
   by simp
 
-lemma borel_measurableI:
+lemma%unimportant borel_measurableI:
   fixes f :: "'a \<Rightarrow> 'x::topological_space"
   assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
   shows "f \<in> borel_measurable M"
@@ -377,30 +377,30 @@
     using assms[of S] by simp
 qed
 
-lemma borel_measurable_const:
+lemma%unimportant borel_measurable_const:
   "(\<lambda>x. c) \<in> borel_measurable M"
   by auto
 
-lemma borel_measurable_indicator:
+lemma%unimportant borel_measurable_indicator:
   assumes A: "A \<in> sets M"
   shows "indicator A \<in> borel_measurable M"
   unfolding indicator_def [abs_def] using A
   by (auto intro!: measurable_If_set)
 
-lemma borel_measurable_count_space[measurable (raw)]:
+lemma%unimportant borel_measurable_count_space[measurable (raw)]:
   "f \<in> borel_measurable (count_space S)"
   unfolding measurable_def by auto
 
-lemma borel_measurable_indicator'[measurable (raw)]:
+lemma%unimportant borel_measurable_indicator'[measurable (raw)]:
   assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M"
   shows "(\<lambda>x. indicator (A x) (f x)) \<in> borel_measurable M"
   unfolding indicator_def[abs_def]
   by (auto intro!: measurable_If)
 
-lemma borel_measurable_indicator_iff:
+lemma%important borel_measurable_indicator_iff:
   "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
     (is "?I \<in> borel_measurable M \<longleftrightarrow> _")
-proof
+proof%unimportant
   assume "?I \<in> borel_measurable M"
   then have "?I -` {1} \<inter> space M \<in> sets M"
     unfolding measurable_def by auto
@@ -415,12 +415,12 @@
   ultimately show "?I \<in> borel_measurable M" by auto
 qed
 
-lemma borel_measurable_subalgebra:
+lemma%unimportant borel_measurable_subalgebra:
   assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N"
   shows "f \<in> borel_measurable M"
   using assms unfolding measurable_def by auto
 
-lemma borel_measurable_restrict_space_iff_ereal:
+lemma%unimportant borel_measurable_restrict_space_iff_ereal:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
   shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
@@ -428,7 +428,7 @@
   by (subst measurable_restrict_space_iff)
      (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong)
 
-lemma borel_measurable_restrict_space_iff_ennreal:
+lemma%unimportant borel_measurable_restrict_space_iff_ennreal:
   fixes f :: "'a \<Rightarrow> ennreal"
   assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
   shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
@@ -436,7 +436,7 @@
   by (subst measurable_restrict_space_iff)
      (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong)
 
-lemma borel_measurable_restrict_space_iff:
+lemma%unimportant borel_measurable_restrict_space_iff:
   fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
   assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
   shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
@@ -445,27 +445,27 @@
      (auto simp: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps
        cong del: if_weak_cong)
 
-lemma cbox_borel[measurable]: "cbox a b \<in> sets borel"
+lemma%unimportant cbox_borel[measurable]: "cbox a b \<in> sets borel"
   by (auto intro: borel_closed)
 
-lemma box_borel[measurable]: "box a b \<in> sets borel"
+lemma%unimportant box_borel[measurable]: "box a b \<in> sets borel"
   by (auto intro: borel_open)
 
-lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
+lemma%unimportant borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
   by (auto intro: borel_closed dest!: compact_imp_closed)
 
-lemma borel_sigma_sets_subset:
+lemma%unimportant borel_sigma_sets_subset:
   "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
   using sets.sigma_sets_subset[of A borel] by simp
 
-lemma borel_eq_sigmaI1:
+lemma%important borel_eq_sigmaI1:
   fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
   assumes borel_eq: "borel = sigma UNIV X"
   assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (F ` A))"
   assumes F: "\<And>i. i \<in> A \<Longrightarrow> F i \<in> sets borel"
   shows "borel = sigma UNIV (F ` A)"
   unfolding borel_def
-proof (intro sigma_eqI antisym)
+proof%unimportant (intro sigma_eqI antisym)
   have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
     unfolding borel_def by simp
   also have "\<dots> = sigma_sets UNIV X"
@@ -477,7 +477,7 @@
     unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
 qed auto
 
-lemma borel_eq_sigmaI2:
+lemma%unimportant borel_eq_sigmaI2:
   fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"
     and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
   assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`B)"
@@ -487,7 +487,7 @@
   using assms
   by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` B" and F="(\<lambda>(i, j). F i j)"]) auto
 
-lemma borel_eq_sigmaI3:
+lemma%unimportant borel_eq_sigmaI3:
   fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
   assumes borel_eq: "borel = sigma UNIV X"
   assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"
@@ -495,7 +495,7 @@
   shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
   using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto
 
-lemma borel_eq_sigmaI4:
+lemma%unimportant borel_eq_sigmaI4:
   fixes F :: "'i \<Rightarrow> 'a::topological_space set"
     and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
   assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`A)"
@@ -504,7 +504,7 @@
   shows "borel = sigma UNIV (range F)"
   using assms by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` A" and F=F]) auto
 
-lemma borel_eq_sigmaI5:
+lemma%unimportant borel_eq_sigmaI5:
   fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"
   assumes borel_eq: "borel = sigma UNIV (range G)"
   assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
@@ -512,12 +512,12 @@
   shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
   using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
 
-lemma second_countable_borel_measurable:
+lemma%important second_countable_borel_measurable:
   fixes X :: "'a::second_countable_topology set set"
   assumes eq: "open = generate_topology X"
   shows "borel = sigma UNIV X"
   unfolding borel_def
-proof (intro sigma_eqI sigma_sets_eqI)
+proof%unimportant (intro sigma_eqI sigma_sets_eqI)
   interpret X: sigma_algebra UNIV "sigma_sets UNIV X"
     by (rule sigma_algebra_sigma_sets) simp
 
@@ -556,7 +556,7 @@
   qed auto
 qed (auto simp: eq intro: generate_topology.Basis)
 
-lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)"
+lemma%unimportant borel_eq_closed: "borel = sigma UNIV (Collect closed)"
   unfolding borel_def
 proof (intro sigma_eqI sigma_sets_eqI, safe)
   fix x :: "'a set" assume "open x"
@@ -572,13 +572,13 @@
   finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
 qed simp_all
 
-lemma borel_eq_countable_basis:
+lemma%important borel_eq_countable_basis:
   fixes B::"'a::topological_space set set"
   assumes "countable B"
   assumes "topological_basis B"
   shows "borel = sigma UNIV B"
   unfolding borel_def
-proof (intro sigma_eqI sigma_sets_eqI, safe)
+proof%unimportant (intro sigma_eqI sigma_sets_eqI, safe)
   interpret countable_basis using assms by unfold_locales
   fix X::"'a set" assume "open X"
   from open_countable_basisE[OF this] guess B' . note B' = this
@@ -590,7 +590,7 @@
   thus "b \<in> sigma_sets UNIV (Collect open)" by auto
 qed simp_all
 
-lemma borel_measurable_continuous_on_restrict:
+lemma%unimportant borel_measurable_continuous_on_restrict:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
   assumes f: "continuous_on A f"
   shows "f \<in> borel_measurable (restrict_space borel A)"
@@ -602,16 +602,16 @@
     by (force simp add: sets_restrict_space space_restrict_space)
 qed
 
-lemma borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"
   by (drule borel_measurable_continuous_on_restrict) simp
 
-lemma borel_measurable_continuous_on_if:
+lemma%unimportant borel_measurable_continuous_on_if:
   "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> continuous_on (- A) g \<Longrightarrow>
     (\<lambda>x. if x \<in> A then f x else g x) \<in> borel_measurable borel"
   by (auto simp add: measurable_If_restrict_space_iff Collect_neg_eq
            intro!: borel_measurable_continuous_on_restrict)
 
-lemma borel_measurable_continuous_countable_exceptions:
+lemma%unimportant borel_measurable_continuous_countable_exceptions:
   fixes f :: "'a::t1_space \<Rightarrow> 'b::topological_space"
   assumes X: "countable X"
   assumes "continuous_on (- X) f"
@@ -623,23 +623,23 @@
     by (intro borel_measurable_continuous_on_if assms continuous_intros)
 qed auto
 
-lemma borel_measurable_continuous_on:
+lemma%unimportant borel_measurable_continuous_on:
   assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f (g x)) \<in> borel_measurable M"
   using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def)
 
-lemma borel_measurable_continuous_on_indicator:
+lemma%unimportant borel_measurable_continuous_on_indicator:
   fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
   shows "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel"
   by (subst borel_measurable_restrict_space_iff[symmetric])
      (auto intro: borel_measurable_continuous_on_restrict)
 
-lemma borel_measurable_Pair[measurable (raw)]:
+lemma%important borel_measurable_Pair[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
   assumes f[measurable]: "f \<in> borel_measurable M"
   assumes g[measurable]: "g \<in> borel_measurable M"
   shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M"
-proof (subst borel_eq_countable_basis)
+proof%unimportant (subst borel_eq_countable_basis)
   let ?B = "SOME B::'b set set. countable B \<and> topological_basis B"
   let ?C = "SOME B::'c set set. countable B \<and> topological_basis B"
   let ?P = "(\<lambda>(b, c). b \<times> c) ` (?B \<times> ?C)"
@@ -660,21 +660,21 @@
   qed auto
 qed
 
-lemma borel_measurable_continuous_Pair:
+lemma%important borel_measurable_continuous_Pair:
   fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
   assumes [measurable]: "f \<in> borel_measurable M"
   assumes [measurable]: "g \<in> borel_measurable M"
   assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))"
   shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
-proof -
+proof%unimportant -
   have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto
   show ?thesis
     unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto
 qed
 
-subsection \<open>Borel spaces on order topologies\<close>
+subsection%important \<open>Borel spaces on order topologies\<close>
 
-lemma [measurable]:
+lemma%unimportant [measurable]:
   fixes a b :: "'a::linorder_topology"
   shows lessThan_borel: "{..< a} \<in> sets borel"
     and greaterThan_borel: "{a <..} \<in> sets borel"
@@ -688,7 +688,7 @@
   by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan
                    closed_atMost closed_atLeast closed_atLeastAtMost)+
 
-lemma borel_Iio:
+lemma%unimportant borel_Iio:
   "borel = sigma UNIV (range lessThan :: 'a::{linorder_topology, second_countable_topology} set set)"
   unfolding second_countable_borel_measurable[OF open_generated_order]
 proof (intro sigma_eqI sigma_sets_eqI)
@@ -726,7 +726,7 @@
   qed auto
 qed auto
 
-lemma borel_Ioi:
+lemma%unimportant borel_Ioi:
   "borel = sigma UNIV (range greaterThan :: 'a::{linorder_topology, second_countable_topology} set set)"
   unfolding second_countable_borel_measurable[OF open_generated_order]
 proof (intro sigma_eqI sigma_sets_eqI)
@@ -764,29 +764,29 @@
   qed auto
 qed auto
 
-lemma borel_measurableI_less:
+lemma%unimportant borel_measurableI_less:
   fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
   shows "(\<And>y. {x\<in>space M. f x < y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
   unfolding borel_Iio
   by (rule measurable_measure_of) (auto simp: Int_def conj_commute)
 
-lemma borel_measurableI_greater:
+lemma%important borel_measurableI_greater:
   fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
   shows "(\<And>y. {x\<in>space M. y < f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
   unfolding borel_Ioi
-  by (rule measurable_measure_of) (auto simp: Int_def conj_commute)
+  by%unimportant (rule measurable_measure_of) (auto simp: Int_def conj_commute)
 
-lemma borel_measurableI_le:
+lemma%unimportant borel_measurableI_le:
   fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
   shows "(\<And>y. {x\<in>space M. f x \<le> y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
   by (rule borel_measurableI_greater) (auto simp: not_le[symmetric])
 
-lemma borel_measurableI_ge:
+lemma%unimportant borel_measurableI_ge:
   fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
   shows "(\<And>y. {x\<in>space M. y \<le> f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
   by (rule borel_measurableI_less) (auto simp: not_le[symmetric])
 
-lemma borel_measurable_less[measurable]:
+lemma%unimportant borel_measurable_less[measurable]:
   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
@@ -800,7 +800,7 @@
   finally show ?thesis .
 qed
 
-lemma
+lemma%important
   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
   assumes f[measurable]: "f \<in> borel_measurable M"
   assumes g[measurable]: "g \<in> borel_measurable M"
@@ -808,23 +808,23 @@
     and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
     and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
   unfolding eq_iff not_less[symmetric]
-  by measurable
+  by%unimportant measurable
 
-lemma borel_measurable_SUP[measurable (raw)]:
+lemma%important borel_measurable_SUP[measurable (raw)]:
   fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}"
   assumes [simp]: "countable I"
   assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
   shows "(\<lambda>x. SUP i:I. F i x) \<in> borel_measurable M"
-  by (rule borel_measurableI_greater) (simp add: less_SUP_iff)
+  by%unimportant (rule borel_measurableI_greater) (simp add: less_SUP_iff)
 
-lemma borel_measurable_INF[measurable (raw)]:
+lemma%unimportant borel_measurable_INF[measurable (raw)]:
   fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}"
   assumes [simp]: "countable I"
   assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
   shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M"
   by (rule borel_measurableI_less) (simp add: INF_less_iff)
 
-lemma borel_measurable_cSUP[measurable (raw)]:
+lemma%unimportant borel_measurable_cSUP[measurable (raw)]:
   fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
   assumes [simp]: "countable I"
   assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
@@ -846,13 +846,13 @@
   qed
 qed
 
-lemma borel_measurable_cINF[measurable (raw)]:
+lemma%important borel_measurable_cINF[measurable (raw)]:
   fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
   assumes [simp]: "countable I"
   assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
   assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_below ((\<lambda>i. F i x) ` I)"
   shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M"
-proof cases
+proof%unimportant cases
   assume "I = {}" then show ?thesis
     unfolding \<open>I = {}\<close> image_empty by simp
 next
@@ -868,7 +868,7 @@
   qed
 qed
 
-lemma borel_measurable_lfp[consumes 1, case_names continuity step]:
+lemma%unimportant borel_measurable_lfp[consumes 1, case_names continuity step]:
   fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})"
   assumes "sup_continuous F"
   assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"
@@ -885,7 +885,7 @@
   finally show ?thesis .
 qed
 
-lemma borel_measurable_gfp[consumes 1, case_names continuity step]:
+lemma%unimportant borel_measurable_gfp[consumes 1, case_names continuity step]:
   fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})"
   assumes "inf_continuous F"
   assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"
@@ -902,56 +902,56 @@
   finally show ?thesis .
 qed
 
-lemma borel_measurable_max[measurable (raw)]:
+lemma%unimportant borel_measurable_max[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. max (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
   by (rule borel_measurableI_less) simp
 
-lemma borel_measurable_min[measurable (raw)]:
+lemma%unimportant borel_measurable_min[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. min (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
   by (rule borel_measurableI_greater) simp
 
-lemma borel_measurable_Min[measurable (raw)]:
+lemma%unimportant borel_measurable_Min[measurable (raw)]:
   "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Min ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 proof (induct I rule: finite_induct)
   case (insert i I) then show ?case
     by (cases "I = {}") auto
 qed auto
 
-lemma borel_measurable_Max[measurable (raw)]:
+lemma%unimportant borel_measurable_Max[measurable (raw)]:
   "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Max ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 proof (induct I rule: finite_induct)
   case (insert i I) then show ?case
     by (cases "I = {}") auto
 qed auto
 
-lemma borel_measurable_sup[measurable (raw)]:
+lemma%unimportant borel_measurable_sup[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. sup (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
   unfolding sup_max by measurable
 
-lemma borel_measurable_inf[measurable (raw)]:
+lemma%unimportant borel_measurable_inf[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. inf (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
   unfolding inf_min by measurable
 
-lemma [measurable (raw)]:
+lemma%unimportant [measurable (raw)]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
   assumes "\<And>i. f i \<in> borel_measurable M"
   shows borel_measurable_liminf: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
     and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
   unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto
 
-lemma measurable_convergent[measurable (raw)]:
+lemma%unimportant measurable_convergent[measurable (raw)]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
   assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "Measurable.pred M (\<lambda>x. convergent (\<lambda>i. f i x))"
   unfolding convergent_ereal by measurable
 
-lemma sets_Collect_convergent[measurable]:
+lemma%unimportant sets_Collect_convergent[measurable]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
   assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M"
   by measurable
 
-lemma borel_measurable_lim[measurable (raw)]:
+lemma%unimportant borel_measurable_lim[measurable (raw)]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
   assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
@@ -962,7 +962,7 @@
     by simp
 qed
 
-lemma borel_measurable_LIMSEQ_order:
+lemma%unimportant borel_measurable_LIMSEQ_order:
   fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
   assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
   and u: "\<And>i. u i \<in> borel_measurable M"
@@ -973,16 +973,16 @@
   with u show ?thesis by (simp cong: measurable_cong)
 qed
 
-subsection \<open>Borel spaces on topological monoids\<close>
+subsection%important \<open>Borel spaces on topological monoids\<close>
 
-lemma borel_measurable_add[measurable (raw)]:
+lemma%unimportant borel_measurable_add[measurable (raw)]:
   fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, topological_monoid_add}"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
   using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
 
-lemma borel_measurable_sum[measurable (raw)]:
+lemma%unimportant borel_measurable_sum[measurable (raw)]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, topological_comm_monoid_add}"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
@@ -991,30 +991,30 @@
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma borel_measurable_suminf_order[measurable (raw)]:
+lemma%important borel_measurable_suminf_order[measurable (raw)]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}"
   assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"
   unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
 
-subsection \<open>Borel spaces on Euclidean spaces\<close>
+subsection%important \<open>Borel spaces on Euclidean spaces\<close>
 
-lemma borel_measurable_inner[measurable (raw)]:
+lemma%important borel_measurable_inner[measurable (raw)]:
   fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_inner}"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x \<bullet> g x) \<in> borel_measurable M"
   using assms
-  by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+  by%unimportant (rule borel_measurable_continuous_Pair) (intro continuous_intros)
 
 notation
   eucl_less (infix "<e" 50)
 
-lemma box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
+lemma%important box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
   and box_co: "{x. a \<le> x \<and> x <e b} = {a..} \<inter> {x. x <e b}"
   by auto
 
-lemma eucl_ivals[measurable]:
+lemma%important eucl_ivals[measurable]:
   fixes a b :: "'a::ordered_euclidean_space"
   shows "{x. x <e a} \<in> sets borel"
     and "{x. a <e x} \<in> sets borel"
@@ -1026,7 +1026,7 @@
   unfolding box_oc box_co
   by (auto intro: borel_open borel_closed)
 
-lemma
+lemma%unimportant (*FIX ME this has no name *)
   fixes i :: "'a::{second_countable_topology, real_inner}"
   shows hafspace_less_borel: "{x. a < x \<bullet> i} \<in> sets borel"
     and hafspace_greater_borel: "{x. x \<bullet> i < a} \<in> sets borel"
@@ -1034,7 +1034,7 @@
     and hafspace_greater_eq_borel: "{x. x \<bullet> i \<le> a} \<in> sets borel"
   by simp_all
 
-lemma borel_eq_box:
+lemma%unimportant borel_eq_box:
   "borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a :: euclidean_space set))"
     (is "_ = ?SIGMA")
 proof (rule borel_eq_sigmaI1[OF borel_def])
@@ -1047,7 +1047,7 @@
     done
 qed (auto simp: box_def)
 
-lemma halfspace_gt_in_halfspace:
+lemma%unimportant halfspace_gt_in_halfspace:
   assumes i: "i \<in> A"
   shows "{x::'a. a < x \<bullet> i} \<in>
     sigma_sets UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> A))"
@@ -1073,7 +1073,7 @@
     by (auto intro!: Diff sigma_sets_Inter i)
 qed
 
-lemma borel_eq_halfspace_less:
+lemma%unimportant borel_eq_halfspace_less:
   "borel = sigma UNIV ((\<lambda>(a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> Basis))"
   (is "_ = ?SIGMA")
 proof (rule borel_eq_sigmaI2[OF borel_eq_box])
@@ -1086,7 +1086,7 @@
   finally show "box a b \<in> sets ?SIGMA" .
 qed auto
 
-lemma borel_eq_halfspace_le:
+lemma%unimportant borel_eq_halfspace_le:
   "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i \<le> a}) ` (UNIV \<times> Basis))"
   (is "_ = ?SIGMA")
 proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
@@ -1110,7 +1110,7 @@
     by (intro sets.countable_UN) (auto intro: i)
 qed auto
 
-lemma borel_eq_halfspace_ge:
+lemma%unimportant borel_eq_halfspace_ge:
   "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
   (is "_ = ?SIGMA")
 proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
@@ -1120,10 +1120,10 @@
     using i by (intro sets.compl_sets) auto
 qed auto
 
-lemma borel_eq_halfspace_greater:
+lemma%important borel_eq_halfspace_greater:
   "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
   (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
+proof%unimportant (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
   fix a :: real and i :: 'a assume "(a, i) \<in> (UNIV \<times> Basis)"
   then have i: "i \<in> Basis" by auto
   have *: "{x::'a. x\<bullet>i \<le> a} = space ?SIGMA - {x::'a. a < x\<bullet>i}" by auto
@@ -1131,7 +1131,7 @@
     by (intro sets.compl_sets) (auto intro: i)
 qed auto
 
-lemma borel_eq_atMost:
+lemma%unimportant borel_eq_atMost:
   "borel = sigma UNIV (range (\<lambda>a. {..a::'a::ordered_euclidean_space}))"
   (is "_ = ?SIGMA")
 proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
@@ -1150,7 +1150,7 @@
     by (intro sets.countable_UN) auto
 qed auto
 
-lemma borel_eq_greaterThan:
+lemma%unimportant borel_eq_greaterThan:
   "borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. a <e x}))"
   (is "_ = ?SIGMA")
 proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
@@ -1177,7 +1177,7 @@
     done
 qed auto
 
-lemma borel_eq_lessThan:
+lemma%unimportant borel_eq_lessThan:
   "borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. x <e a}))"
   (is "_ = ?SIGMA")
 proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
@@ -1203,10 +1203,10 @@
     done
 qed auto
 
-lemma borel_eq_atLeastAtMost:
+lemma%important borel_eq_atLeastAtMost:
   "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} ::'a::ordered_euclidean_space set))"
   (is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
+proof%unimportant (rule borel_eq_sigmaI5[OF borel_eq_atMost])
   fix a::'a
   have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
   proof (safe, simp_all add: eucl_le[where 'a='a])
@@ -1225,12 +1225,12 @@
        (auto intro!: sigma_sets_top)
 qed auto
 
-lemma borel_set_induct[consumes 1, case_names empty interval compl union]:
+lemma%important borel_set_induct[consumes 1, case_names empty interval compl union]:
   assumes "A \<in> sets borel"
   assumes empty: "P {}" and int: "\<And>a b. a \<le> b \<Longrightarrow> P {a..b}" and compl: "\<And>A. A \<in> sets borel \<Longrightarrow> P A \<Longrightarrow> P (-A)" and
           un: "\<And>f. disjoint_family f \<Longrightarrow> (\<And>i. f i \<in> sets borel) \<Longrightarrow>  (\<And>i. P (f i)) \<Longrightarrow> P (\<Union>i::nat. f i)"
   shows "P (A::real set)"
-proof-
+proof%unimportant -
   let ?G = "range (\<lambda>(a,b). {a..b::real})"
   have "Int_stable ?G" "?G \<subseteq> Pow UNIV" "A \<in> sigma_sets UNIV ?G"
       using assms(1) by (auto simp add: borel_eq_atLeastAtMost Int_stable_def)
@@ -1247,7 +1247,7 @@
   qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost)
 qed
 
-lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
+lemma%unimportant borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
 proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
   fix i :: real
   have "{..i} = (\<Union>j::nat. {-j <.. i})"
@@ -1257,10 +1257,10 @@
   finally show "{..i} \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" .
 qed simp
 
-lemma eucl_lessThan: "{x::real. x <e a} = lessThan a"
+lemma%unimportant eucl_lessThan: "{x::real. x <e a} = lessThan a"
   by (simp add: eucl_less_def lessThan_def)
 
-lemma borel_eq_atLeastLessThan:
+lemma%unimportant borel_eq_atLeastLessThan:
   "borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
 proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan])
   have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
@@ -1271,7 +1271,7 @@
     by (auto intro: sigma_sets.intros(2-) simp: eucl_lessThan)
 qed auto
 
-lemma borel_measurable_halfspacesI:
+lemma%unimportant borel_measurable_halfspacesI:
   fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
   assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))"
   and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
@@ -1286,112 +1286,112 @@
     by (auto intro!: measurable_measure_of simp: S_eq F)
 qed
 
-lemma borel_measurable_iff_halfspace_le:
+lemma%unimportant borel_measurable_iff_halfspace_le:
   fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
   shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i \<le> a} \<in> sets M)"
   by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
 
-lemma borel_measurable_iff_halfspace_less:
+lemma%unimportant borel_measurable_iff_halfspace_less:
   fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i < a} \<in> sets M)"
   by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
 
-lemma borel_measurable_iff_halfspace_ge:
+lemma%unimportant borel_measurable_iff_halfspace_ge:
   fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
   shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a \<le> f w \<bullet> i} \<in> sets M)"
   by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
 
-lemma borel_measurable_iff_halfspace_greater:
+lemma%unimportant borel_measurable_iff_halfspace_greater:
   fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a < f w \<bullet> i} \<in> sets M)"
   by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
 
-lemma borel_measurable_iff_le:
+lemma%unimportant borel_measurable_iff_le:
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
   using borel_measurable_iff_halfspace_le[where 'c=real] by simp
 
-lemma borel_measurable_iff_less:
+lemma%unimportant borel_measurable_iff_less:
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
   using borel_measurable_iff_halfspace_less[where 'c=real] by simp
 
-lemma borel_measurable_iff_ge:
+lemma%unimportant borel_measurable_iff_ge:
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
   using borel_measurable_iff_halfspace_ge[where 'c=real]
   by simp
 
-lemma borel_measurable_iff_greater:
+lemma%unimportant borel_measurable_iff_greater:
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
   using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
 
-lemma borel_measurable_euclidean_space:
+lemma%important borel_measurable_euclidean_space:
   fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M)"
-proof safe
+proof%unimportant safe
   assume f: "\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M"
   then show "f \<in> borel_measurable M"
     by (subst borel_measurable_iff_halfspace_le) auto
 qed auto
 
-subsection "Borel measurable operators"
+subsection%important "Borel measurable operators"
 
-lemma borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"
-  by (intro borel_measurable_continuous_on1 continuous_intros)
+lemma%important borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"
+  by%unimportant (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel"
   by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"])
      (auto intro!: continuous_on_sgn continuous_on_id)
 
-lemma borel_measurable_uminus[measurable (raw)]:
+lemma%important borel_measurable_uminus[measurable (raw)]:
   fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. - g x) \<in> borel_measurable M"
-  by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros)
+  by%unimportant (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros)
 
-lemma borel_measurable_diff[measurable (raw)]:
+lemma%important borel_measurable_diff[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-  using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
+  using%unimportant borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
 
-lemma borel_measurable_times[measurable (raw)]:
+lemma%unimportant borel_measurable_times[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
   using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
 
-lemma borel_measurable_prod[measurable (raw)]:
+lemma%important borel_measurable_prod[measurable (raw)]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, real_normed_field}"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
-proof cases
+proof%unimportant cases
   assume "finite S"
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma borel_measurable_dist[measurable (raw)]:
+lemma%important borel_measurable_dist[measurable (raw)]:
   fixes g f :: "'a \<Rightarrow> 'b::{second_countable_topology, metric_space}"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M"
-  using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+  using%unimportant f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
 
-lemma borel_measurable_scaleR[measurable (raw)]:
+lemma%unimportant borel_measurable_scaleR[measurable (raw)]:
   fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x *\<^sub>R g x) \<in> borel_measurable M"
   using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
 
-lemma borel_measurable_uminus_eq [simp]:
+lemma%unimportant borel_measurable_uminus_eq [simp]:
   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
   shows "(\<lambda>x. - f x) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
 proof
   assume ?l from borel_measurable_uminus[OF this] show ?r by simp
 qed auto
 
-lemma affine_borel_measurable_vector:
+lemma%unimportant affine_borel_measurable_vector:
   fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
   assumes "f \<in> borel_measurable M"
   shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
@@ -1412,15 +1412,15 @@
   qed simp
 qed
 
-lemma borel_measurable_const_scaleR[measurable (raw)]:
+lemma%unimportant borel_measurable_const_scaleR[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. b *\<^sub>R f x ::'a::real_normed_vector) \<in> borel_measurable M"
   using affine_borel_measurable_vector[of f M 0 b] by simp
 
-lemma borel_measurable_const_add[measurable (raw)]:
+lemma%unimportant borel_measurable_const_add[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. a + f x ::'a::real_normed_vector) \<in> borel_measurable M"
   using affine_borel_measurable_vector[of f M a 1] by simp
 
-lemma borel_measurable_inverse[measurable (raw)]:
+lemma%unimportant borel_measurable_inverse[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra"
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
@@ -1429,27 +1429,27 @@
   apply (auto intro!: continuous_on_inverse continuous_on_id)
   done
 
-lemma borel_measurable_divide[measurable (raw)]:
+lemma%unimportant borel_measurable_divide[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
     (\<lambda>x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \<in> borel_measurable M"
   by (simp add: divide_inverse)
 
-lemma borel_measurable_abs[measurable (raw)]:
+lemma%unimportant borel_measurable_abs[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
   unfolding abs_real_def by simp
 
-lemma borel_measurable_nth[measurable (raw)]:
+lemma%unimportant borel_measurable_nth[measurable (raw)]:
   "(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"
   by (simp add: cart_eq_inner_axis)
 
-lemma convex_measurable:
+lemma%important convex_measurable:
   fixes A :: "'a :: euclidean_space set"
   shows "X \<in> borel_measurable M \<Longrightarrow> X ` space M \<subseteq> A \<Longrightarrow> open A \<Longrightarrow> convex_on A q \<Longrightarrow>
     (\<lambda>x. q (X x)) \<in> borel_measurable M"
-  by (rule measurable_compose[where f=X and N="restrict_space borel A"])
+  by%unimportant (rule measurable_compose[where f=X and N="restrict_space borel A"])
      (auto intro!: borel_measurable_continuous_on_restrict convex_on_continuous measurable_restrict_space2)
 
-lemma borel_measurable_ln[measurable (raw)]:
+lemma%unimportant borel_measurable_ln[measurable (raw)]:
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. ln (f x :: real)) \<in> borel_measurable M"
   apply (rule measurable_compose[OF f])
@@ -1457,15 +1457,15 @@
   apply (auto intro!: continuous_on_ln continuous_on_id)
   done
 
-lemma borel_measurable_log[measurable (raw)]:
+lemma%unimportant borel_measurable_log[measurable (raw)]:
   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M"
   unfolding log_def by auto
 
-lemma borel_measurable_exp[measurable]:
+lemma%unimportant borel_measurable_exp[measurable]:
   "(exp::'a::{real_normed_field,banach}\<Rightarrow>'a) \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp)
 
-lemma measurable_real_floor[measurable]:
+lemma%unimportant measurable_real_floor[measurable]:
   "(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
 proof -
   have "\<And>a x. \<lfloor>x\<rfloor> = a \<longleftrightarrow> (real_of_int a \<le> x \<and> x < real_of_int (a + 1))"
@@ -1474,44 +1474,44 @@
     by (auto simp: vimage_def measurable_count_space_eq2_countable)
 qed
 
-lemma measurable_real_ceiling[measurable]:
+lemma%unimportant measurable_real_ceiling[measurable]:
   "(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
   unfolding ceiling_def[abs_def] by simp
 
-lemma borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
   by simp
 
-lemma borel_measurable_root [measurable]: "root n \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_root [measurable]: "root n \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_power [measurable (raw)]:
+lemma%unimportant borel_measurable_power [measurable (raw)]:
   fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M"
   by (intro borel_measurable_continuous_on [OF _ f] continuous_intros)
 
-lemma borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
   by (intro borel_measurable_continuous_on1 continuous_intros)
 
-lemma borel_measurable_complex_iff:
+lemma%important borel_measurable_complex_iff:
   "f \<in> borel_measurable M \<longleftrightarrow>
     (\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M"
   apply auto
@@ -1520,21 +1520,21 @@
   apply auto
   done
 
-lemma powr_real_measurable [measurable]:
+lemma%important powr_real_measurable [measurable]:
   assumes "f \<in> measurable M borel" "g \<in> measurable M borel"
   shows   "(\<lambda>x. f x powr g x :: real) \<in> measurable M borel"
-  using assms by (simp_all add: powr_def)
+  using%unimportant assms by (simp_all add: powr_def)
 
-lemma measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel"
+lemma%unimportant measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel"
   by simp
 
-subsection "Borel space on the extended reals"
+subsection%important "Borel space on the extended reals"
 
-lemma borel_measurable_ereal[measurable (raw)]:
+lemma%unimportant borel_measurable_ereal[measurable (raw)]:
   assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
   using continuous_on_ereal f by (rule borel_measurable_continuous_on) (rule continuous_on_id)
 
-lemma borel_measurable_real_of_ereal[measurable (raw)]:
+lemma%unimportant borel_measurable_real_of_ereal[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M"
@@ -1543,7 +1543,7 @@
   apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV)
   done
 
-lemma borel_measurable_ereal_cases:
+lemma%unimportant borel_measurable_ereal_cases:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x)))) \<in> borel_measurable M"
@@ -1554,20 +1554,20 @@
   with f H show ?thesis by simp
 qed
 
-lemma
+lemma%unimportant (*FIX ME needs a name *)
   fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M"
   shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
     and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"
     and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
   by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If)
 
-lemma borel_measurable_uminus_eq_ereal[simp]:
+lemma%unimportant borel_measurable_uminus_eq_ereal[simp]:
   "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
 proof
   assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp
 qed auto
 
-lemma set_Collect_ereal2:
+lemma%important set_Collect_ereal2:
   fixes f g :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -1577,7 +1577,7 @@
     "{x \<in> space borel. H (ereal x) (-\<infinity>)} \<in> sets borel"
     "{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel"
   shows "{x \<in> space M. H (f x) (g x)} \<in> sets M"
-proof -
+proof%unimportant -
   let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = -\<infinity> then H y (-\<infinity>) else H y (ereal (real_of_ereal (g x)))"
   let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = -\<infinity> then ?G (-\<infinity>) x else ?G (ereal (real_of_ereal (f x))) x"
   { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
@@ -1586,7 +1586,7 @@
     by (subst *) (simp del: space_borel split del: if_split)
 qed
 
-lemma borel_measurable_ereal_iff:
+lemma%unimportant borel_measurable_ereal_iff:
   shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
 proof
   assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
@@ -1594,15 +1594,15 @@
   show "f \<in> borel_measurable M" by auto
 qed auto
 
-lemma borel_measurable_erealD[measurable_dest]:
+lemma%unimportant borel_measurable_erealD[measurable_dest]:
   "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
   unfolding borel_measurable_ereal_iff by simp
 
-lemma borel_measurable_ereal_iff_real:
+lemma%important borel_measurable_ereal_iff_real:
   fixes f :: "'a \<Rightarrow> ereal"
   shows "f \<in> borel_measurable M \<longleftrightarrow>
     ((\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
-proof safe
+proof%unimportant safe
   assume *: "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
   have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto
   with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
@@ -1612,15 +1612,15 @@
   finally show "f \<in> borel_measurable M" .
 qed simp_all
 
-lemma borel_measurable_ereal_iff_Iio:
+lemma%unimportant borel_measurable_ereal_iff_Iio:
   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
   by (auto simp: borel_Iio measurable_iff_measure_of)
 
-lemma borel_measurable_ereal_iff_Ioi:
+lemma%unimportant borel_measurable_ereal_iff_Ioi:
   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
   by (auto simp: borel_Ioi measurable_iff_measure_of)
 
-lemma vimage_sets_compl_iff:
+lemma%unimportant vimage_sets_compl_iff:
   "f -` A \<inter> space M \<in> sets M \<longleftrightarrow> f -` (- A) \<inter> space M \<in> sets M"
 proof -
   { fix A assume "f -` A \<inter> space M \<in> sets M"
@@ -1630,15 +1630,15 @@
     by (metis double_complement)
 qed
 
-lemma borel_measurable_iff_Iic_ereal:
+lemma%unimportant borel_measurable_iff_Iic_ereal:
   "(f::'a\<Rightarrow>ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
   unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp
 
-lemma borel_measurable_iff_Ici_ereal:
+lemma%unimportant borel_measurable_iff_Ici_ereal:
   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
   unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp
 
-lemma borel_measurable_ereal2:
+lemma%important borel_measurable_ereal2:
   fixes f g :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -1648,7 +1648,7 @@
     "(\<lambda>x. H (ereal (real_of_ereal (f x))) (-\<infinity>)) \<in> borel_measurable M"
     "(\<lambda>x. H (ereal (real_of_ereal (f x))) (\<infinity>)) \<in> borel_measurable M"
   shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
-proof -
+proof%unimportant -
   let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = - \<infinity> then H y (-\<infinity>) else H y (ereal (real_of_ereal (g x)))"
   let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = - \<infinity> then ?G (-\<infinity>) x else ?G (ereal (real_of_ereal (f x))) x"
   { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
@@ -1656,14 +1656,14 @@
   from assms show ?thesis unfolding * by simp
 qed
 
-lemma [measurable(raw)]:
+lemma%unimportant [measurable(raw)]:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows borel_measurable_ereal_add: "(\<lambda>x. f x + g x) \<in> borel_measurable M"
     and borel_measurable_ereal_times: "(\<lambda>x. f x * g x) \<in> borel_measurable M"
   by (simp_all add: borel_measurable_ereal2)
 
-lemma [measurable(raw)]:
+lemma%unimportant [measurable(raw)]:
   fixes f g :: "'a \<Rightarrow> ereal"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
@@ -1671,42 +1671,42 @@
     and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M"
   using assms by (simp_all add: minus_ereal_def divide_ereal_def)
 
-lemma borel_measurable_ereal_sum[measurable (raw)]:
+lemma%unimportant borel_measurable_ereal_sum[measurable (raw)]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
   using assms by (induction S rule: infinite_finite_induct) auto
 
-lemma borel_measurable_ereal_prod[measurable (raw)]:
+lemma%unimportant borel_measurable_ereal_prod[measurable (raw)]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
   using assms by (induction S rule: infinite_finite_induct) auto
 
-lemma borel_measurable_extreal_suminf[measurable (raw)]:
+lemma%unimportant borel_measurable_extreal_suminf[measurable (raw)]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
   unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
 
-subsection "Borel space on the extended non-negative reals"
+subsection%important "Borel space on the extended non-negative reals"
 
 text \<open> @{type ennreal} is a topological monoid, so no rules for plus are required, also all order
   statements are usually done on type classes. \<close>
 
-lemma measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma%unimportant measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel"
   by (intro borel_measurable_continuous_on1 continuous_on_enn2ereal)
 
-lemma measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma%unimportant measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
   by (intro borel_measurable_continuous_on1 continuous_on_e2ennreal)
 
-lemma borel_measurable_enn2real[measurable (raw)]:
+lemma%unimportant borel_measurable_enn2real[measurable (raw)]:
   "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. enn2real (f x)) \<in> M \<rightarrow>\<^sub>M borel"
   unfolding enn2real_def[abs_def] by measurable
 
-definition [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
+definition%important [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
 
-lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel"
+lemma%unimportant is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel"
   unfolding is_borel_def[abs_def]
 proof (safe intro!: rel_funI ext dest!: rel_fun_eq_pcr_ennreal[THEN iffD1])
   fix f and M :: "'a measure"
@@ -1718,14 +1718,14 @@
   includes ennreal.lifting
 begin
 
-lemma measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma%unimportant measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"
   unfolding is_borel_def[symmetric]
   by transfer simp
 
-lemma borel_measurable_ennreal_iff[simp]:
+lemma%important borel_measurable_ennreal_iff[simp]:
   assumes [simp]: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
   shows "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel \<longleftrightarrow> f \<in> M \<rightarrow>\<^sub>M borel"
-proof safe
+proof%unimportant safe
   assume "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel"
   then have "(\<lambda>x. enn2real (ennreal (f x))) \<in> M \<rightarrow>\<^sub>M borel"
     by measurable
@@ -1733,44 +1733,44 @@
     by (rule measurable_cong[THEN iffD1, rotated]) auto
 qed measurable
 
-lemma borel_measurable_times_ennreal[measurable (raw)]:
+lemma%unimportant borel_measurable_times_ennreal[measurable (raw)]:
   fixes f g :: "'a \<Rightarrow> ennreal"
   shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x * g x) \<in> M \<rightarrow>\<^sub>M borel"
   unfolding is_borel_def[symmetric] by transfer simp
 
-lemma borel_measurable_inverse_ennreal[measurable (raw)]:
+lemma%unimportant borel_measurable_inverse_ennreal[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> ennreal"
   shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. inverse (f x)) \<in> M \<rightarrow>\<^sub>M borel"
   unfolding is_borel_def[symmetric] by transfer simp
 
-lemma borel_measurable_divide_ennreal[measurable (raw)]:
+lemma%unimportant borel_measurable_divide_ennreal[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> ennreal"
   shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x / g x) \<in> M \<rightarrow>\<^sub>M borel"
   unfolding divide_ennreal_def by simp
 
-lemma borel_measurable_minus_ennreal[measurable (raw)]:
+lemma%unimportant borel_measurable_minus_ennreal[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> ennreal"
   shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x - g x) \<in> M \<rightarrow>\<^sub>M borel"
   unfolding is_borel_def[symmetric] by transfer simp
 
-lemma borel_measurable_prod_ennreal[measurable (raw)]:
+lemma%important borel_measurable_prod_ennreal[measurable (raw)]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ennreal"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
-  using assms by (induction S rule: infinite_finite_induct) auto
+  using%unimportant assms by (induction S rule: infinite_finite_induct) auto
 
 end
 
 hide_const (open) is_borel
 
-subsection \<open>LIMSEQ is borel measurable\<close>
+subsection%important \<open>LIMSEQ is borel measurable\<close>
 
-lemma borel_measurable_LIMSEQ_real:
+lemma%important borel_measurable_LIMSEQ_real:
   fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
   assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
   and u: "\<And>i. u i \<in> borel_measurable M"
   shows "u' \<in> borel_measurable M"
-proof -
+proof%unimportant -
   have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)"
     using u' by (simp add: lim_imp_Liminf)
   moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M"
@@ -1778,13 +1778,13 @@
   ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)
 qed
 
-lemma borel_measurable_LIMSEQ_metric:
+lemma%important borel_measurable_LIMSEQ_metric:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b :: metric_space"
   assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
   assumes lim: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. f i x) \<longlonglongrightarrow> g x"
   shows "g \<in> borel_measurable M"
   unfolding borel_eq_closed
-proof (safe intro!: measurable_measure_of)
+proof%unimportant (safe intro!: measurable_measure_of)
   fix A :: "'b set" assume "closed A"
 
   have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M"
@@ -1809,13 +1809,13 @@
   qed simp
 qed auto
 
-lemma sets_Collect_Cauchy[measurable]:
+lemma%important sets_Collect_Cauchy[measurable]:
   fixes f :: "nat \<Rightarrow> 'a => 'b::{metric_space, second_countable_topology}"
   assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
   unfolding metric_Cauchy_iff2 using f by auto
 
-lemma borel_measurable_lim_metric[measurable (raw)]:
+lemma%unimportant borel_measurable_lim_metric[measurable (raw)]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
@@ -1837,17 +1837,17 @@
     unfolding * by measurable
 qed
 
-lemma borel_measurable_suminf[measurable (raw)]:
+lemma%unimportant borel_measurable_suminf[measurable (raw)]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"
   unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
 
-lemma Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
+lemma%unimportant Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
   by (simp add: pred_def)
 
 (* Proof by Jeremy Avigad and Luke Serafin *)
-lemma isCont_borel_pred[measurable]:
+lemma%unimportant isCont_borel_pred[measurable]:
   fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
   shows "Measurable.pred borel (isCont f)"
 proof (subst measurable_cong)
@@ -1898,21 +1898,21 @@
 qed (intro pred_intros_countable closed_Collect_all closed_Collect_le open_Collect_less
            Collect_closed_imp_pred_borel closed_Collect_imp open_Collect_conj continuous_intros)
 
-lemma isCont_borel:
+lemma%unimportant isCont_borel:
   fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
   shows "{x. isCont f x} \<in> sets borel"
   by simp
 
-lemma is_real_interval:
+lemma%important is_real_interval:
   assumes S: "is_interval S"
   shows "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or> S = {a<..} \<or> S = {a..} \<or>
     S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}"
   using S unfolding is_interval_1 by (blast intro: interval_cases)
 
-lemma real_interval_borel_measurable:
+lemma%important real_interval_borel_measurable:
   assumes "is_interval (S::real set)"
   shows "S \<in> sets borel"
-proof -
+proof%unimportant -
   from assms is_real_interval have "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or>
     S = {a<..} \<or> S = {a..} \<or> S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}" by auto
   then guess a ..
@@ -1924,7 +1924,7 @@
 text \<open>The next lemmas hold in any second countable linorder (including ennreal or ereal for instance),
 but in the current state they are restricted to reals.\<close>
 
-lemma borel_measurable_mono_on_fnc:
+lemma%important borel_measurable_mono_on_fnc:
   fixes f :: "real \<Rightarrow> real" and A :: "real set"
   assumes "mono_on f A"
   shows "f \<in> borel_measurable (restrict_space borel A)"
@@ -1935,18 +1935,18 @@
               intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset)
   done
 
-lemma borel_measurable_piecewise_mono:
+lemma%unimportant borel_measurable_piecewise_mono:
   fixes f::"real \<Rightarrow> real" and C::"real set set"
   assumes "countable C" "\<And>c. c \<in> C \<Longrightarrow> c \<in> sets borel" "\<And>c. c \<in> C \<Longrightarrow> mono_on f c" "(\<Union>C) = UNIV"
   shows "f \<in> borel_measurable borel"
-by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms)
+  by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms)
 
-lemma borel_measurable_mono:
+lemma%unimportant borel_measurable_mono:
   fixes f :: "real \<Rightarrow> real"
   shows "mono f \<Longrightarrow> f \<in> borel_measurable borel"
   using borel_measurable_mono_on_fnc[of f UNIV] by (simp add: mono_def mono_on_def)
 
-lemma measurable_bdd_below_real[measurable (raw)]:
+lemma%unimportant measurable_bdd_below_real[measurable (raw)]:
   fixes F :: "'a \<Rightarrow> 'i \<Rightarrow> real"
   assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> M \<rightarrow>\<^sub>M borel"
   shows "Measurable.pred M (\<lambda>x. bdd_below ((\<lambda>i. F i x)`I))"
@@ -1957,12 +1957,12 @@
     using countable_int by measurable
 qed
 
-lemma borel_measurable_cINF_real[measurable (raw)]:
+lemma%important borel_measurable_cINF_real[measurable (raw)]:
   fixes F :: "_ \<Rightarrow> _ \<Rightarrow> real"
   assumes [simp]: "countable I"
   assumes F[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
   shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M"
-proof (rule measurable_piecewise_restrict)
+proof%unimportant (rule measurable_piecewise_restrict)
   let ?\<Omega> = "{x\<in>space M. bdd_below ((\<lambda>i. F i x)`I)}"
   show "countable {?\<Omega>, - ?\<Omega>}" "space M \<subseteq> \<Union>{?\<Omega>, - ?\<Omega>}" "\<And>X. X \<in> {?\<Omega>, - ?\<Omega>} \<Longrightarrow> X \<inter> space M \<in> sets M"
     by auto
@@ -1982,7 +1982,7 @@
   qed
 qed
 
-lemma borel_Ici: "borel = sigma UNIV (range (\<lambda>x::real. {x ..}))"
+lemma%unimportant borel_Ici: "borel = sigma UNIV (range (\<lambda>x::real. {x ..}))"
 proof (safe intro!: borel_eq_sigmaI1[OF borel_Iio])
   fix x :: real
   have eq: "{..<x} = space (sigma UNIV (range atLeast)) - {x ..}"
@@ -1991,7 +1991,7 @@
     unfolding eq by (intro sets.compl_sets) auto
 qed auto
 
-lemma borel_measurable_pred_less[measurable (raw)]:
+lemma%unimportant borel_measurable_pred_less[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
   shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> Measurable.pred M (\<lambda>w. f w < g w)"
   unfolding Measurable.pred_def by (rule borel_measurable_less)
@@ -1999,19 +1999,19 @@
 no_notation
   eucl_less (infix "<e" 50)
 
-lemma borel_measurable_Max2[measurable (raw)]:
+lemma%important borel_measurable_Max2[measurable (raw)]:
   fixes f::"_ \<Rightarrow> _ \<Rightarrow> 'a::{second_countable_topology, dense_linorder, linorder_topology}"
   assumes "finite I"
     and [measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. Max{f i x |i. i \<in> I}) \<in> borel_measurable M"
-by (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)
+by%unimportant (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)
 
-lemma measurable_compose_n [measurable (raw)]:
+lemma%unimportant measurable_compose_n [measurable (raw)]:
   assumes "T \<in> measurable M M"
   shows "(T^^n) \<in> measurable M M"
 by (induction n, auto simp add: measurable_compose[OF _ assms])
 
-lemma measurable_real_imp_nat:
+lemma%unimportant measurable_real_imp_nat:
   fixes f::"'a \<Rightarrow> nat"
   assumes [measurable]: "(\<lambda>x. real(f x)) \<in> borel_measurable M"
   shows "f \<in> measurable M (count_space UNIV)"
@@ -2023,7 +2023,7 @@
   then show ?thesis using measurable_count_space_eq2_countable by blast
 qed
 
-lemma measurable_equality_set [measurable]:
+lemma%unimportant measurable_equality_set [measurable]:
   fixes f g::"_\<Rightarrow> 'a::{second_countable_topology, t2_space}"
   assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "{x \<in> space M. f x = g x} \<in> sets M"
@@ -2038,7 +2038,7 @@
   then show ?thesis unfolding A_def by simp
 qed
 
-lemma measurable_inequality_set [measurable]:
+lemma%unimportant measurable_inequality_set [measurable]:
   fixes f g::"_ \<Rightarrow> 'a::{second_countable_topology, linorder_topology}"
   assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
@@ -2066,7 +2066,7 @@
   ultimately show "{x \<in> space M. f x > g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
 qed
 
-lemma measurable_limit [measurable]:
+lemma%unimportant measurable_limit [measurable]:
   fixes f::"nat \<Rightarrow> 'a \<Rightarrow> 'b::first_countable_topology"
   assumes [measurable]: "\<And>n::nat. f n \<in> borel_measurable M"
   shows "Measurable.pred M (\<lambda>x. (\<lambda>n. f n x) \<longlonglongrightarrow> c)"
@@ -2106,11 +2106,11 @@
   then show ?thesis by auto
 qed
 
-lemma measurable_limit2 [measurable]:
+lemma%important measurable_limit2 [measurable]:
   fixes u::"nat \<Rightarrow> 'a \<Rightarrow> real"
   assumes [measurable]: "\<And>n. u n \<in> borel_measurable M" "v \<in> borel_measurable M"
   shows "Measurable.pred M (\<lambda>x. (\<lambda>n. u n x) \<longlonglongrightarrow> v x)"
-proof -
+proof%unimportant -
   define w where "w = (\<lambda>n x. u n x - v x)"
   have [measurable]: "w n \<in> borel_measurable M" for n unfolding w_def by auto
   have "((\<lambda>n. u n x) \<longlonglongrightarrow> v x) \<longleftrightarrow> ((\<lambda>n. w n x) \<longlonglongrightarrow> 0)" for x
@@ -2118,7 +2118,7 @@
   then show ?thesis using measurable_limit by auto
 qed
 
-lemma measurable_P_restriction [measurable (raw)]:
+lemma%unimportant measurable_P_restriction [measurable (raw)]:
   assumes [measurable]: "Measurable.pred M P" "A \<in> sets M"
   shows "{x \<in> A. P x} \<in> sets M"
 proof -
@@ -2127,7 +2127,7 @@
   then show ?thesis by auto
 qed
 
-lemma measurable_sum_nat [measurable (raw)]:
+lemma%unimportant measurable_sum_nat [measurable (raw)]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> nat"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> measurable M (count_space UNIV)"
   shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> measurable M (count_space UNIV)"
@@ -2137,7 +2137,7 @@
 qed simp
 
 
-lemma measurable_abs_powr [measurable]:
+lemma%unimportant measurable_abs_powr [measurable]:
   fixes p::real
   assumes [measurable]: "f \<in> borel_measurable M"
   shows "(\<lambda>x. \<bar>f x\<bar> powr p) \<in> borel_measurable M"
@@ -2145,7 +2145,7 @@
 
 text \<open>The next one is a variation around \verb+measurable_restrict_space+.\<close>
 
-lemma measurable_restrict_space3:
+lemma%unimportant measurable_restrict_space3:
   assumes "f \<in> measurable M N" and
           "f \<in> A \<rightarrow> B"
   shows "f \<in> measurable (restrict_space M A) (restrict_space N B)"
@@ -2157,12 +2157,12 @@
 
 text \<open>The next one is a variation around \verb+measurable_piecewise_restrict+.\<close>
 
-lemma measurable_piecewise_restrict2:
+lemma%important measurable_piecewise_restrict2:
   assumes [measurable]: "\<And>n. A n \<in> sets M"
       and "space M = (\<Union>(n::nat). A n)"
           "\<And>n. \<exists>h \<in> measurable M N. (\<forall>x \<in> A n. f x = h x)"
   shows "f \<in> measurable M N"
-proof (rule measurableI)
+proof%unimportant (rule measurableI)
   fix B assume [measurable]: "B \<in> sets N"
   {
     fix n::nat